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

Commit aa8980d

Browse files
b-mehtaOliver Nashleanprover-community-botpecherskyurkud
committed
chore(category_theory/monad): comonadic adjunction (#5770)
Defines comonadic adjunctions dual to what's already there Co-authored-by: Oliver Nash <oliver.nash@gmail.com> Co-authored-by: leanprover-community-bot <leanprover.community@gmail.com> Co-authored-by: Yakov Pechersky <ffxen158@gmail.com> Co-authored-by: Yury G. Kudryashov <urkud@urkud.name> Co-authored-by: tb65536 <tb65536@uw.edu> Co-authored-by: Floris van Doorn <fpvdoorn@gmail.com> Co-authored-by: Benjamin Davidson <68528197+benjamindavidson@users.noreply.github.com> Co-authored-by: Eric Wieser <wieser.eric@gmail.com> Co-authored-by: Mohamed Al-Fahim <malfahim8@gmail.com> Co-authored-by: hrmacbeth <25316162+hrmacbeth@users.noreply.github.com> Co-authored-by: Johan Commelin <johan@commelin.net> Co-authored-by: Mario Carneiro <di.gama@gmail.com>
1 parent 3b6d6d7 commit aa8980d

File tree

1 file changed

+47
-3
lines changed

1 file changed

+47
-3
lines changed

src/category_theory/monad/adjunction.lean

Lines changed: 47 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -12,7 +12,7 @@ open category
1212
universes v₁ v₂ u₁ u₂ -- declare the `v`'s first; see `category_theory.category` for an explanation
1313

1414
variables {C : Type u₁} [category.{v₁} C] {D : Type u₂} [category.{v₂} D]
15-
variables (R : D ⥤ C)
15+
variables (L : C ⥤ D) (R : D ⥤ C)
1616

1717
namespace adjunction
1818

@@ -23,6 +23,13 @@ instance monad (R : D ⥤ C) [is_right_adjoint R] : monad (left_adjoint R ⋙ R)
2323
assoc' := λ X, by { dsimp, rw [←R.map_comp], simp },
2424
right_unit' := λ X, by { dsimp, rw [←R.map_comp], simp } }
2525

26+
@[simps]
27+
instance comonad (L : C ⥤ D) [is_left_adjoint L] : comonad (right_adjoint L ⋙ L) :=
28+
{ ε := (of_left_adjoint L).counit,
29+
δ := whisker_right (whisker_left (right_adjoint L) (of_left_adjoint L).unit) L,
30+
coassoc' := λ X, by { dsimp, rw ← L.map_comp, simp },
31+
right_counit' := λ X, by { dsimp, rw ← L.map_comp, simp } }
32+
2633
end adjunction
2734

2835
namespace monad
@@ -54,16 +61,53 @@ def comparison_forget [is_right_adjoint R] : comparison R ⋙ forget (left_adjoi
5461

5562
end monad
5663

57-
/-- A right adjoint functor `R : D ⥤ C` is *monadic* if the comparison function `monad.comparison R` from `D` to the
58-
category of Eilenberg-Moore algebras for the adjunction is an equivalence. -/
64+
namespace comonad
65+
66+
/--
67+
Gven any adjunction `L ⊣ R`, there is a comparison functor `category_theory.comonad.comparison L`
68+
sending objects `X : C` to Eilenberg-Moore coalgebras for `L ⋙ R` with underlying object
69+
`L.obj X`.
70+
-/
71+
@[simps]
72+
def comparison [is_left_adjoint L] : C ⥤ coalgebra (right_adjoint L ⋙ L) :=
73+
{ obj := λ X,
74+
{ A := L.obj X,
75+
a := L.map ((adjunction.of_left_adjoint L).unit.app X),
76+
coassoc' := by { dsimp, rw [← L.map_comp, ← adjunction.unit_naturality, L.map_comp], refl } },
77+
map := λ X Y f,
78+
{ f := L.map f,
79+
h' := by { dsimp, rw ← L.map_comp, simp } } }
80+
81+
/--
82+
The underlying object of `(comonad.comparison L).obj X` is just `L.obj X`.
83+
-/
84+
@[simps]
85+
def comparison_forget [is_left_adjoint L] : comparison L ⋙ forget (right_adjoint L ⋙ L) ≅ L :=
86+
{ hom := { app := λ X, 𝟙 _, },
87+
inv := { app := λ X, 𝟙 _, } }
88+
89+
end comonad
90+
91+
/--
92+
A right adjoint functor `R : D ⥤ C` is *monadic* if the comparison functor `monad.comparison R`
93+
from `D` to the category of Eilenberg-Moore algebras for the adjunction is an equivalence.
94+
-/
5995
class monadic_right_adjoint (R : D ⥤ C) extends is_right_adjoint R :=
6096
(eqv : is_equivalence (monad.comparison R))
6197

98+
/--
99+
A left adjoint functor `L : C ⥤ D` is *comonadic* if the comparison functor `comonad.comparison L`
100+
from `C` to the category of Eilenberg-Moore algebras for the adjunction is an equivalence.
101+
-/
102+
class comonadic_left_adjoint (L : C ⥤ D) extends is_left_adjoint L :=
103+
(eqv : is_equivalence (comonad.comparison L))
104+
62105
-- TODO: This holds more generally for idempotent adjunctions, not just reflective adjunctions.
63106
instance μ_iso_of_reflective [reflective R] : is_iso (μ_ (left_adjoint R ⋙ R)) :=
64107
by { dsimp [adjunction.monad], apply_instance }
65108

66109
attribute [instance] monadic_right_adjoint.eqv
110+
attribute [instance] comonadic_left_adjoint.eqv
67111

68112
namespace reflective
69113

0 commit comments

Comments
 (0)