Documentation

Mathlib.Data.Prod.Basic

Extra facts about Prod #

This file defines Prod.swap : α × β → β × α and proves various simple lemmas about Prod. It also defines better delaborators for product projections.

@[deprecated Prod.map_apply]
theorem Prod_map {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} (f : αβ) (g : γδ) (x : α) (y : γ) :
Prod.map f g (x, y) = (f x, g y)

Alias of Prod.map_apply.

def Prod.mk.injArrow {α : Type u_1} {β : Type u_2} {x₁ : α} {y₁ : β} {x₂ : α} {y₂ : β} :
(x₁, y₁) = (x₂, y₂)P : Sort u_5⦄ → (x₁ = x₂y₁ = y₂P)P
Equations
Instances For
    @[simp]
    theorem Prod.mk.eta {α : Type u_1} {β : Type u_2} {p : α × β} :
    (p.1, p.2) = p
    theorem Prod.forall' {α : Type u_1} {β : Type u_2} {p : αβProp} :
    (∀ (x : α × β), p x.1 x.2) ∀ (a : α) (b : β), p a b
    theorem Prod.exists' {α : Type u_1} {β : Type u_2} {p : αβProp} :
    (∃ (x : α × β), p x.1 x.2) ∃ (a : α), ∃ (b : β), p a b
    @[simp]
    theorem Prod.snd_comp_mk {α : Type u_1} {β : Type u_2} (x : α) :
    Prod.snd Prod.mk x = id
    @[simp]
    theorem Prod.fst_comp_mk {α : Type u_1} {β : Type u_2} (x : α) :
    Prod.fst Prod.mk x = Function.const β x
    @[simp]
    theorem Prod.map_mk {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} (f : αγ) (g : βδ) (a : α) (b : β) :
    Prod.map f g (a, b) = (f a, g b)
    theorem Prod.map_apply' {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} (f : αγ) (g : βδ) (p : α × β) :
    Prod.map f g p = (f p.1, g p.2)
    theorem Prod.map_fst' {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} (f : αγ) (g : βδ) :
    Prod.fst Prod.map f g = f Prod.fst
    theorem Prod.map_snd' {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} (f : αγ) (g : βδ) :
    Prod.snd Prod.map f g = g Prod.snd
    theorem Prod.map_comp_map {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {ε : Type u_5} {ζ : Type u_6} (f : αβ) (f' : γδ) (g : βε) (g' : δζ) :
    Prod.map g g' Prod.map f f' = Prod.map (g f) (g' f')

    Composing a Prod.map with another Prod.map is equal to a single Prod.map of composed functions.

    theorem Prod.map_map {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {ε : Type u_5} {ζ : Type u_6} (f : αβ) (f' : γδ) (g : βε) (g' : δζ) (x : α × γ) :
    Prod.map g g' (Prod.map f f' x) = Prod.map (g f) (g' f') x

    Composing a Prod.map with another Prod.map is equal to a single Prod.map of composed functions, fully applied.

    theorem Prod.mk.inj_iff {α : Type u_1} {β : Type u_2} {a₁ : α} {a₂ : α} {b₁ : β} {b₂ : β} :
    (a₁, b₁) = (a₂, b₂) a₁ = a₂ b₁ = b₂
    theorem Prod.mk.inj_left {α : Type u_5} {β : Type u_6} (a : α) :
    theorem Prod.mk.inj_right {α : Type u_5} {β : Type u_6} (b : β) :
    Function.Injective fun (a : α) => (a, b)
    theorem Prod.mk_inj_left {α : Type u_1} {β : Type u_2} {a : α} {b₁ : β} {b₂ : β} :
    (a, b₁) = (a, b₂) b₁ = b₂
    theorem Prod.mk_inj_right {α : Type u_1} {β : Type u_2} {a₁ : α} {a₂ : α} {b : β} :
    (a₁, b) = (a₂, b) a₁ = a₂
    theorem Prod.map_def {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {f : αγ} {g : βδ} :
    Prod.map f g = fun (p : α × β) => (f p.1, g p.2)
    theorem Prod.id_prod {α : Type u_1} {β : Type u_2} :
    (fun (p : α × β) => (p.1, p.2)) = id
    @[simp]
    theorem Prod.map_id {α : Type u_1} {β : Type u_2} :
    Prod.map id id = id
    @[simp]
    theorem Prod.map_id' {α : Type u_1} {β : Type u_2} :
    (Prod.map (fun (a : α) => a) fun (b : β) => b) = fun (x : α × β) => x
    @[simp]
    theorem Prod.map_iterate {α : Type u_1} {β : Type u_2} (f : αα) (g : ββ) (n : ) :
    @[deprecated Prod.map_iterate]
    theorem Prod.iterate_prod_map {α : Type u_1} {β : Type u_2} (f : αα) (g : ββ) (n : ) :

    Alias of Prod.map_iterate.

    theorem Prod.fst_surjective {α : Type u_1} {β : Type u_2} [h : Nonempty β] :
    theorem Prod.snd_surjective {α : Type u_1} {β : Type u_2} [h : Nonempty α] :
    theorem Prod.fst_injective {α : Type u_1} {β : Type u_2} [Subsingleton β] :
    theorem Prod.snd_injective {α : Type u_1} {β : Type u_2} [Subsingleton α] :
    def Prod.swap {α : Type u_1} {β : Type u_2} :
    α × ββ × α

    Swap the factors of a product. swap (a, b) = (b, a)

    Equations
    • p.swap = (p.2, p.1)
    Instances For
      @[simp]
      theorem Prod.swap_swap {α : Type u_1} {β : Type u_2} (x : α × β) :
      x.swap.swap = x
      @[simp]
      theorem Prod.fst_swap {α : Type u_1} {β : Type u_2} {p : α × β} :
      p.swap.1 = p.2
      @[simp]
      theorem Prod.snd_swap {α : Type u_1} {β : Type u_2} {p : α × β} :
      p.swap.2 = p.1
      @[simp]
      theorem Prod.swap_prod_mk {α : Type u_1} {β : Type u_2} {a : α} {b : β} :
      (a, b).swap = (b, a)
      @[simp]
      theorem Prod.swap_swap_eq {α : Type u_1} {β : Type u_2} :
      Prod.swap Prod.swap = id
      @[simp]
      theorem Prod.swap_leftInverse {α : Type u_1} {β : Type u_2} :
      Function.LeftInverse Prod.swap Prod.swap
      @[simp]
      theorem Prod.swap_rightInverse {α : Type u_1} {β : Type u_2} :
      Function.RightInverse Prod.swap Prod.swap
      theorem Prod.swap_injective {α : Type u_1} {β : Type u_2} :
      theorem Prod.swap_surjective {α : Type u_1} {β : Type u_2} :
      theorem Prod.swap_bijective {α : Type u_1} {β : Type u_2} :
      @[simp]
      theorem Prod.swap_inj {α : Type u_1} {β : Type u_2} {p : α × β} {q : α × β} :
      p.swap = q.swap p = q
      theorem Prod.map_comp_swap {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} (f : αβ) (g : γδ) :
      Prod.map f g Prod.swap = Prod.swap Prod.map g f

      For two functions f and g, the composition of Prod.map f g with Prod.swap is equal to the composition of Prod.swap with Prod.map g f.

      theorem Function.Semiconj.swap_map {α : Type u_1} {β : Type u_2} (f : αα) (g : ββ) :
      Function.Semiconj Prod.swap (Prod.map f g) (Prod.map g f)
      theorem Prod.eq_iff_fst_eq_snd_eq {α : Type u_1} {β : Type u_2} {p : α × β} {q : α × β} :
      p = q p.1 = q.1 p.2 = q.2
      theorem Prod.fst_eq_iff {α : Type u_1} {β : Type u_2} {p : α × β} {x : α} :
      p.1 = x p = (x, p.2)
      theorem Prod.snd_eq_iff {α : Type u_1} {β : Type u_2} {p : α × β} {x : β} :
      p.2 = x p = (p.1, x)
      theorem Prod.lex_iff {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {x : α × β} {y : α × β} :
      Prod.Lex r s x y r x.1 y.1 x.1 = y.1 s x.2 y.2
      instance Prod.Lex.decidable {α : Type u_1} {β : Type u_2} [DecidableEq α] (r : ααProp) (s : ββProp) [DecidableRel r] [DecidableRel s] :
      Equations
      theorem Prod.Lex.refl_left {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) [IsRefl α r] (x : α × β) :
      Prod.Lex r s x x
      instance Prod.instIsReflLex {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} [IsRefl α r] :
      IsRefl (α × β) (Prod.Lex r s)
      Equations
      • =
      theorem Prod.Lex.refl_right {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) [IsRefl β s] (x : α × β) :
      Prod.Lex r s x x
      instance Prod.instIsReflLex_1 {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} [IsRefl β s] :
      IsRefl (α × β) (Prod.Lex r s)
      Equations
      • =
      instance Prod.isIrrefl {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} [IsIrrefl α r] [IsIrrefl β s] :
      IsIrrefl (α × β) (Prod.Lex r s)
      Equations
      • =
      theorem Prod.Lex.trans {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} [IsTrans α r] [IsTrans β s] {x : α × β} {y : α × β} {z : α × β} :
      Prod.Lex r s x yProd.Lex r s y zProd.Lex r s x z
      instance Prod.instIsTransLex {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} [IsTrans α r] [IsTrans β s] :
      IsTrans (α × β) (Prod.Lex r s)
      Equations
      • =
      instance Prod.instIsAntisymmLexOfIsStrictOrder {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} [IsStrictOrder α r] [IsAntisymm β s] :
      IsAntisymm (α × β) (Prod.Lex r s)
      Equations
      • =
      instance Prod.isTotal_left {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} [IsTotal α r] :
      IsTotal (α × β) (Prod.Lex r s)
      Equations
      • =
      instance Prod.isTotal_right {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} [IsTrichotomous α r] [IsTotal β s] :
      IsTotal (α × β) (Prod.Lex r s)
      Equations
      • =
      instance Prod.IsTrichotomous {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} [IsTrichotomous α r] [IsTrichotomous β s] :
      IsTrichotomous (α × β) (Prod.Lex r s)
      Equations
      • =
      theorem Function.Injective.prodMap {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {f : αγ} {g : βδ} (hf : Function.Injective f) (hg : Function.Injective g) :
      theorem Function.Surjective.prodMap {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {f : αγ} {g : βδ} (hf : Function.Surjective f) (hg : Function.Surjective g) :
      theorem Function.Bijective.prodMap {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {f : αγ} {g : βδ} (hf : Function.Bijective f) (hg : Function.Bijective g) :
      theorem Function.LeftInverse.prodMap {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {f₁ : αβ} {g₁ : γδ} {f₂ : βα} {g₂ : δγ} (hf : Function.LeftInverse f₁ f₂) (hg : Function.LeftInverse g₁ g₂) :
      Function.LeftInverse (Prod.map f₁ g₁) (Prod.map f₂ g₂)
      theorem Function.RightInverse.prodMap {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {f₁ : αβ} {g₁ : γδ} {f₂ : βα} {g₂ : δγ} :
      Function.RightInverse f₁ f₂Function.RightInverse g₁ g₂Function.RightInverse (Prod.map f₁ g₁) (Prod.map f₂ g₂)
      theorem Function.Involutive.prodMap {α : Type u_1} {β : Type u_2} {f : αα} {g : ββ} :
      @[deprecated Function.Injective.prodMap]
      theorem Function.Injective.Prod_map {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {f : αγ} {g : βδ} (hf : Function.Injective f) (hg : Function.Injective g) :

      Alias of Function.Injective.prodMap.

      @[deprecated Function.Surjective.prodMap]
      theorem Function.Surjective.Prod_map {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {f : αγ} {g : βδ} (hf : Function.Surjective f) (hg : Function.Surjective g) :

      Alias of Function.Surjective.prodMap.

      @[deprecated Function.Bijective.prodMap]
      theorem Function.Bijective.Prod_map {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {f : αγ} {g : βδ} (hf : Function.Bijective f) (hg : Function.Bijective g) :

      Alias of Function.Bijective.prodMap.

      @[deprecated Function.LeftInverse.prodMap]
      theorem Function.LeftInverse.Prod_map {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {f₁ : αβ} {g₁ : γδ} {f₂ : βα} {g₂ : δγ} (hf : Function.LeftInverse f₁ f₂) (hg : Function.LeftInverse g₁ g₂) :
      Function.LeftInverse (Prod.map f₁ g₁) (Prod.map f₂ g₂)

      Alias of Function.LeftInverse.prodMap.

      @[deprecated Function.RightInverse.prodMap]
      theorem Function.RightInverse.Prod_map {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {f₁ : αβ} {g₁ : γδ} {f₂ : βα} {g₂ : δγ} :
      Function.RightInverse f₁ f₂Function.RightInverse g₁ g₂Function.RightInverse (Prod.map f₁ g₁) (Prod.map f₂ g₂)

      Alias of Function.RightInverse.prodMap.

      @[deprecated Function.Involutive.prodMap]
      theorem Function.Involutive.Prod_map {α : Type u_1} {β : Type u_2} {f : αα} {g : ββ} :

      Alias of Function.Involutive.prodMap.

      @[simp]
      theorem Prod.map_injective {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} [Nonempty α] [Nonempty β] {f : αγ} {g : βδ} :
      @[simp]
      theorem Prod.map_surjective {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} [Nonempty γ] [Nonempty δ] {f : αγ} {g : βδ} :
      @[simp]
      theorem Prod.map_bijective {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} [Nonempty α] [Nonempty β] {f : αγ} {g : βδ} :
      @[simp]
      theorem Prod.map_leftInverse {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} [Nonempty β] [Nonempty δ] {f₁ : αβ} {g₁ : γδ} {f₂ : βα} {g₂ : δγ} :
      @[simp]
      theorem Prod.map_rightInverse {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} [Nonempty α] [Nonempty γ] {f₁ : αβ} {g₁ : γδ} {f₂ : βα} {g₂ : δγ} :
      @[simp]
      theorem Prod.map_involutive {α : Type u_1} {β : Type u_2} [Nonempty α] [Nonempty β] {f : αα} {g : ββ} :

      Delaborator for Prod.fst x as x.1.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For

        Delaborator for Prod.snd x as x.2.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For