mathlib documentation

category_theory.​equivalence

category_theory.​equivalence

structure category_theory.​equivalence (C : Type u₁) [category_theory.category C] (D : Type u₂) [category_theory.category D] :
Type (max u₁ u₂ v₁ v₂)

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.

@[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.​functor_unit_comp {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] (e : C D) (X : C) :

@[simp]

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.​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) :

@[simp]
theorem category_theory.​equivalence.​inv_fun_map {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] (e : C D) (X Y : C) (f : X Y) :

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) :
(𝟭 C F G)(G F 𝟭 D)(C D)

Equations
@[simp]

@[simp]

@[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_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) :

@[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) :
(e.fun_inv_id_assoc F).inv.app X = F.map (e.unit.app X)

@[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) :

@[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) :
f e.unit.app Y = f' e.unit.app Y f = f'

@[simp]
theorem category_theory.​equivalence.​cancel_unit_inv_right {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] (e : C D) {X Y : C} (f f' : X e.inverse.obj (e.functor.obj Y)) :
f e.unit_inv.app Y = f' e.unit_inv.app Y f = f'

@[simp]
theorem category_theory.​equivalence.​cancel_counit_right {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] (e : C D) {X Y : D} (f f' : X e.functor.obj (e.inverse.obj Y)) :
f e.counit.app Y = f' e.counit.app Y f = f'

@[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]
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 : C} (f : W X) (g : X Y) (f' : W X') (g' : X' Y) :
f g e.unit.app Y = f' g' e.unit.app Y f g = f' g'

@[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) :
f g e.counit_inv.app Y = f' g' e.counit_inv.app Y f g = f' g'

@[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) :
f g h e.unit.app Z = f' g' h' e.unit.app Z f g h = f' g' h'

@[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) :
f g h e.counit_inv.app Z = f' g' h' e.counit_inv.app Z f g h = f' g' h'

def category_theory.​equivalence.​pow {C : Type u₁} [category_theory.category C] :
(C C)(C C)

Powers of an auto-equivalence.

Equations
@[simp]
theorem category_theory.​equivalence.​pow_one {C : Type u₁} [category_theory.category C] (e : C C) :
e ^ 1 = e

@[simp]
theorem category_theory.​equivalence.​pow_minus_one {C : Type u₁} [category_theory.category C] (e : C C) :
e ^ -1 = e.symm

@[simp]

@[simp]

@[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) :

@[class]
structure category_theory.​ess_surj {C : Type u₁} [category_theory.category C] {D : Type u₂} [category_theory.category D] :
C DType (max u₁ u₂ v₂)

Instances

Equations
@[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) :
e.functor.map f = e.functor.map g f = g

@[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) :
e.inverse.map f = e.inverse.map g f = g