mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
Worked on defining delay monad correctly
This commit is contained in:
parent
22448c3bab
commit
0a2cc10ec8
1 changed files with 129 additions and 6 deletions
|
@ -1,20 +1,24 @@
|
|||
<!--
|
||||
```agda
|
||||
open import Level
|
||||
open import Categories.Category.Core
|
||||
open import Categories.Category
|
||||
open import Categories.Monad
|
||||
open import Categories.Category.Distributive
|
||||
open import Categories.Category.Extensive.Bundle
|
||||
open import Categories.Category.Extensive
|
||||
open import Categories.Category.BinaryProducts
|
||||
open import Categories.Category.Cocartesian
|
||||
open import Categories.Category.Cartesian
|
||||
open import Categories.Category.Cartesian
|
||||
open import Categories.Object.Terminal
|
||||
open import Categories.Object.Coproduct
|
||||
open import Categories.Category.Construction.F-Coalgebras
|
||||
open import Categories.Functor.Coalgebra
|
||||
open import Categories.Functor
|
||||
open import Categories.Functor.Algebra
|
||||
open import Categories.Monad.Construction.Kleisli
|
||||
open import Categories.Category.Construction.F-Coalgebras
|
||||
open import Categories.NaturalTransformation
|
||||
import Categories.Morphism as M
|
||||
import Categories.Morphism.Reasoning as MR
|
||||
```
|
||||
|
@ -50,6 +54,129 @@ First I postulate the Functor *D*, maybe I should derive it...
|
|||
- how to express **Theorem 8** in agda?
|
||||
|
||||
```agda
|
||||
|
||||
-- TODO use this to get i₂ as F-Coalgebra, then use ⊤-id below!!
|
||||
Coalg-cop : ∀ {A B : Obj} → (F : Endofunctor C) → (alg₁ : F-Coalgebra-on F A) → (alg₂ : F-Coalgebra-on F B) → Coproduct (F-Coalgebras F) (record { A = A ; α = alg₁ }) (record { A = B ; α = alg₂ })
|
||||
Coalg-cop {A} {B} F alg₁ alg₂ = record
|
||||
{ A+B = record { A = A + B ; α = [ (F₁ i₁ ∘ alg₁) , F₁ i₂ ∘ alg₂ ] }
|
||||
; i₁ = record { f = i₁ ; commutes = inject₁ }
|
||||
; i₂ = record { f = i₂ ; commutes = inject₂ }
|
||||
; [_,_] = λ {CA} h i → record { f = [ F-Coalgebra-Morphism.f h , F-Coalgebra-Morphism.f i ] ; commutes = begin
|
||||
F-Coalgebra.α CA ∘ [ F-Coalgebra-Morphism.f h , F-Coalgebra-Morphism.f i ] ≈⟨ ∘[] ⟩
|
||||
[ F-Coalgebra.α CA ∘ F-Coalgebra-Morphism.f h , F-Coalgebra.α CA ∘ F-Coalgebra-Morphism.f i ] ≈⟨ ⟺ ([]-cong₂ (⟺ (F-Coalgebra-Morphism.commutes h)) ((⟺ (F-Coalgebra-Morphism.commutes i)))) ⟩
|
||||
[ F₁ (F-Coalgebra-Morphism.f h) ∘ alg₁ , F₁ (F-Coalgebra-Morphism.f i) ∘ alg₂ ] ≈⟨ ⟺ ([]-cong₂ ((F-resp-≈ inject₁) ⟩∘⟨refl) ((F-resp-≈ inject₂) ⟩∘⟨refl)) ⟩
|
||||
[ F₁ ([ F-Coalgebra-Morphism.f h , F-Coalgebra-Morphism.f i ] ∘ i₁) ∘ alg₁ , F₁ ([ F-Coalgebra-Morphism.f h , F-Coalgebra-Morphism.f i ] ∘ i₂) ∘ alg₂ ] ≈⟨ ⟺ ([]-cong₂ (pullˡ (⟺ homomorphism)) (pullˡ (⟺ homomorphism))) ⟩
|
||||
[ F₁ ([ F-Coalgebra-Morphism.f h , F-Coalgebra-Morphism.f i ]) ∘ F₁ i₁ ∘ alg₁ , F₁ ([ F-Coalgebra-Morphism.f h , F-Coalgebra-Morphism.f i ]) ∘ F₁ i₂ ∘ alg₂ ] ≈⟨ ⟺ ∘[] ⟩
|
||||
F₁ ([ F-Coalgebra-Morphism.f h , F-Coalgebra-Morphism.f i ]) ∘ [ F₁ i₁ ∘ alg₁ , F₁ i₂ ∘ alg₂ ] ∎ }
|
||||
; inject₁ = inject₁
|
||||
; inject₂ = inject₂
|
||||
; unique = λ eq₁ eq₂ → +-unique eq₁ eq₂
|
||||
}
|
||||
where open Functor F
|
||||
|
||||
delayF : Obj → Endofunctor C
|
||||
delayF Y = record
|
||||
{ F₀ = λ X → Y + X
|
||||
; F₁ = λ f → idC +₁ f
|
||||
; identity = CC.coproduct.unique id-comm-sym id-comm-sym
|
||||
; homomorphism = ⟺ (+₁∘+₁ ○ +₁-cong₂ identity² refl)
|
||||
; F-resp-≈ = λ eq → +₁-cong₂ refl eq
|
||||
}
|
||||
|
||||
record DelayM : Set (o ⊔ ℓ ⊔ e) where
|
||||
field
|
||||
algebras : ∀ (A : Obj) → Terminal (F-Coalgebras (delayF A))
|
||||
|
||||
module D A = Functor (delayF A)
|
||||
|
||||
module _ (X : Obj) where
|
||||
open Terminal (algebras X) using (⊤; !)
|
||||
open F-Coalgebra ⊤ using (α) renaming (A to DX)
|
||||
|
||||
-- TODO figure out how to name things...
|
||||
out' : DX ⇒ X + DX
|
||||
out' = α
|
||||
|
||||
D₀ = DX
|
||||
|
||||
out-≅ : DX ≅ X + DX
|
||||
out-≅ = colambek {F = delayF X} (algebras X)
|
||||
|
||||
open _≅_ out-≅ using () renaming (to to out⁻¹; from to out) public
|
||||
|
||||
now : X ⇒ DX
|
||||
now = out⁻¹ ∘ i₁
|
||||
|
||||
later : DX ⇒ DX
|
||||
later = out⁻¹ ∘ i₂
|
||||
|
||||
module _ {Y : Obj} where
|
||||
coit : Y ⇒ X + Y → Y ⇒ DX
|
||||
coit f = F-Coalgebra-Morphism.f (! {A = record { A = Y ; α = f }})
|
||||
|
||||
coit-commutes : ∀ (f : Y ⇒ X + Y) → out ∘ (coit f) ≈ (idC +₁ coit f) ∘ f
|
||||
coit-commutes f = F-Coalgebra-Morphism.commutes (! {A = record { A = Y ; α = f }})
|
||||
|
||||
monad : Monad C
|
||||
monad = Kleisli⇒Monad C (record
|
||||
{ F₀ = D₀
|
||||
; unit = λ {X} → now X
|
||||
; extend = extend
|
||||
; identityʳ = {! !}
|
||||
; identityˡ = {! !}
|
||||
; assoc = {! !}
|
||||
; sym-assoc = {! !}
|
||||
; extend-≈ = {! !}
|
||||
})
|
||||
where
|
||||
module _ {X Y : Obj} (f : X ⇒ D₀ Y) where
|
||||
open Terminal (algebras Y) using (!; ⊤-id)
|
||||
alg : F-Coalgebra (delayF Y)
|
||||
alg = record { A = D₀ X + D₀ Y ; α = (idC +₁ idC +₁ [ idC , idC ]) ∘ (idC +₁ _≅_.to +-assoc ∘ _≅_.to +-comm) ∘ _≅_.to +-assoc ∘ ((out Y ∘ f) +₁ idC) ∘ _≅_.to +-assoc ∘ (out X +₁ idC) } -- _≅_.to +-assoc
|
||||
extend : D₀ X ⇒ D₀ Y
|
||||
extend = F-Coalgebra-Morphism.f (! {A = alg}) ∘ i₁ {B = D₀ Y}
|
||||
+-assocˡ∘i₁∘i₁ : +-assocˡ ∘ i₁ ∘ i₁ ≈ i₁
|
||||
+-assocˡ∘i₁∘i₁ = begin +-assocˡ ∘ i₁ ∘ i₁ ≈⟨ pullˡ inject₁ ⟩ [ i₁ , i₂ ∘ i₁ ] ∘ i₁ ≈⟨ inject₁ ⟩ i₁ ∎
|
||||
-- i₂-morphism = record { f = i₂ ; commutes = {! !} }
|
||||
i₂∘! : F-Coalgebra-Morphism.f (! {A = alg}) ∘ i₂ ≈ idC
|
||||
i₂∘! = ⊤-id (F-Coalgebras (delayF Y) [ ! {A = alg} ∘ record { f = i₂ ; commutes = {! inject₂ !} } ])
|
||||
identityʳ' : extend ∘ now X ≈ f
|
||||
identityʳ' = begin
|
||||
(F-Coalgebra-Morphism.f (! {A = alg}) ∘ i₁ {B = D₀ Y}) ∘ (out⁻¹ X ∘ i₁) ≈⟨ introˡ (_≅_.isoˡ (out-≅ Y)) ⟩
|
||||
(out⁻¹ Y ∘ out Y) ∘ (F-Coalgebra-Morphism.f (! {A = alg}) ∘ i₁ {B = D₀ Y}) ∘ (out⁻¹ X ∘ i₁) ≈⟨ pullʳ (pullˡ (pullˡ (F-Coalgebra-Morphism.commutes (! {A = alg})))) ⟩
|
||||
out⁻¹ Y ∘ (((idC +₁ (F-Coalgebra-Morphism.f !)) ∘ F-Coalgebra.α alg) ∘ i₁) ∘ out⁻¹ X ∘ i₁ ≈⟨ refl⟩∘⟨ (pullʳ (pullʳ (pullʳ (pullʳ (pullʳ (pullʳ +₁∘i₁)))))) ⟩∘⟨refl ⟩
|
||||
out⁻¹ Y ∘ ((idC +₁ F-Coalgebra-Morphism.f !) ∘ (idC +₁ idC +₁ [ idC , idC ]) ∘ (idC +₁ M._≅_.to +-assoc ∘ M._≅_.to +-comm) ∘ M._≅_.to +-assoc ∘ (out Y ∘ f +₁ idC) ∘ M._≅_.to +-assoc ∘ (i₁ ∘ out X)) ∘ out⁻¹ X ∘ i₁ ≈⟨ refl⟩∘⟨ (pullʳ (pullʳ (pullʳ (pullʳ (pullʳ (pullʳ (pullʳ (pullˡ (M._≅_.isoʳ (out-≅ X)))))))))) ⟩
|
||||
out⁻¹ Y ∘ (idC +₁ F-Coalgebra-Morphism.f !) ∘ (idC +₁ idC +₁ [ idC , idC ]) ∘ (idC +₁ M._≅_.to +-assoc ∘ M._≅_.to +-comm) ∘ M._≅_.to +-assoc ∘ (out Y ∘ f +₁ idC) ∘ M._≅_.to +-assoc ∘ i₁ ∘ idC ∘ i₁ ≈⟨ refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ identityˡ))))))) ⟩
|
||||
out⁻¹ Y ∘ (idC +₁ F-Coalgebra-Morphism.f !) ∘ (idC +₁ idC +₁ [ idC , idC ]) ∘ (idC +₁ +-assocˡ ∘ M._≅_.to +-comm) ∘ M._≅_.to +-assoc ∘ (out Y ∘ f +₁ idC) ∘ M._≅_.to +-assoc ∘ i₁ ∘ i₁ ≈⟨ refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ (+-assocˡ∘i₁∘i₁)))))) ⟩
|
||||
out⁻¹ Y ∘ (idC +₁ F-Coalgebra-Morphism.f !) ∘ (idC +₁ idC +₁ [ idC , idC ]) ∘ (idC +₁ +-assocˡ ∘ M._≅_.to +-comm) ∘ M._≅_.to +-assoc ∘ (out Y ∘ f +₁ idC) ∘ i₁ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ +₁∘i₁ ⟩
|
||||
out⁻¹ Y ∘ (idC +₁ F-Coalgebra-Morphism.f !) ∘ (idC +₁ idC +₁ [ idC , idC ]) ∘ (idC +₁ +-assocˡ ∘ M._≅_.to +-comm) ∘ M._≅_.to +-assoc ∘ (i₁ ∘ (out Y ∘ f)) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ pullˡ inject₁ ⟩
|
||||
out⁻¹ Y ∘ (idC +₁ F-Coalgebra-Morphism.f !) ∘ (idC +₁ idC +₁ [ idC , idC ]) ∘ (idC +₁ +-assocˡ ∘ M._≅_.to +-comm) ∘ [ i₁ , i₂ ∘ i₁ ] ∘ (out Y ∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ pullˡ ∘[] ⟩
|
||||
out⁻¹ Y ∘ (idC +₁ F-Coalgebra-Morphism.f !) ∘ (idC +₁ idC +₁ [ idC , idC ]) ∘ [ (idC +₁ +-assocˡ ∘ M._≅_.to +-comm) ∘ i₁ , (idC +₁ +-assocˡ ∘ M._≅_.to +-comm) ∘ i₂ ∘ i₁ ] ∘ (out Y ∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ ([]-cong₂ +₁∘i₁ (pullˡ +₁∘i₂)) ⟩∘⟨refl ⟩
|
||||
out⁻¹ Y ∘ (idC +₁ F-Coalgebra-Morphism.f !) ∘ (idC +₁ idC +₁ [ idC , idC ]) ∘ [ i₁ ∘ idC , (i₂ ∘ (+-assocˡ ∘ M._≅_.to +-comm)) ∘ i₁ ] ∘ (out Y ∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ ([]-cong₂ identityʳ (pullʳ (pullʳ inject₁))) ⟩∘⟨refl ⟩
|
||||
out⁻¹ Y ∘ (idC +₁ F-Coalgebra-Morphism.f !) ∘ (idC +₁ idC +₁ [ idC , idC ]) ∘ [ i₁ , i₂ ∘ (+-assocˡ ∘ i₂) ] ∘ (out Y ∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ ([]-cong₂ refl (refl⟩∘⟨ inject₂)) ⟩∘⟨refl ⟩
|
||||
out⁻¹ Y ∘ (idC +₁ F-Coalgebra-Morphism.f !) ∘ (idC +₁ idC +₁ [ idC , idC ]) ∘ [ i₁ , i₂ ∘ i₂ ∘ i₂ ] ∘ (out Y ∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pullˡ ∘[] ⟩
|
||||
out⁻¹ Y ∘ (idC +₁ F-Coalgebra-Morphism.f !) ∘ [ (idC +₁ idC +₁ [ idC , idC ]) ∘ i₁ , (idC +₁ idC +₁ [ idC , idC ]) ∘ i₂ ∘ i₂ ∘ i₂ ] ∘ (out Y ∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ []-cong₂ +₁∘i₁ (pullˡ +₁∘i₂) ⟩∘⟨refl ⟩
|
||||
out⁻¹ Y ∘ (idC +₁ F-Coalgebra-Morphism.f !) ∘ [ i₁ ∘ idC , (i₂ ∘ (idC +₁ [ idC , idC ])) ∘ i₂ ∘ i₂ ] ∘ (out Y ∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ []-cong₂ identityʳ (pullˡ (pullʳ +₁∘i₂)) ⟩∘⟨refl ⟩
|
||||
out⁻¹ Y ∘ (idC +₁ F-Coalgebra-Morphism.f !) ∘ [ i₁ , (i₂ ∘ (i₂ ∘ [ idC , idC ])) ∘ i₂ ] ∘ (out Y ∘ f) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ []-cong₂ refl (pullʳ (pullʳ inject₂)) ⟩∘⟨refl ⟩
|
||||
out⁻¹ Y ∘ (idC +₁ F-Coalgebra-Morphism.f !) ∘ [ i₁ , i₂ ∘ (i₂ ∘ idC) ] ∘ (out Y ∘ f) ≈⟨ refl⟩∘⟨ pullˡ ∘[] ⟩
|
||||
out⁻¹ Y ∘ [ (idC +₁ F-Coalgebra-Morphism.f !) ∘ i₁ , (idC +₁ F-Coalgebra-Morphism.f !) ∘ i₂ ∘ (i₂ ∘ idC) ] ∘ (out Y ∘ f) ≈⟨ refl⟩∘⟨ (([]-cong₂ +₁∘i₁ (pullˡ +₁∘i₂)) ⟩∘⟨refl) ⟩
|
||||
out⁻¹ Y ∘ [ i₁ ∘ idC , (i₂ ∘ F-Coalgebra-Morphism.f !) ∘ (i₂ ∘ idC) ] ∘ (out Y ∘ f) ≈⟨ refl⟩∘⟨ (([]-cong₂ identityʳ (refl⟩∘⟨ identityʳ ○ assoc ○ refl⟩∘⟨ ⊤-id {! !})) ⟩∘⟨refl) ⟩
|
||||
out⁻¹ Y ∘ [ i₁ , i₂ ∘ idC ] ∘ (out Y ∘ f) ≈⟨ refl⟩∘⟨ (([]-cong₂ refl {! !}) ⟩∘⟨refl) ⟩
|
||||
{! !} ≈⟨ {! !} ⟩
|
||||
{! !} ≈⟨ {! !} ⟩
|
||||
f ∎
|
||||
|
||||
|
||||
|
||||
-- +-assocˡ∘i₁ = begin +-assocˡ ∘ i₁ ≈⟨ inject₁ ⟩ [ i₁ , i₂ ∘ i₁ ] ∎
|
||||
-- +-assocˡ∘i₁∘i₁ = begin +-assocˡ ∘ i₁ ∘ i₁ ≈⟨ inject₁ ⟩ [ i₁ , i₂ ∘ i₁ ] ∎
|
||||
-- given: out ∘ ! ≈ (! +₁ idC) ∘ long
|
||||
```
|
||||
|
||||
Old definitions:
|
||||
|
||||
```agda
|
||||
|
||||
record DelayFunctor : Set (o ⊔ ℓ ⊔ e) where
|
||||
field
|
||||
D : Endofunctor C
|
||||
|
@ -59,7 +186,7 @@ First I postulate the Functor *D*, maybe I should derive it...
|
|||
field
|
||||
now : ∀ {X} → X ⇒ D₀ X
|
||||
later : ∀ {X} → D₀ X ⇒ D₀ X
|
||||
isIso : ∀ {X} → IsIso [ now {X} , later {X} ]
|
||||
isIso : ∀ {X} → IsIso ([ now {X} , later {X} ])
|
||||
|
||||
out : ∀ {X} → D₀ X ⇒ X + D₀ X
|
||||
out {X} = IsIso.inv (isIso {X})
|
||||
|
@ -67,11 +194,7 @@ First I postulate the Functor *D*, maybe I should derive it...
|
|||
field
|
||||
coit : ∀ {X Y} → Y ⇒ X + Y → Y ⇒ D₀ X
|
||||
coit-law : ∀ {X Y} {f : Y ⇒ X + Y} → out ∘ (coit f) ≈ (idC +₁ (coit f)) ∘ f
|
||||
```
|
||||
|
||||
Now let's define the monad:
|
||||
|
||||
```agda
|
||||
record DelayMonad : Set (o ⊔ ℓ ⊔ e) where
|
||||
field
|
||||
D : DelayFunctor
|
||||
|
|
Loading…
Reference in a new issue