mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
46 KiB
46 KiB
Summary
This file introduces the delay monad D
- Proposition 1 Characterization of the delay monad D (here treated as definition)
- Proposition 2 D is commutative
Code
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 open CoLambek### *Proposition 1*: Characterization of the delay monad ***D***
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?
record DelayFunctor : Set (o ⊔ ℓ ⊔ e) where field D : Endofunctor C open Functor D public renaming (F₀ to D₀; F₁ to D₁) field now : ∀ {X} → X ⇒ D₀ X later : ∀ {X} → D₀ X ⇒ D₀ X isIso : ∀ {X} → IsIso [ now {X} , later {X} ] out : ∀ {X} → D₀ X ⇒ X + D₀ X out {X} = IsIso.inv (isIso {X}) 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:
record DelayMonad : Set (o ⊔ ℓ ⊔ e) where field D : DelayFunctor open DelayFunctor D 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-≈ }
Definition 30: Search-Algebras
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'
record SearchAlgebras (DF : DelayFunctor) : Set (o ⊔ ℓ ⊔ e) where open DelayFunctor DF