module Cat.Instances.Coalgebras.Colimits {o ℓ} {C : Precategory o ℓ} {F : Functor C C} (W : Comonad-on F) where
private
module CoEM = Cat.Reasoning (Coalgebras W)
module C = Cat.Reasoning C
module W = Comonad-on W
open Coalgebra-on
open ∫Hom
Colimits in categories of coalgebras🔗
This module dualises the construction of limits in categories of algebras for a monad to colimits in categories of coalgebras for a comonad; see there for explanations.
module _ {jo jℓ} {J : Precategory jo jℓ} (F : Functor J (Coalgebras W)) where
private
module J = Precategory J
module F = Functor F
module FAlg j = Coalgebra-on (F.₀ j .snd)
= πᶠ (Coalgebras-over W)
Forget-CoEM open Functor
open _=>_
: preserves-colimit (πᶠ (Coalgebras-over W)) F
Forget-CoEM-preserves-colimits = left-adjoint-is-cocontinuous (Forget⊣Cofree W)
Forget-CoEM-preserves-colimits
make-coalgebra-colimit: ∀ {K : Functor ⊤Cat C} {eta : Forget-CoEM F∘ F => K F∘ !F}
→ (colim : is-lan !F (Forget-CoEM F∘ F) K eta)
→ (rho : Coalgebra-on W (K .F₀ tt))
→ (∀ j → W.W₁ (is-colimit.ψ colim j) C.∘ FAlg.ρ j ≡ rho .ρ C.∘ is-colimit.ψ colim j)
→ make-is-colimit F (K .F₀ tt , rho)
= coem-colim where
make-coalgebra-colimit colim apex-coalg comm module colim = is-colimit colim
open make-is-colimit
module apex = Coalgebra-on apex-coalg
open _=>_
: make-is-colimit F _
coem-colim .ψ j .fst = colim.ψ j
coem-colim .ψ j .snd = comm j
coem-colim .commutes f = ext (colim.commutes f)
coem-colim .universal eta p .fst =
coem-colim .universal (λ j → eta j .fst) (λ f i → p f i .fst)
colim.factors eta p =
coem-colim (colim.factors _ _)
ext .unique eta p other q =
coem-colim (colim.unique _ _ _ λ j i → q j i .fst)
ext .universal eta p .snd = colim.unique₂ (λ j → W.W₁ (eta j .fst) C.∘ F.F₀ j .snd .ρ)
coem-colim (λ f → C.pullr (sym (F.F₁ f .snd))
.pulll (sym (W.W-∘ _ _) ∙ ap W.W₁ (ap fst (p f))))
∙ C(λ j → C.pullr (sym (comm j))
.pulll (sym (W.W-∘ _ _) ∙ ap W.W₁ (colim.factors _ _)))
∙ C(λ j → C.pullr (colim.factors _ _)
(eta j .snd))
∙ sym
: Colimit (Forget-CoEM F∘ F) → Colimit F
Forget-lift-colimit = to-colimit $ to-is-colimit $ make-coalgebra-colimit
Forget-lift-colimit colim-over (Colimit.has-colimit colim-over) coapex-coalgebra (λ j → sym (colim-over.factors _ _))
where
module colim-over = Colimit colim-over
module colim = Lan colim-over
: Coalgebra-on W colim-over.coapex
coapex-coalgebra .ρ =
coapex-coalgebra .universal (λ j → W.W₁ (colim-over.ψ j) C.∘ FAlg.ρ j) comm where abstract
colim-over: ∀ {x y} (f : J.Hom x y)
comm → (W.W₁ (colim-over.ψ y) C.∘ FAlg.ρ y) C.∘ F.₁ f .fst
.W₁ (colim-over.ψ x) C.∘ FAlg.ρ x
≡ W{x} {y} f =
comm (W.W₁ (colim-over.ψ y) C.∘ FAlg.ρ y) C.∘ F.₁ f .fst ≡⟨ C.pullr (sym (F.₁ f .snd)) ⟩
.W₁ (colim-over.ψ y) C.∘ W.W₁ (F.₁ f .fst) C.∘ FAlg.ρ x ≡⟨ C.pulll (sym (W.W-∘ _ _)) ⟩
W.W₁ (colim-over.ψ y C.∘ F.₁ f .fst) C.∘ FAlg.ρ x ≡⟨ ap W.W₁ (colim-over.commutes f) C.⟩∘⟨refl ⟩
W.W₁ (colim-over.ψ x) C.∘ FAlg.ρ x ∎
W.ρ-counit = colim-over.unique₂ _ colim-over.commutes
coapex-coalgebra (λ j → C.pullr (colim-over.factors _ _)
.pulll (W.counit.is-natural _ _ _)
∙∙ C.cancelr (FAlg.ρ-counit j))
∙∙ C(λ j → C.idl _)
.ρ-comult = colim-over.unique₂ _
coapex-coalgebra (λ f → C.pullr $ C.pullr (sym (F.₁ f .snd))
.pulll (sym (W.W-∘ _ _) ∙ ap W.W₁ (colim-over.commutes f)))
∙ C(λ j → C.pullr (colim-over.factors _ _)
(C.pulll (sym (W.W-∘ _ _) ∙ ap W.W₁ (colim-over.factors _ _) ∙ W.W-∘ _ _)
∙ sym .extendr (sym (FAlg.ρ-comult j))
∙∙ C(C._∘ FAlg.ρ j) (sym (W.comult.is-natural _ _ _))
∙∙ ap (C.assoc _ _ _)))
∙ sym (λ j → C.pullr (colim-over.factors _ _))
module _ {jo jℓ} {J : Precategory jo jℓ} where
open lifts-colimit
open creates-colimit
: lifts-colimits-of J (πᶠ (Coalgebras-over W))
Forget-CoEM-lifts-colimits .lifted = Forget-lift-colimit _ colim
Forget-CoEM-lifts-colimits colim .preserved =
Forget-CoEM-lifts-colimits colim _ (Lan.has-lan (Forget-lift-colimit _ colim))
Forget-CoEM-preserves-colimits
: creates-colimits-of J (πᶠ (Coalgebras-over W))
Forget-CoEM-creates-colimits = conservative+lifts→creates-colimits
Forget-CoEM-creates-colimits Forget-CoEM-is-conservative Forget-CoEM-lifts-colimits
Eilenberg-Moore-is-cocomplete: ∀ {a b} → is-cocomplete a b C → is-cocomplete a b (Coalgebras W)
= lifts-colimits→cocomplete
Eilenberg-Moore-is-cocomplete (πᶠ (Coalgebras-over W)) Forget-CoEM-lifts-colimits