2023-08-19 12:15:34 +02:00
<!--
```agda
2023-08-16 14:54:50 +02:00
open import Level
2023-09-05 18:17:09 +02:00
open import Categories.Category
open import Categories.Monad
2023-08-16 14:54:50 +02:00
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
2023-09-05 18:17:09 +02:00
open import Categories.Category.Cartesian
2023-08-16 14:54:50 +02:00
open import Categories.Object.Terminal
2023-09-05 18:17:09 +02:00
open import Categories.Object.Coproduct
2023-08-16 14:54:50 +02:00
open import Categories.Category.Construction.F-Coalgebras
open import Categories.Functor.Coalgebra
open import Categories.Functor
2023-08-21 16:22:33 +02:00
open import Categories.Functor.Algebra
2023-08-16 14:54:50 +02:00
open import Categories.Monad.Construction.Kleisli
2023-08-21 16:22:33 +02:00
open import Categories.Category.Construction.F-Coalgebras
2023-09-05 18:17:09 +02:00
open import Categories.NaturalTransformation
2023-08-16 14:54:50 +02:00
import Categories.Morphism as M
import Categories.Morphism.Reasoning as MR
2023-08-19 12:15:34 +02:00
```
-->
2023-08-16 14:54:50 +02:00
2023-08-19 16:01:48 +02:00
## Summary
This file introduces the delay monad ** *D***
- [X] *Proposition 1* Characterization of the delay monad ** *D*** (here treated as definition)
- [ ] *Proposition 2* ** *D*** is commutative
## Code
2023-08-19 12:15:34 +02:00
```agda
2023-08-16 14:54:50 +02:00
module Monad.Instance.Delay {o ℓ e} (ED : ExtensiveDistributiveCategory o ℓ e) where
open ExtensiveDistributiveCategory ED renaming (U to C; id to idC)
open Cocartesian (Extensive.cocartesian extensive)
open Cartesian (ExtensiveDistributiveCategory.cartesian ED)
open BinaryProducts products
open M C
open MR C
open Equiv
open HomReasoning
2023-08-21 16:22:33 +02:00
open CoLambek
2023-08-19 16:01:48 +02:00
```
### *Proposition 1*: Characterization of the delay monad ***D***
2023-08-21 16:22:33 +02:00
First I postulate the Functor *D* , maybe I should derive it...
**TODO**:
- how to define using final coalgebra(s)?
- DX can be defined as retract of infinite streams, how?
- how to express **Theorem 8** in agda?
2023-08-19 16:01:48 +02:00
```agda
2023-09-05 18:17:09 +02:00
-- 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
2023-08-21 16:22:33 +02:00
record DelayFunctor : Set (o ⊔ ℓ ⊔ e) where
field
D : Endofunctor C
open Functor D public renaming (F₀ to D₀; F₁ to D₁)
2023-08-16 14:54:50 +02:00
field
now : ∀ {X} → X ⇒ D₀ X
later : ∀ {X} → D₀ X ⇒ D₀ X
2023-09-05 18:17:09 +02:00
isIso : ∀ {X} → IsIso ([ now {X} , later {X} ])
2023-08-21 16:22:33 +02:00
2023-08-16 14:54:50 +02:00
out : ∀ {X} → D₀ X ⇒ X + D₀ X
out {X} = IsIso.inv (isIso {X})
2023-08-21 16:22:33 +02:00
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
record DelayMonad : Set (o ⊔ ℓ ⊔ e) where
field
D : DelayFunctor
open DelayFunctor D
2023-08-16 14:54:50 +02:00
field
_* : ∀ {X Y} → X ⇒ D₀ Y → D₀ X ⇒ D₀ Y
*-law : ∀ {X Y} {f : X ⇒ D₀ Y} → out ∘ (f * ) ≈ [ out ∘ f , i₂ ∘ (f *) ] ∘ out
*-unique : ∀ {X Y} (f : X ⇒ D₀ Y) (h : D₀ X ⇒ D₀ Y) → h ≈ f *
*-resp-≈ : ∀ {X Y} {f h : X ⇒ D₀ Y} → f ≈ h → f * ≈ h *
unitLaw : ∀ {X} → out {X} ∘ now {X} ≈ i₁
unitLaw = begin
out ∘ now ≈⟨ refl⟩∘⟨ sym inject₁ ⟩
out ∘ [ now , later ] ∘ i₁ ≈⟨ cancelˡ (IsIso.isoˡ isIso) ⟩
i₁ ∎
toMonad : KleisliTriple C
toMonad = record
{ F₀ = D₀
; unit = now
; extend = _*
; identityʳ = λ {X} {Y} {k} → begin
k * ∘ now ≈⟨ introˡ (IsIso.isoʳ isIso) ⟩∘⟨refl ⟩
(([ now , later ] ∘ out) ∘ k *) ∘ now ≈⟨ pullʳ * -law ⟩∘⟨refl ⟩
([ now , later ] ∘ [ out ∘ k , i₂ ∘ (k *) ] ∘ out) ∘ now ≈⟨ pullʳ (pullʳ unitLaw) ⟩
[ now , later ] ∘ [ out ∘ k , i₂ ∘ (k *) ] ∘ i₁ ≈⟨ refl⟩∘⟨ inject₁ ⟩
[ now , later ] ∘ out ∘ k ≈⟨ cancelˡ (IsIso.isoʳ isIso) ⟩
k ∎
; identityˡ = λ {X} → sym (*-unique now idC)
; assoc = λ {X} {Y} {Z} {f} {g} → sym (*-unique ((g *) ∘ f) ((g * ) ∘ (f *)))
; sym-assoc = λ {X} {Y} {Z} {f} {g} → *-unique ((g * ) ∘ f) ((g *) ∘ (f * ))
; extend-≈ = *-resp-≈
2023-08-19 12:15:34 +02:00
}
2023-08-19 16:01:48 +02:00
```
2023-08-21 16:22:33 +02:00
### Definition 30: Search-Algebras
```agda
record SearchAlgebra (DF : DelayFunctor) : Set (o ⊔ ℓ ⊔ e) where
open DelayFunctor DF
field
FA : F-Algebra D
open F-Algebra FA
field
now-id : α ∘ now ≈ idC
later-same : α ∘ later ≈ α
```
### Proposition 31 : the category of uniform-iteration algebras coincides with the category of search-algebras
TODOs:
- [ ] Define SearchAlgebras (and SearchAlgebra morphisms)
- [ ] show StrongEquivalence
- [ ] Show 'ElgotAlgebra⇔Search+***D***'
```agda
record SearchAlgebras (DF : DelayFunctor) : Set (o ⊔ ℓ ⊔ e) where
open DelayFunctor DF
```