Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 616cb5e

Browse files
committed
chore(category_theory/equivalence) explicit transitivity transformation (#3176)
Modifies the construction of the transitive equivalence to be explicit in what exactly the natural transformations are. The motivation for this is two-fold: firstly we get auto-generated projection lemmas for extracting the functor and inverse, and the natural transformations aren't obscured through `adjointify_η`.
1 parent abae5a3 commit 616cb5e

File tree

1 file changed

+22
-8
lines changed

1 file changed

+22
-8
lines changed

src/category_theory/equivalence.lean

Lines changed: 22 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -158,14 +158,28 @@ protected definition mk (F : C ⥤ D) (G : D ⥤ C)
158158

159159
variables {E : Type u₃} [category.{v₃} E]
160160

161-
@[trans] def trans (e : C ≌ D) (f : D ≌ E) : C ≌ E :=
162-
begin
163-
apply equivalence.mk (e.functor ⋙ f.functor) (f.inverse ⋙ e.inverse),
164-
{ refine iso.trans e.unit_iso _,
165-
exact iso_whisker_left e.functor (iso_whisker_right f.unit_iso e.inverse) },
166-
{ refine iso.trans _ f.counit_iso,
167-
exact iso_whisker_left f.inverse (iso_whisker_right e.counit_iso f.functor) }
168-
end
161+
@[trans, simps] def trans (e : C ≌ D) (f : D ≌ E) : C ≌ E :=
162+
{ functor := e.functor ⋙ f.functor,
163+
inverse := f.inverse ⋙ e.inverse,
164+
unit_iso :=
165+
begin
166+
refine iso.trans e.unit_iso _,
167+
exact iso_whisker_left e.functor (iso_whisker_right f.unit_iso e.inverse) ,
168+
end,
169+
counit_iso :=
170+
begin
171+
refine iso.trans _ f.counit_iso,
172+
exact iso_whisker_left f.inverse (iso_whisker_right e.counit_iso f.functor)
173+
end,
174+
-- We wouldn't have need to give this proof if we'd used `equivalence.mk`,
175+
-- but we choose to avoid using that here, for the sake of good structure projection `simp` lemmas.
176+
functor_unit_iso_comp' := λ X,
177+
begin
178+
dsimp,
179+
rw [← f.functor.map_comp_assoc, e.functor.map_comp, functor_unit, fun_inv_map,
180+
inv_hom_id_app_assoc, assoc, inv_hom_id_app, counit_functor, ← functor.map_comp],
181+
erw [comp_id, hom_inv_id_app, functor.map_id],
182+
end }
169183

170184
def fun_inv_id_assoc (e : C ≌ D) (F : C ⥤ E) : e.functor ⋙ e.inverse ⋙ F ≅ F :=
171185
(functor.associator _ _ _).symm ≪≫ iso_whisker_right e.unit_iso.symm F ≪≫ F.left_unitor

0 commit comments

Comments
 (0)