bsc-leon-vatthauer/src/Monad/Instance/Delay/Properties.lagda.md

127 lines
5.9 KiB
Markdown
Raw Normal View History

<!--
```agda
open import Level
open import Categories.Functor
open import Category.Instance.AmbientCategory using (Ambient)
open import Categories.Monad
open import Categories.Monad.Construction.Kleisli
open import Categories.Monad.Relative renaming (Monad to RMonad)
open import Data.Product using (_,_; Σ; Σ-syntax)
open import Categories.Functor.Algebra
open import Categories.Functor.Coalgebra
open import Categories.Object.Terminal
open import Categories.NaturalTransformation.Core
open import Misc.FreeObject
```
-->
```agda
module Monad.Instance.Delay.Properties {o e} (ambient : Ambient o e) where
open Ambient ambient
open import Categories.Diagram.Coequalizer C
open import Monad.Instance.Delay ambient
open import Algebra.Search ambient
open import Algebra.ElgotAlgebra ambient
open import Algebra.Properties ambient
open F-Coalgebra-Morphism using () renaming (f to u; commutes to coalg-commutes)
open import Monad.Instance.Delay.Quotienting ambient
```
# Properties of the quotiented delay monad
```agda
module _ (D : DelayM) where
open DelayM D renaming (functor to D-Functor; monad to D-Monad; kleisli to D-Kleisli)
open Functor D-Functor using () renaming (F₁ to D₁; homomorphism to D-homomorphism; F-resp-≈ to D-resp-≈; identity to D-identity)
open RMonad D-Kleisli using (extend; extend-≈) renaming (assoc to k-assoc; identityʳ to k-identityʳ; identityˡ to k-identityˡ)
open Monad D-Monad using () renaming (assoc to M-assoc; identityʳ to M-identityʳ)
open HomReasoning
open M C
open MR C
open Equiv
module _ (coeqs : ∀ X → Coequalizer (extend (ι {X})) (D₁ π₁)) where
open Quotiented D coeqs
cond-1 : Set (o ⊔ ⊔ e)
cond-1 = ∀ X → preserves D-Functor (coeqs X)
-- cond-2' : Set (o ⊔ ⊔ e)
-- cond-2' = ∀ X → Σ[ s-alg-on ∈ Search-Algebra-on D (Ď₀ X) ] is-F-Algebra-Morphism {F = D-Functor} (record { A = D₀ X ; α = μ.η X }) (record { A = Ď₀ X ; α = Search-Algebra-on.α s-alg-on }) (ρ {X})
record cond-2' (X : Obj) : Set (o ⊔ ⊔ e) where
field
s-alg-on : Search-Algebra-on D (Ď₀ X)
ρ-algebra-morphism : is-F-Algebra-Morphism {F = D-Functor} (record { A = D₀ X ; α = μ.η X }) (record { A = Ď₀ X ; α = Search-Algebra-on.α s-alg-on }) (ρ {X})
cond-2 = ∀ X → cond-2' X
record cond-3' (X : Obj) : Set (suc o ⊔ suc ⊔ suc e) where
-- Ď₀ X is stable free elgot algebra
field
elgot : Elgot-Algebra-on (Ď₀ X)
elgot-alg = record { A = Ď₀ X ; algebra = elgot }
open Elgot-Algebra-on elgot
field
isFO : IsFreeObject elgotForgetfulF X elgot-alg
freeobject = IsFreeObject⇒FreeObject elgotForgetfulF X elgot-alg isFO
field
isStable : IsStableFreeElgotAlgebra freeobject
-- ρ is D-algebra morphism
field
ρ-algebra-morphism : is-F-Algebra-Morphism {F = D-Functor} (record { A = D₀ X ; α = μ.η X }) (record { A = Ď₀ X ; α = out # }) (ρ {X})
ρ-law : ρ ≈ ((ρ ∘ now +₁ idC) ∘ out) #
cond-3 = ∀ X → cond-3' X
record cond-4 : Set (o ⊔ ⊔ e) where
2⇒3 : cond-2 → cond-3
2⇒3 c-2 X = record
{ elgot = {! !}
; isFO = {! !}
; isStable = {! !}
; ρ-algebra-morphism = {! !}
; ρ-law = {! !}
}
1⇒2 : cond-1 → cond-2
1⇒2 c-1 X = record { s-alg-on = s-alg-on ; ρ-algebra-morphism = begin ρ ∘ μ.η X ≈⟨ D-universal ⟩ Search-Algebra-on.α s-alg-on ∘ D₁ ρ ∎ }
where
open Coequalizer (coeqs X) renaming (universal to coeq-universal)
open IsCoequalizer (c-1 X) using () renaming (equality to D-equality; coequalize to D-coequalize; universal to D-universal; unique to D-unique)
s-alg-on : Search-Algebra-on D (Ď₀ X)
s-alg-on = record
{ α = α'
; identity₁ = ρ-epi (α' ∘ now) idC (begin
(α' ∘ now) ∘ ρ ≈⟨ pullʳ (η.commute ρ) ⟩
α' ∘ D₁ ρ ∘ now ≈⟨ pullˡ (sym D-universal) ⟩
(ρ ∘ μ.η X) ∘ now ≈⟨ cancelʳ M-identityʳ ⟩
ρ ≈⟨ sym identityˡ ⟩
idC ∘ ρ ∎)
; identity₂ = IsCoequalizer⇒Epi (c-1 X) (α' ∘ ▷) α' (begin
(α' ∘ ▷) ∘ D₁ ρ ≈⟨ pullʳ (▷∘extend-comm (now ∘ ρ)) ⟩
α' ∘ D₁ ρ ∘ ▷ ≈⟨ pullˡ (sym D-universal) ⟩
(ρ ∘ μ.η X) ∘ ▷ ≈⟨ pullʳ (sym (▷∘extend-comm idC)) ⟩
ρ ∘ ▷ ∘ extend idC ≈⟨ pullˡ ρ▷ ⟩
ρ ∘ extend idC ≈⟨ D-universal ⟩
α' ∘ D₁ ρ ∎)
}
where
μ∘Dι : μ.η X ∘ D₁ ι ≈ extend ι
μ∘Dι = sym k-assoc ○ extend-≈ (cancelˡ k-identityʳ)
eq₁ : D₁ (extend ι) ≈ D₁ (μ.η X) ∘ D₁ (D₁ ι)
eq₁ = sym (begin
D₁ (μ.η X) ∘ D₁ (D₁ ι) ≈⟨ sym D-homomorphism ⟩
D₁ (μ.η X ∘ D₁ ι) ≈⟨ D-resp-≈ μ∘Dι
D₁ (extend ι) ∎)
α' = D-coequalize {h = ρ {X} ∘ μ.η X} (begin
(ρ ∘ μ.η X) ∘ D₁ (extend ι) ≈⟨ (refl⟩∘⟨ eq₁) ⟩
(ρ ∘ μ.η X) ∘ D₁ (μ.η X) ∘ D₁ (D₁ ι) ≈⟨ pullʳ (pullˡ M-assoc) ⟩
ρ ∘ (μ.η X ∘ μ.η (D₀ X)) ∘ D₁ (D₁ ι) ≈⟨ (refl⟩∘⟨ pullʳ (μ.commute ι)) ⟩
ρ ∘ μ.η X ∘ (D₁ ι) ∘ μ.η (X × N) ≈⟨ (refl⟩∘⟨ pullˡ μ∘Dι) ⟩
ρ ∘ extend ι ∘ μ.η (X × N) ≈⟨ pullˡ equality ⟩
(ρ ∘ D₁ π₁) ∘ μ.η (X × N) ≈⟨ (pullʳ (sym (μ.commute π₁)) ○ sym-assoc) ⟩
(ρ ∘ μ.η X) ∘ D₁ (D₁ π₁) ∎)
```