@[instance]
def
category_theory.prod
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D] :
category_theory.category (C × D)
prod C D
gives the cartesian product of two categories.
Equations
- category_theory.prod C D = {to_category_struct := {to_has_hom := {hom := λ (X Y : C × D), (X.fst ⟶ Y.fst) × (X.snd ⟶ Y.snd)}, id := λ (X : C × D), (𝟙 X.fst, 𝟙 X.snd), comp := λ (_x _x_1 _x_2 : C × D) (f : _x ⟶ _x_1) (g : _x_1 ⟶ _x_2), (f.fst ≫ g.fst, f.snd ≫ g.snd)}, id_comp' := _, comp_id' := _, assoc' := _}
@[simp]
theorem
category_theory.prod_id
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(X : C)
(Y : D) :
@[simp]
theorem
category_theory.prod_comp
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
{P Q R : C}
{S T U : D}
(f : (P, S) ⟶ (Q, T))
(g : (Q, T) ⟶ (R, U)) :
@[simp]
theorem
category_theory.prod_id_fst
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(X : C × D) :
@[simp]
theorem
category_theory.prod_id_snd
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(X : C × D) :
@[simp]
theorem
category_theory.prod_comp_fst
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
{X Y Z : C × D}
(f : X ⟶ Y)
(g : Y ⟶ Z) :
@[simp]
theorem
category_theory.prod_comp_snd
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
{X Y Z : C × D}
(f : X ⟶ Y)
(g : Y ⟶ Z) :
@[instance]
def
category_theory.uniform_prod
(C : Type u₁)
[category_theory.category C]
(D : Type u₁)
[category_theory.category D] :
category_theory.category (C × D)
prod.category.uniform C D
is an additional instance specialised so both factors have the same
universe levels. This helps typeclass resolution.
Equations
@[simp]
theorem
category_theory.prod.sectl_map
(C : Type u₁)
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(Z : D)
(X Y : C)
(f : X ⟶ Y) :
(category_theory.prod.sectl C Z).map f = (f, 𝟙 Z)
def
category_theory.prod.sectl
(C : Type u₁)
[category_theory.category C]
{D : Type u₂}
[category_theory.category D] :
sectl C Z
is the functor C ⥤ C × D
given by X ↦ (X, Z)
.
@[simp]
theorem
category_theory.prod.sectl_obj
(C : Type u₁)
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(Z : D)
(X : C) :
(category_theory.prod.sectl C Z).obj X = (X, Z)
def
category_theory.prod.sectr
{C : Type u₁}
[category_theory.category C]
(Z : C)
(D : Type u₂)
[category_theory.category D] :
sectr Z D
is the functor D ⥤ C × D
given by Y ↦ (Z, Y)
.
@[simp]
theorem
category_theory.prod.sectr_obj
{C : Type u₁}
[category_theory.category C]
(Z : C)
(D : Type u₂)
[category_theory.category D]
(X : D) :
(category_theory.prod.sectr Z D).obj X = (Z, X)
@[simp]
theorem
category_theory.prod.sectr_map
{C : Type u₁}
[category_theory.category C]
(Z : C)
(D : Type u₂)
[category_theory.category D]
(X Y : D)
(f : X ⟶ Y) :
(category_theory.prod.sectr Z D).map f = (𝟙 Z, f)
@[simp]
theorem
category_theory.prod.fst_obj
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(X : C × D) :
(category_theory.prod.fst C D).obj X = X.fst
def
category_theory.prod.fst
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D] :
fst
is the functor (X, Y) ↦ X
.
@[simp]
theorem
category_theory.prod.fst_map
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(X Y : C × D)
(f : X ⟶ Y) :
(category_theory.prod.fst C D).map f = f.fst
@[simp]
theorem
category_theory.prod.snd_map
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(X Y : C × D)
(f : X ⟶ Y) :
(category_theory.prod.snd C D).map f = f.snd
@[simp]
theorem
category_theory.prod.snd_obj
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(X : C × D) :
(category_theory.prod.snd C D).obj X = X.snd
def
category_theory.prod.snd
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D] :
snd
is the functor (X, Y) ↦ Y
.
@[simp]
theorem
category_theory.prod.swap_obj
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(X : C × D) :
(category_theory.prod.swap C D).obj X = (X.snd, X.fst)
def
category_theory.prod.swap
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D] :
@[simp]
theorem
category_theory.prod.swap_map
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(_x _x_1 : C × D)
(f : _x ⟶ _x_1) :
(category_theory.prod.swap C D).map f = (f.snd, f.fst)
@[simp]
theorem
category_theory.prod.symmetry_hom_app
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(X : C × D) :
(category_theory.prod.symmetry C D).hom.app X = 𝟙 X
@[simp]
theorem
category_theory.prod.symmetry_inv_app
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(X : C × D) :
(category_theory.prod.symmetry C D).inv.app X = 𝟙 X
def
category_theory.prod.symmetry
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D] :
category_theory.prod.swap C D ⋙ category_theory.prod.swap D C ≅ 𝟭 (C × D)
Equations
- category_theory.prod.symmetry C D = {hom := {app := λ (X : C × D), 𝟙 X, naturality' := _}, inv := {app := λ (X : C × D), 𝟙 X, naturality' := _}, hom_inv_id' := _, inv_hom_id' := _}
def
category_theory.prod.braiding
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D] :
Equations
- category_theory.prod.braiding C D = category_theory.equivalence.mk (category_theory.prod.swap C D) (category_theory.prod.swap D C) (category_theory.nat_iso.of_components (λ (X : C × D), category_theory.eq_to_iso _) _) (category_theory.nat_iso.of_components (λ (X : D × C), category_theory.eq_to_iso _) _)
@[instance]
def
category_theory.prod.swap_is_equivalence
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D] :
@[simp]
theorem
category_theory.evaluation_obj_obj
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(X : C)
(F : C ⥤ D) :
((category_theory.evaluation C D).obj X).obj F = F.obj X
def
category_theory.evaluation
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D] :
@[simp]
theorem
category_theory.evaluation_map_app
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(X Y : C)
(f : X ⟶ Y)
(F : C ⥤ D) :
((category_theory.evaluation C D).map f).app F = F.map f
@[simp]
theorem
category_theory.evaluation_obj_map
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(X : C)
(F G : C ⥤ D)
(α : F ⟶ G) :
((category_theory.evaluation C D).obj X).map α = α.app X
@[simp]
theorem
category_theory.evaluation_uncurried_map
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(x y : C × (C ⥤ D))
(f : x ⟶ y) :
def
category_theory.evaluation_uncurried
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D] :
@[simp]
theorem
category_theory.evaluation_uncurried_obj
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D]
(p : C × (C ⥤ D)) :
@[simp]
theorem
category_theory.functor.prod_map
{A : Type u₁}
[category_theory.category A]
{B : Type u₂}
[category_theory.category B]
{C : Type u₃}
[category_theory.category C]
{D : Type u₄}
[category_theory.category D]
(F : A ⥤ B)
(G : C ⥤ D)
(_x _x_1 : A × C)
(f : _x ⟶ _x_1) :
def
category_theory.functor.prod
{A : Type u₁}
[category_theory.category A]
{B : Type u₂}
[category_theory.category B]
{C : Type u₃}
[category_theory.category C]
{D : Type u₄}
[category_theory.category D] :
The cartesian product of two functors.
@[simp]
theorem
category_theory.functor.prod_obj
{A : Type u₁}
[category_theory.category A]
{B : Type u₂}
[category_theory.category B]
{C : Type u₃}
[category_theory.category C]
{D : Type u₄}
[category_theory.category D]
(F : A ⥤ B)
(G : C ⥤ D)
(X : A × C) :
def
category_theory.nat_trans.prod
{A : Type u₁}
[category_theory.category A]
{B : Type u₂}
[category_theory.category B]
{C : Type u₃}
[category_theory.category C]
{D : Type u₄}
[category_theory.category D]
{F G : A ⥤ B}
{H I : C ⥤ D} :
The cartesian product of two natural transformations.
Equations
- category_theory.nat_trans.prod α β = {app := λ (X : A × C), (α.app X.fst, β.app X.snd), naturality' := _}
@[simp]
theorem
category_theory.nat_trans.prod_app
{A : Type u₁}
[category_theory.category A]
{B : Type u₂}
[category_theory.category B]
{C : Type u₃}
[category_theory.category C]
{D : Type u₄}
[category_theory.category D]
{F G : A ⥤ B}
{H I : C ⥤ D}
(α : F ⟶ G)
(β : H ⟶ I)
(X : A × C) :