structure
category_theory.equivalence
(C : Type u₁)
[category_theory.category C]
(D : Type u₂)
[category_theory.category D] :
Type (max u₁ u₂ v₁ v₂)
- functor : C ⥤ D
- inverse : D ⥤ C
- unit_iso : 𝟭 C ≅ c.functor ⋙ c.inverse
- counit_iso : c.inverse ⋙ c.functor ≅ 𝟭 D
- functor_unit_iso_comp' : (∀ (X : C), c.functor.map (c.unit_iso.hom.app X) ≫ c.counit_iso.hom.app (c.functor.obj X) = 𝟙 (c.functor.obj X)) . "obviously"
We define an equivalence as a (half)-adjoint equivalence, a pair of functors with
a unit and counit which are natural isomorphisms and the triangle law Fη ≫ εF = 1
, or in other
words the composite F ⟶ FGF ⟶ F
is the identity.
The triangle equation is written as a family of equalities between morphisms, it is more
complicated if we write it as an equality of natural transformations, because then we would have
to insert natural transformations like F ⟶ F1
.
def
category_theory.equivalence.unit
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D) :
def
category_theory.equivalence.counit
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D) :
def
category_theory.equivalence.unit_inv
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D) :
def
category_theory.equivalence.counit_inv
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D) :
@[simp]
theorem
category_theory.equivalence.equivalence_mk'_unit
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(functor : C ⥤ D)
(inverse : D ⥤ C)
(unit_iso : 𝟭 C ≅ functor ⋙ inverse)
(counit_iso : inverse ⋙ functor ≅ 𝟭 D)
(f : (∀ (X : C), functor.map (unit_iso.hom.app X) ≫ counit_iso.hom.app (functor.obj X) = 𝟙 (functor.obj X)) . "obviously") :
{functor := functor, inverse := inverse, unit_iso := unit_iso, counit_iso := counit_iso, functor_unit_iso_comp' := f}.unit = unit_iso.hom
@[simp]
theorem
category_theory.equivalence.equivalence_mk'_counit
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(functor : C ⥤ D)
(inverse : D ⥤ C)
(unit_iso : 𝟭 C ≅ functor ⋙ inverse)
(counit_iso : inverse ⋙ functor ≅ 𝟭 D)
(f : (∀ (X : C), functor.map (unit_iso.hom.app X) ≫ counit_iso.hom.app (functor.obj X) = 𝟙 (functor.obj X)) . "obviously") :
{functor := functor, inverse := inverse, unit_iso := unit_iso, counit_iso := counit_iso, functor_unit_iso_comp' := f}.counit = counit_iso.hom
@[simp]
theorem
category_theory.equivalence.equivalence_mk'_unit_inv
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(functor : C ⥤ D)
(inverse : D ⥤ C)
(unit_iso : 𝟭 C ≅ functor ⋙ inverse)
(counit_iso : inverse ⋙ functor ≅ 𝟭 D)
(f : (∀ (X : C), functor.map (unit_iso.hom.app X) ≫ counit_iso.hom.app (functor.obj X) = 𝟙 (functor.obj X)) . "obviously") :
{functor := functor, inverse := inverse, unit_iso := unit_iso, counit_iso := counit_iso, functor_unit_iso_comp' := f}.unit_inv = unit_iso.inv
@[simp]
theorem
category_theory.equivalence.equivalence_mk'_counit_inv
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(functor : C ⥤ D)
(inverse : D ⥤ C)
(unit_iso : 𝟭 C ≅ functor ⋙ inverse)
(counit_iso : inverse ⋙ functor ≅ 𝟭 D)
(f : (∀ (X : C), functor.map (unit_iso.hom.app X) ≫ counit_iso.hom.app (functor.obj X) = 𝟙 (functor.obj X)) . "obviously") :
{functor := functor, inverse := inverse, unit_iso := unit_iso, counit_iso := counit_iso, functor_unit_iso_comp' := f}.counit_inv = counit_iso.inv
@[simp]
theorem
category_theory.equivalence.counit_inv_functor_comp
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
(X : C) :
theorem
category_theory.equivalence.functor_unit
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
(X : C) :
theorem
category_theory.equivalence.counit_functor
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
(X : C) :
@[simp]
theorem
category_theory.equivalence.unit_inverse_comp
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
(Y : D) :
The other triangle equality. The proof follows the following proof in Globular: http://globular.science/1905.001
@[simp]
theorem
category_theory.equivalence.inverse_counit_inv_comp
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
(Y : D) :
theorem
category_theory.equivalence.unit_inverse
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
(Y : D) :
theorem
category_theory.equivalence.inverse_counit
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
(Y : D) :
@[simp]
theorem
category_theory.equivalence.fun_inv_map
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
(X Y : D)
(f : X ⟶ Y) :
def
category_theory.equivalence.adjointify_η
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{F : C ⥤ D}
{G : D ⥤ C} :
Equations
- category_theory.equivalence.adjointify_η η ε = (((((η ≪≫ category_theory.iso_whisker_left F G.left_unitor.symm) ≪≫ category_theory.iso_whisker_left F (category_theory.iso_whisker_right ε.symm G)) ≪≫ category_theory.iso_whisker_left F (G.associator F G)) ≪≫ (F.associator G (F ⋙ G)).symm) ≪≫ category_theory.iso_whisker_right η.symm (F ⋙ G)) ≪≫ (F ⋙ G).left_unitor
theorem
category_theory.equivalence.adjointify_η_ε
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{F : C ⥤ D}
{G : D ⥤ C}
(η : 𝟭 C ≅ F ⋙ G)
(ε : G ⋙ F ≅ 𝟭 D)
(X : C) :
def
category_theory.equivalence.mk
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
(G : D ⥤ C) :
Equations
- category_theory.equivalence.mk F G η ε = {functor := F, inverse := G, unit_iso := category_theory.equivalence.adjointify_η η ε, counit_iso := ε, functor_unit_iso_comp' := _}
@[simp]
@[simp]
@[simp]
Equations
- category_theory.equivalence.refl = {functor := 𝟭 C _inst_1, inverse := 𝟭 C _inst_1, unit_iso := category_theory.iso.refl (𝟭 C), counit_iso := category_theory.iso.refl (𝟭 C ⋙ 𝟭 C), functor_unit_iso_comp' := _}
@[simp]
@[simp]
theorem
category_theory.equivalence.symm_inverse
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D) :
@[simp]
theorem
category_theory.equivalence.symm_unit_iso
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D) :
e.symm.unit_iso = e.counit_iso.symm
@[simp]
theorem
category_theory.equivalence.symm_functor
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D) :
def
category_theory.equivalence.symm
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D] :
Equations
- e.symm = {functor := e.inverse, inverse := e.functor, unit_iso := e.counit_iso.symm, counit_iso := e.unit_iso.symm, functor_unit_iso_comp' := _}
@[simp]
theorem
category_theory.equivalence.symm_counit_iso
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D) :
e.symm.counit_iso = e.unit_iso.symm
@[simp]
theorem
category_theory.equivalence.trans_counit_iso
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{E : Type u₃}
[category_theory.category E]
(e : C ≌ D)
(f : D ≌ E) :
@[simp]
theorem
category_theory.equivalence.trans_inverse
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{E : Type u₃}
[category_theory.category E]
(e : C ≌ D)
(f : D ≌ E) :
@[simp]
theorem
category_theory.equivalence.trans_unit_iso
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{E : Type u₃}
[category_theory.category E]
(e : C ≌ D)
(f : D ≌ E) :
@[simp]
theorem
category_theory.equivalence.trans_functor
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{E : Type u₃}
[category_theory.category E]
(e : C ≌ D)
(f : D ≌ E) :
def
category_theory.equivalence.trans
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{E : Type u₃}
[category_theory.category E] :
Equations
- e.trans f = {functor := e.functor ⋙ f.functor, inverse := f.inverse ⋙ e.inverse, unit_iso := e.unit_iso ≪≫ category_theory.iso_whisker_left e.functor (category_theory.iso_whisker_right f.unit_iso e.inverse), counit_iso := category_theory.iso_whisker_left f.inverse (category_theory.iso_whisker_right e.counit_iso f.functor) ≪≫ f.counit_iso, functor_unit_iso_comp' := _}
def
category_theory.equivalence.fun_inv_id_assoc
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{E : Type u₃}
[category_theory.category E]
(e : C ≌ D)
(F : C ⥤ E) :
Equations
- e.fun_inv_id_assoc F = (e.functor.associator e.inverse F).symm ≪≫ category_theory.iso_whisker_right e.unit_iso.symm F ≪≫ F.left_unitor
@[simp]
theorem
category_theory.equivalence.fun_inv_id_assoc_hom_app
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{E : Type u₃}
[category_theory.category E]
(e : C ≌ D)
(F : C ⥤ E)
(X : C) :
@[simp]
theorem
category_theory.equivalence.fun_inv_id_assoc_inv_app
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{E : Type u₃}
[category_theory.category E]
(e : C ≌ D)
(F : C ⥤ E)
(X : C) :
def
category_theory.equivalence.inv_fun_id_assoc
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{E : Type u₃}
[category_theory.category E]
(e : C ≌ D)
(F : D ⥤ E) :
Equations
- e.inv_fun_id_assoc F = (e.inverse.associator e.functor F).symm ≪≫ category_theory.iso_whisker_right e.counit_iso F ≪≫ F.left_unitor
@[simp]
theorem
category_theory.equivalence.inv_fun_id_assoc_hom_app
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{E : Type u₃}
[category_theory.category E]
(e : C ≌ D)
(F : D ⥤ E)
(X : D) :
@[simp]
theorem
category_theory.equivalence.inv_fun_id_assoc_inv_app
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{E : Type u₃}
[category_theory.category E]
(e : C ≌ D)
(F : D ⥤ E)
(X : D) :
(e.inv_fun_id_assoc F).inv.app X = F.map (e.counit_inv.app X)
@[simp]
theorem
category_theory.equivalence.cancel_unit_right
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
{X Y : C}
(f f' : X ⟶ Y) :
@[simp]
theorem
category_theory.equivalence.cancel_counit_inv_right
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
{X Y : D}
(f f' : X ⟶ Y) :
f ≫ e.counit_inv.app Y = f' ≫ e.counit_inv.app Y ↔ f = f'
@[simp]
@[simp]
theorem
category_theory.equivalence.cancel_counit_inv_right_assoc
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
{W X X' Y : D}
(f : W ⟶ X)
(g : X ⟶ Y)
(f' : W ⟶ X')
(g' : X' ⟶ Y) :
@[simp]
theorem
category_theory.equivalence.cancel_unit_right_assoc'
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
{W X X' Y Y' Z : C}
(f : W ⟶ X)
(g : X ⟶ Y)
(h : Y ⟶ Z)
(f' : W ⟶ X')
(g' : X' ⟶ Y')
(h' : Y' ⟶ Z) :
@[simp]
theorem
category_theory.equivalence.cancel_counit_inv_right_assoc'
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
{W X X' Y Y' Z : D}
(f : W ⟶ X)
(g : X ⟶ Y)
(h : Y ⟶ Z)
(f' : W ⟶ X')
(g' : X' ⟶ Y')
(h' : Y' ⟶ Z) :
Powers of an auto-equivalence.
@[instance]
Equations
@[simp]
theorem
category_theory.equivalence.pow_zero
{C : Type u₁}
[category_theory.category C]
(e : C ≌ C) :
@[simp]
theorem
category_theory.equivalence.pow_one
{C : Type u₁}
[category_theory.category C]
(e : C ≌ C) :
@[simp]
theorem
category_theory.equivalence.pow_minus_one
{C : Type u₁}
[category_theory.category C]
(e : C ≌ C) :
@[class]
structure
category_theory.is_equivalence
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D] :
C ⥤ D → Type (max u₁ u₂ v₁ v₂)
- inverse : D ⥤ C
- unit_iso : 𝟭 C ≅ F ⋙ category_theory.is_equivalence.inverse F
- counit_iso : category_theory.is_equivalence.inverse F ⋙ F ≅ 𝟭 D
- functor_unit_iso_comp' : (∀ (X : C), F.map (category_theory.is_equivalence.unit_iso.hom.app X) ≫ category_theory.is_equivalence.counit_iso.hom.app (F.obj X) = 𝟙 (F.obj X)) . "obviously"
A functor that is part of a (half) adjoint equivalence
Instances
- category_theory.is_equivalence.of_equivalence
- category_theory.is_equivalence.of_equivalence_inverse
- category_theory.functor.is_equivalence_refl
- category_theory.functor.is_equivalence_inv
- category_theory.functor.is_equivalence_trans
- category_theory.prod.swap_is_equivalence
- category_theory.forget₂.category_theory.is_equivalence
- category_theory.monadic_right_adjoint.eqv
- category_theory.prod.associator_is_equivalence
- category_theory.prod.inverse_associator_is_equivalence
- category_theory.sum.swap.is_equivalence
- category_theory.sum.associator_is_equivalence
- category_theory.sum.inverse_associator_is_equivalence
@[instance]
def
category_theory.is_equivalence.of_equivalence
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ≌ D) :
Equations
- category_theory.is_equivalence.of_equivalence F = {inverse := F.inverse, unit_iso := F.unit_iso, counit_iso := F.counit_iso, functor_unit_iso_comp' := _}
@[instance]
def
category_theory.is_equivalence.of_equivalence_inverse
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ≌ D) :
def
category_theory.is_equivalence.mk
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{F : C ⥤ D}
(G : D ⥤ C) :
Equations
- category_theory.is_equivalence.mk G η ε = {inverse := G, unit_iso := category_theory.equivalence.adjointify_η η ε, counit_iso := ε, functor_unit_iso_comp' := _}
def
category_theory.functor.as_equivalence
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.is_equivalence F] :
C ≌ D
Equations
- F.as_equivalence = {functor := F, inverse := category_theory.is_equivalence.inverse F _inst_3, unit_iso := category_theory.is_equivalence.unit_iso _inst_3, counit_iso := category_theory.is_equivalence.counit_iso _inst_3, functor_unit_iso_comp' := _}
@[instance]
def
category_theory.functor.inv
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.is_equivalence F] :
D ⥤ C
Equations
@[instance]
def
category_theory.functor.is_equivalence_inv
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.is_equivalence F] :
@[simp]
theorem
category_theory.functor.as_equivalence_functor
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.is_equivalence F] :
F.as_equivalence.functor = F
@[simp]
theorem
category_theory.functor.as_equivalence_inverse
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.is_equivalence F] :
F.as_equivalence.inverse = F.inv
@[simp]
theorem
category_theory.functor.inv_inv
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.is_equivalence F] :
def
category_theory.functor.fun_inv_id
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.is_equivalence F] :
Equations
def
category_theory.functor.inv_fun_id
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.is_equivalence F] :
Equations
@[instance]
def
category_theory.functor.is_equivalence_trans
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
{E : Type u₃}
[category_theory.category E]
(F : C ⥤ D)
(G : D ⥤ E)
[category_theory.is_equivalence F]
[category_theory.is_equivalence G] :
Equations
@[simp]
theorem
category_theory.equivalence.functor_inv
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(E : C ≌ D) :
@[simp]
theorem
category_theory.equivalence.inverse_inv
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(E : C ≌ D) :
@[simp]
theorem
category_theory.equivalence.functor_as_equivalence
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(E : C ≌ D) :
E.functor.as_equivalence = E
@[simp]
theorem
category_theory.equivalence.inverse_as_equivalence
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(E : C ≌ D) :
E.inverse.as_equivalence = E.symm
@[simp]
theorem
category_theory.is_equivalence.fun_inv_map
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.is_equivalence F]
(X Y : D)
(f : X ⟶ Y) :
@[simp]
theorem
category_theory.is_equivalence.inv_fun_map
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.is_equivalence F]
(X Y : C)
(f : X ⟶ Y) :
@[simp]
theorem
category_theory.is_equivalence.functor_unit_comp
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(E : C ⥤ D)
[category_theory.is_equivalence E]
(Y : C) :
@[simp]
theorem
category_theory.is_equivalence.inv_fun_id_inv_comp
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(E : C ⥤ D)
[category_theory.is_equivalence E]
(Y : C) :
@[class]
structure
category_theory.ess_surj
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D] :
C ⥤ D → Type (max u₁ u₂ v₂)
- obj_preimage : D → C
- iso' : Π (d : D), (F.obj (category_theory.ess_surj.obj_preimage F d) ≅ d) . "obviously"
def
category_theory.functor.obj_preimage
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.ess_surj F] :
D → C
Equations
def
category_theory.functor.fun_obj_preimage_iso
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.ess_surj F]
(d : D) :
F.obj (F.obj_preimage d) ≅ d
Equations
- F.fun_obj_preimage_iso d = category_theory.ess_surj.iso d
def
category_theory.equivalence.ess_surj_of_equivalence
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.is_equivalence F] :
Equations
- category_theory.equivalence.ess_surj_of_equivalence F = {obj_preimage := λ (Y : D), F.inv.obj Y, iso' := λ (Y : D), F.inv_fun_id.app Y}
@[instance]
def
category_theory.equivalence.faithful_of_equivalence
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.is_equivalence F] :
Equations
- _ = _
@[instance]
def
category_theory.equivalence.full_of_equivalence
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.is_equivalence F] :
def
category_theory.equivalence.equivalence_of_fully_faithfully_ess_surj
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(F : C ⥤ D)
[category_theory.full F]
[category_theory.faithful F]
[category_theory.ess_surj F] :
Equations
- category_theory.equivalence.equivalence_of_fully_faithfully_ess_surj F = category_theory.is_equivalence.mk (equivalence_inverse F) (category_theory.nat_iso.of_components (λ (X : C), (category_theory.preimage_iso (F.fun_obj_preimage_iso (F.obj X))).symm) _) (category_theory.nat_iso.of_components (λ (Y : D), F.fun_obj_preimage_iso Y) _)
@[simp]
theorem
category_theory.equivalence.functor_map_inj_iff
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
{X Y : C}
(f g : X ⟶ Y) :
@[simp]
theorem
category_theory.equivalence.inverse_map_inj_iff
{C : Type u₁}
[category_theory.category C]
{D : Type u₂}
[category_theory.category D]
(e : C ≌ D)
{X Y : D}
(f g : X ⟶ Y) :