module Cat.Diagram.Comonad where
open Functor
open _=>_
module _ {o ℓ} {C : Precategory o ℓ} where
open Precategory C
Comonads🔗
A comonad on a category is dual to a monad on instead of a unit and multiplication we have a counit and comultiplication More generally, we can define what it means to equip a fixed functor with the structure of a comonad.
record Comonad-on (W : Functor C C) : Type (o ⊔ ℓ) where
field
: W => Id
counit : W => (W F∘ W) comult
module counit = _=>_ counit renaming (η to ε)
module comult = _=>_ comult renaming (η to δ)
open Functor W renaming (F₀ to W₀ ; F₁ to W₁ ; F-id to W-id ; F-∘ to W-∘) public
open counit using (ε) public
open comult using (δ) public
We also have equations governing the counit and comultiplication. Unsurprisingly, these are dual to the equations of a monad.
field
: ∀ {x} → W₁ (ε x) ∘ δ x ≡ id
δ-unitl : ∀ {x} → ε (W₀ x) ∘ δ x ≡ id
δ-unitr : ∀ {x} → W₁ (δ x) ∘ δ x ≡ δ (W₀ x) ∘ δ x δ-assoc
module _ {o h : _} {C : Precategory o h} {F G : Functor C C} {M : Comonad-on F} {N : Comonad-on G} where
private
module C = Cat.Reasoning C
module M = Comonad-on M
module N = Comonad-on N
unquoteDecl eqv = declare-record-iso eqv (quote Comonad-on)
Comonad-on-path: (p0 : F ≡ G)
→ (∀ x → PathP (λ i → C.Hom (p0 i · x) x) (M.ε x) (N.ε x))
→ (∀ x → PathP (λ i → C.Hom (p0 i · x) (p0 i · (p0 i · x))) (M.δ x) (N.δ x))
→ PathP (λ i → Comonad-on (p0 i)) M N
= injectiveP (λ _ → eqv) $
Comonad-on-path M=N pcounit pcomult (ap₂ _F∘_ M=N M=N) pcomult ,ₚ prop! Nat-pathp M=N refl pcounit ,ₚ Nat-pathp M=N