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

Commit 292fc04

Browse files
committed
feat(category_theory): adjunction convenience defs (#2754)
Transport adjunctions along natural isomorphisms, and `is_left_adjoint` or `is_right_adjoint` versions of existing adjunction properties.
1 parent 2ef444a commit 292fc04

File tree

1 file changed

+85
-1
lines changed

1 file changed

+85
-1
lines changed

src/category_theory/adjunction/basic.lean

Lines changed: 85 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,7 @@
11
/-
22
Copyright (c) 2019 Reid Barton. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
4-
Authors: Reid Barton, Johan Commelin
4+
Authors: Reid Barton, Johan Commelin, Bhavik Mehta
55
-/
66
import category_theory.equivalence
77
import data.equiv.basic
@@ -204,16 +204,87 @@ def id : 𝟭 C ⊣ 𝟭 C :=
204204
unit := 𝟙 _,
205205
counit := 𝟙 _ }
206206

207+
/-- If F and G are naturally isomorphic functors, establish an equivalence of hom-sets. -/
208+
def equiv_homset_left_of_nat_iso
209+
{F F' : C ⥤ D} (iso : F ≅ F') {X : C} {Y : D} :
210+
(F.obj X ⟶ Y) ≃ (F'.obj X ⟶ Y) :=
211+
{ to_fun := λ f, iso.inv.app _ ≫ f,
212+
inv_fun := λ g, iso.hom.app _ ≫ g,
213+
left_inv := λ f, by simp,
214+
right_inv := λ g, by simp }
215+
216+
@[simp]
217+
lemma equiv_homset_left_of_nat_iso_apply {F F' : C ⥤ D} (iso : F ≅ F') {X : C} {Y : D} (f : F.obj X ⟶ Y) :
218+
(equiv_homset_left_of_nat_iso iso) f = iso.inv.app _ ≫ f := rfl
219+
220+
@[simp]
221+
lemma equiv_homset_left_of_nat_iso_symm_apply {F F' : C ⥤ D} (iso : F ≅ F') {X : C} {Y : D} (g : F'.obj X ⟶ Y) :
222+
(equiv_homset_left_of_nat_iso iso).symm g = iso.hom.app _ ≫ g := rfl
223+
224+
/-- If G and H are naturally isomorphic functors, establish an equivalence of hom-sets. -/
225+
def equiv_homset_right_of_nat_iso
226+
{G G' : D ⥤ C} (iso : G ≅ G') {X : C} {Y : D} :
227+
(X ⟶ G.obj Y) ≃ (X ⟶ G'.obj Y) :=
228+
{ to_fun := λ f, f ≫ iso.hom.app _,
229+
inv_fun := λ g, g ≫ iso.inv.app _,
230+
left_inv := λ f, by simp,
231+
right_inv := λ g, by simp }
232+
233+
@[simp]
234+
lemma equiv_homset_right_of_nat_iso_apply {G G' : D ⥤ C} (iso : G ≅ G') {X : C} {Y : D} (f : X ⟶ G.obj Y) :
235+
(equiv_homset_right_of_nat_iso iso) f = f ≫ iso.hom.app _ := rfl
236+
237+
@[simp]
238+
lemma equiv_homset_right_of_nat_iso_symm_apply {G G' : D ⥤ C} (iso : G ≅ G') {X : C} {Y : D} (g : X ⟶ G'.obj Y) :
239+
(equiv_homset_right_of_nat_iso iso).symm g = g ≫ iso.inv.app _ := rfl
240+
241+
/-- Transport an adjunction along an natural isomorphism on the left. -/
242+
def of_nat_iso_left
243+
{F G : C ⥤ D} {H : D ⥤ C} (adj : F ⊣ H) (iso : F ≅ G) :
244+
G ⊣ H :=
245+
adjunction.mk_of_hom_equiv
246+
{ hom_equiv := λ X Y, (equiv_homset_left_of_nat_iso iso.symm).trans (adj.hom_equiv X Y) }
247+
248+
/-- Transport an adjunction along an natural isomorphism on the right. -/
249+
def of_nat_iso_right
250+
{F : C ⥤ D} {G H : D ⥤ C} (adj : F ⊣ G) (iso : G ≅ H) :
251+
F ⊣ H :=
252+
adjunction.mk_of_hom_equiv
253+
{ hom_equiv := λ X Y, (adj.hom_equiv X Y).trans (equiv_homset_right_of_nat_iso iso) }
254+
255+
/-- Transport being a right adjoint along a natural isomorphism. -/
256+
def right_adjoint_of_nat_iso {F G : C ⥤ D} (h : F ≅ G) [r : is_right_adjoint F] : is_right_adjoint G :=
257+
{ left := r.left,
258+
adj := of_nat_iso_right r.adj h }
259+
260+
/-- Transport being a left adjoint along a natural isomorphism. -/
261+
def left_adjoint_of_nat_iso {F G : C ⥤ D} (h : F ≅ G) [r : is_left_adjoint F] : is_left_adjoint G :=
262+
{ right := r.right,
263+
adj := of_nat_iso_left r.adj h }
264+
207265
section
208266
variables {E : Type u₃} [ℰ : category.{v₃} E] (H : D ⥤ E) (I : E ⥤ D)
209267

268+
/-- Show that adjunctions can be composed. -/
210269
def comp (adj₁ : F ⊣ G) (adj₂ : H ⊣ I) : F ⋙ H ⊣ I ⋙ G :=
211270
{ hom_equiv := λ X Z, equiv.trans (adj₂.hom_equiv _ _) (adj₁.hom_equiv _ _),
212271
unit := adj₁.unit ≫
213272
(whisker_left F $ whisker_right adj₂.unit G) ≫ (functor.associator _ _ _).inv,
214273
counit := (functor.associator _ _ _).hom ≫
215274
(whisker_left I $ whisker_right adj₁.counit H) ≫ adj₂.counit }
216275

276+
/-- If `F` and `G` are left adjoints then `F ⋙ G` is a left adjoint too. -/
277+
instance left_adjoint_of_comp {E : Type u₃} [ℰ : category.{v₃} E] (F : C ⥤ D) (G : D ⥤ E)
278+
[Fl : is_left_adjoint F] [Gl : is_left_adjoint G] : is_left_adjoint (F ⋙ G) :=
279+
{ right := Gl.right ⋙ Fl.right,
280+
adj := comp _ _ Fl.adj Gl.adj }
281+
282+
/-- If `F` and `G` are right adjoints then `F ⋙ G` is a right adjoint too. -/
283+
instance right_adjoint_of_comp {E : Type u₃} [ℰ : category.{v₃} E] {F : C ⥤ D} {G : D ⥤ E}
284+
[Fr : is_right_adjoint F] [Gr : is_right_adjoint G] : is_right_adjoint (F ⋙ G) :=
285+
{ left := Gr.left ⋙ Fr.left,
286+
adj := comp _ _ Gr.adj Fr.adj }
287+
217288
end
218289

219290
section construct_left
@@ -297,9 +368,22 @@ end equivalence
297368

298369
namespace functor
299370

371+
/-- An equivalence `E` is left adjoint to its inverse. -/
300372
def adjunction (E : C ⥤ D) [is_equivalence E] : E ⊣ E.inv :=
301373
(E.as_equivalence).to_adjunction
302374

375+
/-- If `F` is an equivalence, it's a left adjoint. -/
376+
@[priority 10]
377+
instance left_adjoint_of_equivalence {F : C ⥤ D} [is_equivalence F] : is_left_adjoint F :=
378+
{ right := _,
379+
adj := functor.adjunction F }
380+
381+
/-- If `F` is an equivalence, it's a right adjoint. -/
382+
@[priority 10]
383+
instance right_adjoint_of_equivalence {F : C ⥤ D} [is_equivalence F] : is_right_adjoint F :=
384+
{ left := _,
385+
adj := functor.adjunction F.inv }
386+
303387
end functor
304388

305389
end category_theory

0 commit comments

Comments
 (0)