bsc-leon-vatthauer/agda/bsc-thesis/Monad.Instance.Delay.Quotienting.md
2024-02-09 17:53:52 +01:00

146 KiB
Raw Blame History

module Monad.Instance.Delay.Quotienting {o  e} (ambient : Ambient o  e) where
  open Ambient ambient
  open import Categories.Diagram.Coequalizer C 
  open import Monad.Instance.Delay ambient
  open F-Coalgebra-Morphism using () renaming (f to u; commutes to coalg-commutes)

Quotienting the delay monad by weak bisimilarity

The coinductive definition of the delay monad yields a 'wrong' kind of equality called strong bisimilarity, which differentiates between computations with different execution times, instead we want two computations to be equal, when they both terminate with the same result (or don't terminate at all). This is called weak bisimilarity. The ideas is then to quotient the delay monad with this 'right' kind of equality and show that the result \tilde{D} is again a (strong) monad.

In category theory existence of coequalizers corresponds to qoutienting, so we will express the quotiented delay monad via the following coequalizer:

Preliminaries

Existence of the coequalizer doesn't suffice, we will need some conditions having to do with preservation of the coequalizer, so let's first define what it means for a coequalizer to be preserved by an endofunctor:

  preserves :  {X Y} {f g : X  Y}  Endofunctor C  Coequalizer f g  Set (o    e)
  preserves {X} {Y} {f} {g} F coeq = IsCoequalizer (Functor.₁ F f) (Functor.₁ F g) (Functor.₁ F (Coequalizer.arr coeq))

We will now show that the following conditions are equivalent:

  1. For every X, the coequalizer is preserved by D
  2. every \tilde{D}X extends to a search-algebra, so that each ρ_X is a D-algebra morphism
  3. for every X, (\tilde{D}X, ρ ∘ now : X → \tilde{D}X) is a stable free elgot algebra on X, ρ_X is a D-algebra morphism and ρ_X = ((ρ_X ∘ now + id) ∘ out)^\#
  4. \tilde{D} extends to a strong monad, so that ρ is a strong monad morphism
  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 Quotiented (coeqs :  X  Coequalizer (extend (ι {X})) (D₁ π₁)) where

      Ď₀ : Obj  Obj
      Ď₀ X = Coequalizer.obj (coeqs X)

      ρ :  {X}  D₀ X  Ď₀ X
      ρ {X} = Coequalizer.arr (coeqs X)

      ρ-epi :  {X}  Epi (ρ {X})
      ρ-epi {X} = Coequalizer⇒Epi (coeqs X)

      ρ▷ :  {X}  ρ    ρ {X}
      ρ▷ {X} = sym (begin 
        ρ ≈⟨ introʳ intro-helper  
        ρ  D₁ π₁  D₁  idC , s  z  Terminal.! terminal  ≈⟨ pullˡ (sym equality)  
        (ρ  extend ι)  D₁  idC , s  z  Terminal.! terminal  ≈⟨ pullʳ (sym k-assoc)  
        ρ  extend (extend ι  now   idC , s  z  Terminal.! terminal ) ≈⟨ (refl⟩∘⟨ extend-≈ (pullˡ k-identityʳ))  
        ρ  extend (ι   idC , s  z  Terminal.! terminal ) ≈⟨ (refl⟩∘⟨ extend-≈ helper)  
        ρ  extend (  now) ≈⟨ (refl⟩∘⟨ sym (▷∘extendʳ now))  
        ρ  extend now   ≈⟨ elim-center k-identityˡ  
        ρ   )
        where
          open Coequalizer (coeqs X) using (equality; coequalize; id-coequalize; coequalize-resp-≈) renaming (universal to coeq-universal; unique to coeq-unique; unique to coeq-unique)
          intro-helper : D₁ π₁  D₁  idC , s  z  Terminal.! terminal   idC
          intro-helper = sym D-homomorphism  D-resp-≈ project₁  D-identity
          helper : ι   idC , s  z  Terminal.! terminal     now
          helper = begin 
            ι   idC , s  z  Terminal.! terminal  ≈⟨ introˡ (_≅_.isoˡ out-≅)  
            (out⁻¹  out)  ι   idC , s  z  Terminal.! terminal  ≈⟨ pullʳ (pullˡ (coalg-commutes ((Terminal.! (coalgebras X) {A = record { A = X × N ; α = _≅_.from nno-iso }})))) 
            out⁻¹  ((idC +₁ ι)  _≅_.from nno-iso)   idC , s  z  Terminal.! terminal  ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ sym (⁂∘⟨⟩  ⟨⟩-cong₂ identity² refl)) 
            out⁻¹  ((idC +₁ ι)  _≅_.from nno-iso)  (idC  s)   idC , z  Terminal.! terminal  ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ sym inject₂ ⟩∘⟨refl) 
            out⁻¹  ((idC +₁ ι)  _≅_.from nno-iso)  (_≅_.to nno-iso  i₂)   idC , z  Terminal.! terminal  ≈⟨ (refl⟩∘⟨ (pullʳ (pullˡ (cancelˡ (_≅_.isoʳ nno-iso))))) 
            out⁻¹  (idC +₁ ι)  i₂   idC , z  Terminal.! terminal  ≈⟨ (refl⟩∘⟨ pullˡ +₁∘i₂) 
            out⁻¹  (i₂  ι)   idC , z  Terminal.! terminal  ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ sym inject₁) 
            out⁻¹  (i₂  ι)  _≅_.to nno-iso  i₁ ≈⟨ (refl⟩∘⟨ intro-center (_≅_.isoˡ out-≅) ⟩∘⟨refl) 
            out⁻¹  (i₂  (out⁻¹  out)  ι)  _≅_.to nno-iso  i₁ ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ pullʳ (coalg-commutes ((Terminal.! (coalgebras X) {A = record { A = X × N ; α = _≅_.from nno-iso }})))) ⟩∘⟨refl) 
            out⁻¹  (i₂  out⁻¹  (idC +₁ ι)  _≅_.from nno-iso)  _≅_.to nno-iso  i₁ ≈⟨ (refl⟩∘⟨ (pullʳ (pullˡ (pullʳ (cancelʳ (_≅_.isoʳ nno-iso)))))) 
            out⁻¹  i₂  (out⁻¹  (idC +₁ ι))  i₁ ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ pullʳ +₁∘i₁) 
            out⁻¹  i₂  out⁻¹  i₁  idC ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ identityʳ) 
            out⁻¹  i₂  out⁻¹  i₁ ≈⟨ ((refl⟩∘⟨ sym-assoc)  assoc²'') 
              now 

      Ď-Functor : Endofunctor C
      Ď-Functor = record
        { F₀ = Ď₀
        ; F₁ = F₁'
        ; identity = λ {X}  Coequalizer.coequalize-resp-≈ (coeqs X) (elimʳ D-identity)  sym (Coequalizer.id-coequalize (coeqs X))
        ; homomorphism = λ {X} {Y} {Z} {f} {g}  sym (Coequalizer.unique (coeqs X) (sym (begin 
          (F₁' g  F₁' f)  ρ ≈⟨ pullʳ (sym (Coequalizer.universal (coeqs X)))  
          F₁' g  ρ  D₁ f ≈⟨ pullˡ (sym (Coequalizer.universal (coeqs Y))) 
          (ρ  D₁ g)  D₁ f ≈⟨ pullʳ (sym D-homomorphism) 
          ρ  D₁ (g  f) )))
        ; F-resp-≈ = λ {X} {Y} {f} {g} eq  Coequalizer.coequalize-resp-≈ (coeqs X) (refl⟩∘⟨ (D-resp-≈ eq))
        }
        where
          F₁' :  {X} {Y} (f : X  Y)  Ď₀ X  Ď₀ Y
          F₁' {X} {Y} f = coequalize {h = ρ  D₁ f} (begin 
            (ρ  D₁ f)  extend ι ≈⟨ pullʳ (sym k-assoc)  
            ρ  extend (extend (now  f)  ι) ≈⟨ refl⟩∘⟨ (extend-≈ (sym (Terminal.!-unique (coalgebras Y) (record { f = extend (now  f)  ι ; commutes = begin 
              out  extend (now  f)  ι ≈⟨ pullˡ (extendlaw (now  f))  
              ([ out  now  f , i₂  D₁ f ]  out)  ι ≈⟨ pullʳ (coalg-commutes (Terminal.! (coalgebras X) {A = record { A = X × N ; α = _≅_.from nno-iso }}))  
              [ out  now  f , i₂  D₁ f ]  (idC +₁ ι)   _≅_.from nno-iso ≈⟨ (([]-cong₂ (pullˡ unitlaw) refl) ⟩∘⟨refl)  
              (f +₁ D₁ f)  (idC +₁ ι)  _≅_.from nno-iso ≈⟨ pullˡ +₁∘+₁  
              (f  idC +₁ D₁ f  ι)  _≅_.from nno-iso ≈⟨ ((+₁-cong₂ id-comm (sym identityʳ)) ⟩∘⟨refl)  
              (idC  f +₁ (D₁ f  ι)  idC)  _≅_.from nno-iso ≈⟨ sym (pullˡ +₁∘+₁)  
              (idC +₁ extend (now  f)  ι)  (f +₁ idC)  _≅_.from nno-iso  })))) 
            ρ  extend (u (Terminal.! (coalgebras Y) {A = record { A = X × N ; α = (f +₁ idC)  _≅_.from nno-iso }})) ≈⟨ (refl⟩∘⟨ (extend-≈ (Terminal.!-unique (coalgebras Y) (record { f = ι  (f  idC) ; commutes = begin 
              out  ι  (f  idC) ≈⟨ pullˡ (coalg-commutes (Terminal.! (coalgebras Y) {A = record { A = Y × N ; α = _≅_.from nno-iso }}))  
              ((idC +₁ ι)  _≅_.from nno-iso)  (f  idC) ≈⟨ assoc  
              (idC +₁ ι)  _≅_.from nno-iso  (f  idC) ≈⟨ (refl⟩∘⟨ (introʳ (_≅_.isoˡ nno-iso)))  
              (idC +₁ ι)  (_≅_.from nno-iso  (f  idC))  [  idC , z  Terminal.! terminal  , idC  s ]  _≅_.from nno-iso ≈⟨ (refl⟩∘⟨ pullʳ (pullˡ ∘[]))  
              (idC +₁ ι)  _≅_.from nno-iso  [ (f  idC)   idC , z  Terminal.! terminal  , (f  idC)  (idC  s) ]  _≅_.from nno-iso ≈⟨ (refl⟩∘⟨ (refl⟩∘⟨ (([]-cong₂ (⁂∘⟨⟩  ⟨⟩-cong₂ identityʳ identityˡ) (⁂∘⁂  ⁂-cong₂ identityʳ identityˡ)) ⟩∘⟨refl)))  
              (idC +₁ ι)  _≅_.from nno-iso  [  f , z  Terminal.! terminal  , (f  s) ]  _≅_.from nno-iso ≈⟨ (refl⟩∘⟨ refl⟩∘⟨ []-cong₂ (⟨⟩-cong₂ refl (refl⟩∘⟨ Terminal.!-unique terminal (Terminal.! terminal  f))) refl ⟩∘⟨refl)  
              (idC +₁ ι)  _≅_.from nno-iso  [  f , z  Terminal.! terminal  f  , (f  s) ]  _≅_.from nno-iso ≈⟨ sym (refl⟩∘⟨ refl⟩∘⟨ []-cong₂ (⟨⟩∘  ⟨⟩-cong₂ identityˡ assoc) (⁂∘⁂  ⁂-cong₂ identityˡ identityʳ) ⟩∘⟨refl)  
              (idC +₁ ι)  _≅_.from nno-iso  [  idC , z  Terminal.! terminal   f , (idC  s)  (f  idC) ]  _≅_.from nno-iso ≈⟨ sym (refl⟩∘⟨ (pullʳ (pullˡ []∘+₁)))  
              (idC +₁ ι)  (_≅_.from nno-iso  [  idC , z  Terminal.! terminal  , idC  s ])  (f +₁ (f  idC))  _≅_.from nno-iso ≈⟨ sym (refl⟩∘⟨ introˡ (_≅_.isoʳ nno-iso))  
              (idC +₁ ι)  (f +₁ (f  idC))  _≅_.from nno-iso ≈⟨ pullˡ (+₁∘+₁  +₁-cong₂ identityˡ refl)  
              (f +₁ ι  (f  idC))  _≅_.from nno-iso ≈⟨ sym (pullˡ (+₁∘+₁  +₁-cong₂ identityˡ identityʳ)) 
              (idC +₁ ι  (f  idC))  (f +₁ idC)  _≅_.from nno-iso  })))) 
            ρ  extend (ι  (f  idC)) ≈⟨ (refl⟩∘⟨ extend-≈ (pushˡ (sym k-identityʳ))) 
            ρ  extend (extend ι  now  (f  idC)) ≈⟨ pushʳ k-assoc 
            (ρ  extend ι)  D₁ (f  idC) ≈⟨ pushˡ (Coequalizer.equality (coeqs Y)) 
            ρ  D₁ π₁  D₁ (f  idC) ≈⟨ refl⟩∘⟨ sym D-homomorphism 
            ρ  D₁ (π₁  (f  idC)) ≈⟨ (refl⟩∘⟨ (D-resp-≈ project₁)) 
            ρ  D₁ (f  π₁) ≈⟨ pushʳ D-homomorphism 
            (ρ  D₁ f)  D₁ π₁ )
            where 
              open Coequalizer (coeqs X) using (coequalize; equality) renaming (universal to coeq-universal)

      ρ-natural : NaturalTransformation D-Functor Ď-Functor
      ρ-natural = ntHelper (record 
        { η = λ X  ρ {X} 
        ; commute = λ {X} {Y} f  Coequalizer.universal (coeqs X)
        })