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

463 KiB
Raw Blame History

{-# OPTIONS --guardedness #-}

open import Level renaming (zero to -zero; suc to -suc)
open import Function.Bundles using (_⟨$⟩_) renaming (Func to _⟶_)
open import Function.Construct.Identity renaming (function to idₛ)
open import Function.Construct.Setoid renaming (setoid to _⇨_; _∙_ to _∘_)
open import Relation.Binary
open import Data.Sum using (_⊎_; inj₁; inj₂)
import Relation.Binary.PropositionalEquality as Eq
open Eq using (_≡_)
open import Function using (id)
open import Categories.Monad
open import Categories.Category.Instance.Setoids
open import Categories.NaturalTransformation hiding (id)
open import Data.Product
open import Data.Product.Relation.Binary.Pointwise.NonDependent
open import Category.Ambient.Setoids

module Monad.Instance.Setoids.Delay {c } where

open Setoid using () renaming (Carrier to _; _≈_ to [_][_≡_])
open module eq (S : Setoid c (c  )) = IsEquivalence (Setoid.isEquivalence S) using () renaming (refl to ≡-refl; sym to ≡-sym; trans to ≡-trans)


mutual

  data Delay (A : Set c) : Set c where
    now   : A         Delay A
    later : Delay A  Delay A

  record Delay (A : Set c) : Set c where
    coinductive
    field
      force : Delay A

open Delay public

module Bisimilarity (A : Setoid c (c  )) where

  never : Delay  A 
  never : Delay  A 

  never = later never
  force never = never

  -- Removes a later constructor, if possible.

  drop-later : Delay  A   Delay  A 
  drop-later (now   x) = now x
  drop-later (later x) = force x

  mutual

    -- adapted from https://www.cse.chalmers.se/nad/listings/delay-monad/Delay-monad.Bisimilarity.html
    
    data __ : Delay  A   Delay  A   Set (c  ) where
      now    :  {x y}  [ A ][ x  y ]  now x  now y
      later  :  {x y}  force x  force y  later x  later y

    record __ (x y : Delay  A ) : Set (c  ) where
      coinductive
      
      field
        force : x  y

  open __ public

  -- strong bisimilarity of now and later leads to a clash
  nowlater :  {'}{Z : Set '}{x :  A }{y : Delay  A }  now x  later y  Z
  nowlater ()

  -refl  :  {x : Delay  A }  x  x
  -refl :  {x : Delay  A }  x  x

  -refl {now x}   = now (≡-refl A) 
  -refl {later x} = later -refl

  force (-refl {now x})   = now (≡-refl A)
  force (-refl {later x}) = later -refl

  -sym  :  {x y : Delay  A }  x  y  y  x
  -sym :  {x y : Delay  A }  x  y  y  x

  -sym (now x≡y) = now (≡-sym A x≡y)
  -sym (later fxfy) = later (-sym fxfy)

  force (-sym xy) = -sym (force xy)

  -trans  :  {x y z : Delay  A }  x  y  y  z  x  z
  -trans :  {x y z : Delay  A }  x  y  y  z  x  z

  -trans (now x≡y) (now y≡z) = now (≡-trans A x≡y y≡z)
  -trans (later xy) (later yz) = later (-trans xy yz)

  force (-trans xy yz) = -trans (force xy) (force yz)

  data _↓_ : Delay  A    A   Set (c  ) where
    now↓   :  {x y} (x≡y : [ A ][ x  y ])  now x  y
    later↓ :  {x y} (x↓y : (force x)  y)  later x  y

  unique↓ :  {a : Delay  A  } {x y :  A }  a  x  a  y  [ A ][ x  y ]
  unique↓ (now↓ a≡x) (now↓ a≡y) = ≡-trans A (≡-sym A a≡x) a≡y
  unique↓ (later↓ fb↓x) (later↓ fb↓y) = unique↓ fb↓x fb↓y

  mutual
  
    data _≈_ : Delay  A   Delay  A   Set (c  ) where
      ↓≈     :  {x y a b} (a≡b : [ A ][ a  b ]) (x↓a : x  a) (y↓b : y  b)  x  y
      later≈ :  {x y} (x≈y : force x ≈′ force y)  later x  later y

    record _≈_ (x y : Delay  A ) : Set (c  ) where
      coinductive
      
      field
        force≈ : x  y

  open _≈_ public

  ≈→≈′ :  {x y : Delay  A }  x  y  x ≈′ y
  force≈ (≈→≈′ x≈y) = x≈y

  ≡↓ :  {x y :  A } {z : Delay  A }  [ A ][ x  y ]  z  x  z  y
  ≡↓ {x} {y} {.(now _)} x≡y (now↓ a≡x) = now↓ (≡-trans A a≡x x≡y)
  ≡↓ {x} {y} {.(later _)} x≡y (later↓ z↓x) = later↓ (≡↓ x≡y z↓x)

  ≈↓  :  {x y : Delay  A  } {a :  A }  x  y  x  a  y  a

  ≈↓ (↓≈ c≡b (now↓ x≡c) y↓b) (now↓ x≡a) = ≡↓ (≡-trans A (≡-trans A (≡-sym A c≡b) (≡-sym A x≡c)) x≡a) y↓b
  ≈↓ (↓≈ c≡b (later↓ x↓c) y↓b) (later↓ x↓a) = ≡↓ (≡-trans A (≡-sym A c≡b) (≡-sym A (unique↓ x↓a x↓c))) y↓b
  ≈↓ (later≈ fx≈fy) (later↓ fx↓a) = later↓ (≈↓ (force≈ fx≈fy) fx↓a)


  ≈-refl  :  {x : Delay  A }  x  x
  ≈′-refl :  {x : Delay  A }  x ≈′ x

  ≈-refl {now x}   = ↓≈ (≡-refl A) (now↓ (≡-refl A)) (now↓ (≡-refl A))
  ≈-refl {later x} = later≈ ≈′-refl

  force≈ ≈′-refl = ≈-refl


  ≈-sym  :  {x y : Delay  A }  x  y  y  x
  ≈′-sym :  {x y : Delay  A }  x ≈′ y  y ≈′ x

  ≈-sym (↓≈ a≡b x↓a y↓b) = ↓≈ (≡-sym A a≡b) y↓b x↓a
  ≈-sym (later≈ x≈y) = later≈ (≈′-sym x≈y)

  force≈ (≈′-sym xy) = ≈-sym (force≈ xy)


  ≈-trans  :  {x y z : Delay  A }  x  y  y  z  x  z
  ≈′-trans :  {x y z : Delay  A }  x ≈′ y  y ≈′ z  x ≈′ z

  ≈-trans (↓≈ a≡b x↓a y↓b) (↓≈ c≡d y↓c z↓d) = ↓≈ (≡-trans A (≡-trans A a≡b (unique↓ y↓b y↓c)) c≡d) x↓a z↓d
  ≈-trans (↓≈ a≡b z↓a (later↓ x↓b)) (later≈ x≈y) = ↓≈ a≡b z↓a (later↓ (≈↓ (force≈ x≈y) x↓b))
  ≈-trans (later≈ x≈y) (↓≈ a≡b (later↓ y↓a) z↓b) = ↓≈ a≡b (later↓ (≈↓ (≈-sym (force≈ x≈y)) y↓a)) z↓b
  ≈-trans (later≈ x≈y) (later≈ y≈z) = later≈ (≈′-trans x≈y y≈z) 

  force≈ (≈′-trans x≈y y≈z) = ≈-trans (force≈ x≈y) (force≈ y≈z)

  mutual
    data _≲_ : Delay  A   Delay  A   Set (c  ) where
      ↓≲     :  {y a} (y↓a : y  a)  now a  y
      later≲ :  {x y} (x≈y : force x ≲′ force y)  later x  later y

    record _≲_ (x y : Delay  A ) : Set (c  ) where
      coinductive
      
      field
        force≲ : x  y

  open _≲_ public

open Bisimilarity renaming (_≈_ to [_][_≈_]; _≈_ to [_][_≈_]; __ to [_][__]; __ to [_][__]; _↓_ to [_][_↓_]; _≲_ to [_][_≲_]; _≲_ to [_][_≲_])


module DelayMonad where
  Delayₛ : Setoid c (c  )  Setoid c (c  )
  Delayₛ A = record { Carrier = Delay  A  ; _≈_ = [_][_≈_] A ; isEquivalence = record { refl = ≈-refl A ; sym = ≈-sym A ; trans = ≈-trans A } }
  Delayₛ : Setoid c (c  )  Setoid c (c  )
  Delayₛ A = record { Carrier = Delay  A  ; _≈_ = [_][__] A ; isEquivalence = record { refl = -refl A ; sym = -sym A ; trans = -trans A } }
  <_> = _⟨$⟩_
  open _⟶_ using (cong)

  ∼⇒≈ :  {A : Setoid c (c  )} {x y : Delay  A }  [ A ][ x  y ]  [ A ][ x  y ]
  ∼′⇒≈′ :  {A : Setoid c (c  )} {x y : Delay  A }  [ A ][ x  y ]  [ A ][ x ≈′ y ]
  ∼⇒≈ {A} {.(now _)} {.(now _)} (now a≡b) = ↓≈ a≡b (now↓ (≡-refl A)) (now↓ (≡-refl A))
  ∼⇒≈ {A} {.(later _)} {.(later _)} (later xy) = later≈ (∼′⇒≈′ xy)
  force≈ (∼′⇒≈′ {A} {x} {y} xy) = ∼⇒≈ (force xy)

  now-cong :  {A : Setoid c (c  )} {x y :  A }  [ A ][ x  y ]  [ A ][ now x  now y ]
  now-cong {A} {x} {y} x≡y = ↓≈ x≡y (now↓ (≡-refl A)) (now↓ (≡-refl A))

  later-cong :  {A : Setoid c (c  )} {x y : Delay  A }  [ A ][ force x ≈′ force y ]  [ A ][ later x  later y ]
  later-cong {A} {x} {y} x≈y = later≈ x≈y

  now-cong :  {A : Setoid c (c  )} {x y :  A }  [ A ][ x  y ]  [ A ][ now x  now y ]
  now-cong {A} {x} {y} x≡y = now x≡y

  now-inj :  {A : Setoid c (c  )} {x y :  A }  [ A ][ now x  now y ]  [ A ][ x  y ]
  now-inj {A} {x} {y} (↓≈ a≡b (now↓ x≡a) (now↓ y≡b)) = ≡-trans A x≡a (≡-trans A a≡b (≡-sym A y≡b))

  liftF :  {A B : Set c}  (A  B)  Delay A  Delay B
  liftF :  {A B : Set c}  (A  B)  Delay A  Delay B
  liftF f (now x) = now (f x)
  liftF f (later x) = later (liftF f x)
  force (liftF f x) = liftF f (force x)

  lift↓ :  {A B : Setoid c (c  )} (f : A  B) {x : Delay  A } {b :  A }  [ A ][ x  b ]  [ B ][ (liftF (< f >) x)  (f ⟨$⟩ b) ]
  lift↓ {A} {B} f {now x} {b} (now↓ x≡a) = now↓ (cong f x≡a)
  lift↓ {A} {B} f {later x} {b} (later↓ x↓b) = later↓ (lift↓ f x↓b)

  lift-cong :  {A B : Setoid c (c  )} (f : A  B) {x y : Delay  A }  [ A ][ x  y ]  [ B ][ liftF < f > x  liftF < f > y ]
  lift-cong :  {A B : Setoid c (c  )} (f : A  B) {x y : Delay  A }  [ A ][ x ≈′ y ]  [ B ][ liftF < f > x ≈′ liftF < f > y ]
  lift-cong {A} {B} f {now x} {now y} (↓≈ a≡b (now↓ x≡a) (now↓ y≡b)) = ↓≈ (cong f a≡b) (now↓ (cong f x≡a)) (now↓ (cong f y≡b))
  lift-cong {A} {B} f {now x} {later y} (↓≈ a≡b (now↓ x≡a) (later↓ y↓b)) = ↓≈ (cong f a≡b) (now↓ (cong f x≡a)) (later↓ (lift↓ f y↓b))
  lift-cong {A} {B} f {later x} {now y} (↓≈ a≡b (later↓ x↓a) (now↓ y≡b)) = ↓≈ (cong f a≡b) (later↓ (lift↓ f x↓a)) (now↓ (cong f y≡b))
  lift-cong {A} {B} f {later x} {later y} (↓≈ a≡b (later↓ x↓a) (later↓ y↓b)) = later≈ (lift-cong f (≈→≈′ A (↓≈ a≡b x↓a y↓b)))
  lift-cong {A} {B} f {later x} {later y} (later≈ x≈y) = later≈ (lift-cong {A} {B} f x≈y)
  force≈ (lift-cong {A} {B} f {x} {y} x≈y) = lift-cong f (force≈ x≈y)

  liftFₛ :  {A B : Setoid c (c  )}  A  B  Delayₛ A  Delayₛ B
  liftFₛ {A} {B} f = record { to = liftF < f > ; cong = lift-cong f }

  liftFₛ :  {A B : Setoid c (c  )}  A  B  Delayₛ A  Delayₛ B
  liftFₛ {A} {B} f = record { to = liftF < f > ; cong = -cong }
    where
      -cong :  {x y}  [ A ][ x  y ]  [ B ][ liftF < f > x  liftF < f > y ]
      -cong :  {x y}  [ A ][ x  y ]  [ B ][ liftF < f > x  liftF < f > y ]
      force (-cong {x} {y} xy) = -cong (force xy)
      -cong {.(now _)} {.(now _)} (now x≡y) = now-cong (cong f x≡y)
      -cong {.(later _)} {.(later _)} (later xy) = later (-cong xy)

  lift-comp :  {A B C : Setoid c (c  )} {f : A  B} {g : B  C} {x y : Delay  A }  [ A ][ x  y ]  [ C ][ liftFₛ (g  f) ⟨$⟩ x  (liftFₛ g  liftFₛ f) ⟨$⟩ y ]
  lift-comp :  {A B C : Setoid c (c  )} {f : A  B} {g : B  C} {x y : Delay  A }  [ A ][ x  y ]  [ C ][ liftFₛ (g  f) ⟨$⟩ x  (liftFₛ g  liftFₛ f) ⟨$⟩ y ]
  lift-comp {A} {B} {C} {f} {g} {.(now _)} {.(now _)} (now x≡y) = now-cong (cong g (cong f (x≡y)))
  lift-comp {A} {B} {C} {f} {g} {.(later _)} {.(later _)} (later xy) = later (lift-comp {A} {B} {C} {f} {g} xy)
  force (lift-comp {A} {B} {C} {f} {g} {x} {y} xy) = lift-comp {A} {B} {C} {f} {g} {x} {y} (force xy)


  -- this needs polymorphic universe levels
  _≋_ :  {c' '} {A B : Setoid c' '}  A  B  A  B  Set (c'  ')
  _≋_ {c'} {'} {A} {B} f g = Setoid._≈_ (A  B) f g

  later-self :  {A : Setoid c (c  )} {x : Delay  A }  [ A ][ force x  later x ]
  later-self {A} {x} with force x in eqx
  ...                | now y = ↓≈ (≡-refl A) (now↓ (≡-refl A)) (later↓ helper)
    where
      helper : [ A ][ force x  y ]
      helper rewrite eqx = now↓ (≡-refl A)
  ...                | later y = later≈ helper
    where
      helper : [ A ][ force y ≈′ force x ]
      force≈ (helper) rewrite eqx = later-self {x = y}

  later-eq :  {A : Setoid c (c  )} {x : Delay  A } {y : Delay  A }  [ A ][ later x  y ]  [ A ][ force x  y ] 
  later-eq {A} {x} {now y} (↓≈ a≡b (later↓ x↓a) (now↓ y≡b)) = ↓≈ a≡b x↓a (now↓ y≡b)
  later-eq {A} {x} {later y} x≈ly = ≈-trans A later-self x≈ly

  lift-id :  {A : Setoid c (c  )}  (liftFₛ (idₛ A))  (idₛ (Delayₛ A))
  lift-id :  {A : Setoid c (c  )} {x : Delay  A }  [ A ][ (liftF id) x ≈′ id x ]
  lift-id {A} {now x} = ≈-refl A
  lift-id {A} {later x} = later≈ lift-id
  force≈ (lift-id {A} {x}) = lift-id

  lift-comp :  {A B C : Setoid c (c  )} {f : A  B} {g : B  C}  liftFₛ (g  f)  (liftFₛ g  liftFₛ f)
  lift-comp :  {A B C : Setoid c (c  )} {f : A  B} {g : B  C} {x : Delay  A }  [ C ][ liftFₛ (g  f) ⟨$⟩ x ≈′ (liftFₛ g  liftFₛ f) ⟨$⟩ x ]
  force≈ (lift-comp {A} {B} {C} {f} {g} {x}) = lift-comp {A} {B} {C} {f} {g}
  lift-comp {A} {B} {C} {f} {g} {now x} = ≈-refl C
  lift-comp {A} {B} {C} {f} {g} {later x} = later≈ (lift-comp {A} {B} {C} {f} {g})

  lift-resp-≈ :  {A B : Setoid c (c  )} {f g : A  B}  f  g  liftFₛ f  liftFₛ g
  lift-resp-≈′ :  {A B : Setoid c (c  )} {f g : A  B}  f  g   {x : Delay  A }  [ B ][ liftFₛ f ⟨$⟩ x ≈′ liftFₛ g ⟨$⟩ x ]
  lift-resp-≈ {A} {B} {f} {g} f≋g {now x} = now-cong f≋g
  lift-resp-≈ {A} {B} {f} {g} f≋g {later x} = later≈ (lift-resp-≈′ {A} {B} {f} {g} f≋g)
  force≈ (lift-resp-≈′ {A} {B} {f} {g} f≋g {x}) = lift-resp-≈ {A} {B} {f} {g} f≋g

  ηₛ :  (A : Setoid c (c  ))  A  Delayₛ A
  ηₛ A = record { to = now ; cong = now-cong }

  η-natural :  {A B : Setoid c (c  )} (f : A  B)  (ηₛ B  f)  (liftFₛ f  ηₛ A)
  η-natural {A} {B} f {x} = now-cong (cong f (≡-refl A))

  μ :  {A : Setoid c (c  )}  Delay (Delay  A )  Delay  A 
  μ′ :  {A : Setoid c (c  )}  Delay (Delay  A )  Delay  A 
  force (μ′ {A} x) = μ {A} (force x)
  μ {A} (now x) = x
  μ {A} (later x) = later (μ′ {A} x)

  μ↓-trans :  {A : Setoid c (c  )} {x : Delay (Delay  A )} {y : Delay  A } {b :  A }  [ Delayₛ A ][ x  y ]  [ A ][ y  b ]  [ A ][ (μ {A} x)  b ]
  μ↓-trans {A} {now x} {y} {b} (now↓ x≡y) y↓b = ≈↓ A (≈-sym A x≡y) y↓b
  μ↓-trans {A} {later x} {now y} {b} (later↓ x↓y) (now↓ y≡b) = later↓ (μ↓-trans x↓y (now↓ y≡b))
  μ↓-trans {A} {later x} {later y} {b} (later↓ x↓y) (later↓ y↓b) = later↓ (μ↓-trans (≡↓ (Delayₛ A) (≈-sym A later-self) x↓y) y↓b)

  μ↓ :  {A : Setoid c (c  )} {x : Delay (Delay  A )} {y : Delay  A }  [ Delayₛ A ][ x  y ]  [ A ][ (μ {A} x)  y ]
  μ↓′ :  {A : Setoid c (c  )} {x : Delay (Delay  A )} {y : Delay  A }  [ Delayₛ A ][ x  y ]  [ A ][ (μ {A} x) ≈′ y ]
  force≈ (μ↓′ {A} {x} {y} x↓y) = μ↓ x↓y
  μ↓ {A} {now x} {y} (now↓ x≡y) = x≡y
  μ↓ {A} {later x} {now y} (later↓ x↓y) = ≈-trans A (≈-sym A later-self) (↓≈ (≡-refl A) (μ↓-trans x↓y (now↓ (≡-refl A))) (now↓ (≡-refl A)))
  μ↓ {A} {later x} {later y} (later↓ x↓y) = later≈ (μ↓′ {A} {force x} {force y} (≡↓ (Delayₛ A) (≈-sym A later-self) x↓y))

  μ-cong :  (A : Setoid c (c  )) {x y : Delay (Delay  A )}  [ Delayₛ A ][ x  y ]  [ A ][ μ {A} x  μ {A} y ]
  μ-cong :  (A : Setoid c (c  )) {x y : Delay (Delay  A )}  [ Delayₛ A ][ x ≈′ y ]  [ A ][ μ {A} x ≈′ μ {A} y ]
  μ-cong A {now x} {now y} x≈y = now-inj x≈y
  μ-cong A {now x} {later y} (↓≈ a≡b (now↓ x≡a) (later↓ y↓b)) = ≈-trans A (≈-sym A (μ↓ (≡↓ (Delayₛ A) (≈-trans A (≈-sym A a≡b) (≈-sym A x≡a)) y↓b))) later-self
  μ-cong A {later x} {now y} (↓≈ a≡b (later↓ x↓a) (now↓ y≡b)) = ≈-trans A (≈-sym A later-self) (μ↓ (≡↓ (Delayₛ A) (≈-trans A a≡b (≈-sym A y≡b)) x↓a))
  μ-cong A {later x} {later y} (↓≈ a≡b (later↓ x↓a) (later↓ y↓b)) = later≈ (μ-cong A (≈→≈′ (Delayₛ A) (↓≈ a≡b x↓a y↓b)))
  μ-cong A {later x} {later y} (later≈ x≈y) = later≈ (μ-cong A x≈y)
  force≈ (μ-cong A {x} {y} x≈y) = μ-cong A (force≈ x≈y)

  μₛ :  (A : Setoid c (c  ))  Delayₛ (Delayₛ A)  Delayₛ A
  μₛ A = record { to = μ {A} ; cong = μ-cong A }

  μₛ∼ :  (A : Setoid c (c  ))  Delayₛ (Delayₛ A)  Delayₛ A
  μₛ∼ A = record { to = μ {A} ; cong = μ-cong A }
    where
      μ-cong :  (A : Setoid c (c  )) {x y : Delay (Delay  A )}  [ Delayₛ A ][ x  y ]  [ A ][ μ {A} x  μ {A} y ]
      μ-cong :  (A : Setoid c (c  )) {x y : Delay (Delay  A )}  [ Delayₛ A ][ x  y ]  [ A ][ μ {A} x  μ {A} y ]
      force (μ-cong A {x} {y} xy) = μ-cong A (force xy)
      μ-cong A {.(now _)} {.(now _)} (now xy) = xy
      μ-cong A {.(later _)} {.(later _)} (later xy) = later (μ-cong A xy)

  μ-natural :  {A B : Setoid c (c  )} (f : A  B)  (μₛ B  liftFₛ (liftFₛ f))  (liftFₛ f  μₛ A)
  μ-natural :  {A B : Setoid c (c  )} (f : A  B)   {x : Delay (Delay  A )}  [ B ][ (μₛ B  liftFₛ (liftFₛ f)) ⟨$⟩ x ≈′ (liftFₛ f  μₛ A) ⟨$⟩ x ]
  force≈ (μ-natural {A} {B} f {x}) = μ-natural f {x}
  μ-natural {A} {B} f {now x} = ≈-refl B
  μ-natural {A} {B} f {later x} = later≈ (μ-natural f {force x})

  μ-assoc' :  {A : Setoid c (c  )} {x : Delay (Delay (Delay  A ))}  [ A ][ (μₛ A  (liftFₛ (μₛ A))) ⟨$⟩ x  (μₛ A  μₛ (Delayₛ A)) ⟨$⟩ x ]
  μ-assoc' :  {A : Setoid c (c  )} {x : Delay (Delay (Delay  A ))}  [ A ][ (μₛ A  (liftFₛ (μₛ A))) ⟨$⟩ x  (μₛ A  μₛ (Delayₛ A)) ⟨$⟩ x ]
  force (μ-assoc' {A} {x}) = μ-assoc' {A} {x}
  μ-assoc' {A} {now x} = -refl A
  μ-assoc' {A} {later x} = later (μ-assoc' {A} {force x})

  μ-assoc :  {A : Setoid c (c  )}  (μₛ A  (liftFₛ (μₛ A)))  (μₛ A  μₛ (Delayₛ A))
  μ-assoc {A} {x} = ∼⇒≈ (μ-assoc' {A} {x})

  identityˡ↓ :  {A : Setoid c (c  )} {x : Delay  A } {y :  A }  [ A ][ x  y ]  [ A ][ ((μₛ A) ⟨$⟩ ((liftFₛ (ηₛ A)) ⟨$⟩ x))  y ]
  identityˡ↓ {A} {now x} {y} x↓y = x↓y
  identityˡ↓ {A} {later x} {y} (later↓ x↓y) = later↓ (identityˡ↓ x↓y)

  identityˡ :  {A : Setoid c (c  )} {x : Delay  A }  [ A ][ (μₛ A  liftFₛ (ηₛ A)) ⟨$⟩ x  x ]
  identityˡ :  {A : Setoid c (c  )} {x : Delay  A }  [ A ][ (μₛ A  liftFₛ (ηₛ A)) ⟨$⟩ x  x ]
  identityˡ {A} {now x} = -refl A
  identityˡ {A} {later x} = later identityˡ
  force (identityˡ {A} {x}) = identityˡ

  identityˡ :  {A : Setoid c (c  )}  (μₛ A  liftFₛ (ηₛ A))  idₛ (Delayₛ A)
  identityˡ :  {A : Setoid c (c  )} {x : Delay  A }  [ A ][ (μₛ A  liftFₛ (ηₛ A)) ⟨$⟩ x ≈′ x ]
  force≈ (identityˡ {A} {x}) = identityˡ
  identityˡ {A} {now x} = ≈-refl A
  identityˡ {A} {later x} = later≈ identityˡ

  identityʳ :  {A : Setoid c (c  )}  (μₛ A  ηₛ (Delayₛ A))  idₛ (Delayₛ A)
  identityʳ {A} {now x} = ≈-refl A
  identityʳ {A} {later x} = ≈-refl A

  delayMonad : Monad (Setoids c (c  ))
  delayMonad = record
    { F = record
      { F₀ = Delayₛ
      ; F₁ = liftFₛ
      ; identity = lift-id
      ; homomorphism = λ {X} {Y} {Z} {f} {g}  lift-comp {X} {Y} {Z} {f} {g}
      ; F-resp-≈ = λ {A} {B} {f} {g}  lift-resp-≈ {A} {B} {f} {g}
      }
    ; η = ntHelper (record { η = ηₛ ; commute = η-natural })
    ; μ = ntHelper (record { η = μₛ ; commute = μ-natural })
    ; assoc = λ {A} {x}  μ-assoc {A} {x}
    ; sym-assoc = λ {A} {x}  ≈-sym A (μ-assoc {A} {x})
    ; identityˡ = identityˡ
    ; identityʳ = identityʳ
    }

open DelayMonad

module extra {A : Setoid c (c  )} where
  ≲→≈  : {x y : Delay  A }   [ A ][ x  y ]  [ A ][ x  y ]
  ≲→≈′ : {x y : Delay  A }  [ A ][ x ≲′ y ]  [ A ][ x ≈′ y ]
  ≲→≈ (↓≲ y↓a)     = ↓≈ (≡-refl A) (now↓ (≡-refl A)) y↓a
  ≲→≈ (later≲ x≈y) = later≈ (≲→≈′ x≈y)
  force≈ (≲→≈′ x≲y) = ≲→≈ (force≲ x≲y)

  race : Delay  A   Delay  A   Delay  A 
  race : Delay  A   Delay  A   Delay  A 
  race (now a) _ = now a
  race (later x) (now a) = now a
  race (later x) (later y) = later (race x y)
  force (race x y) = race (force x) (force y)

  race-sym :  {x y}  [ A ][ x  y ]  [ A ][ race x y  race y x ]
  race-sym :  {x y}  [ A ][ x  y ]  [ A ][ race x y  race y x ]
  force (race-sym {x} {y} xy) = race-sym {x} {y} (force xy)
  race-sym {now x} {now y} xy = xy
  race-sym {now x} {later y} xy = -refl A
  race-sym {later x} {now y} xy = -refl A
  race-sym {later x} {later y} (later xy) = later (race-sym xy)

  race-sym≈ :  {x y}  [ A ][ x  y ]  [ A ][ race x y  race y x ]
  race-sym≈ :  {x y}  [ A ][ x ≈′ y ]  [ A ][ race x y  race y x ]
  force (race-sym≈ {x} {y} xy) = race-sym≈ {x} {y} (force≈ xy)
  race-sym≈ {now x} {now y} (↓≈ a≡b (now↓ x≡a) (now↓ y≡b)) = now (≡-trans A x≡a (≡-trans A a≡b (≡-sym A y≡b)))
  race-sym≈ {now x} {later y} x≈y = -refl A
  race-sym≈ {later x} {now y} x≈y = -refl A
  race-sym≈ {later x} {later y} (↓≈ a≡b (later↓ x↓a) (later↓ y↓b)) = later (race-sym≈ (≈→≈′ A (↓≈ a≡b x↓a y↓b)))
  race-sym≈ {later x} {later y} (later≈ x≈y) = later (race-sym≈ x≈y)

  ≈→≲₀ :  {x y a b} (x↓a : [ A ][ x  a ]) (y↓b : [ A ][ y  b ]) (a≡b : [ A ][ a  b ])  [ A ][ race x y  y ]
  ≈→≲₀ (now↓ x≡a) y↓b a≡b = ↓≲ (≡↓ A (≡-sym A (≡-trans A x≡a a≡b)) y↓b)
  ≈→≲₀ (later↓ x↓a) (now↓ x≡y) a≡b = ↓≲ (now↓ (≡-refl A))
  ≈→≲₀ (later↓ x↓a) (later↓ y↓b) a≡b = later≲ (record { force≲ = ≈→≲₀ x↓a y↓b a≡b })

  ≈→≲ : {x y : Delay  A }  [ A ][ x  y ]  [ A ][ race x y  y ]
  ≈′→≲′ : {x y : Delay  A }  [ A ][ x ≈′ y ]  [ A ][ race x y ≲′ y ]
  ≈→≲ (↓≈ a≡b x↓a y↓b) = ≈→≲₀ x↓a y↓b a≡b
  ≈→≲ (later≈ x≈y) = later≲ (≈′→≲′ x≈y)
  force≲ (≈′→≲′ x≈y) = ≈→≲ (force≈ x≈y)

  delta₀  : {x : Delay  A } {a :  A }   (x↓a : [ A ][ x  a ])   {c}
  delta₀ {x} (now↓ x≡y) = zero
  delta₀ (later↓ x↓a)   = suc (delta₀ x↓a)

  delta  : {x y : Delay  A }  [ A ][ x  y ]  Delay ( A  ×  {c})
  delta : {x y : Delay  A }  [ A ][ x ≲′ y ]  Delay ( A  ×  {c})
  delta (↓≲ {x}{a} x↓a) = now (a , delta₀ x↓a)
  delta (later≲ x≲y)   = later (delta x≲y)
  force (delta x≲y) = delta (force≲ x≲y)

  ι :  A  ×  {c}  Delay  A 
  ι :  A  ×  {c}  Delay  A 
  force (ι p) = ι p
  ι (x , zero)  = now x
  ι (x , suc n) = later (ι (x , n))

  ιₛ' : (A ×ₛ (-setoid {c}))  Delayₛ A
  ιₛ' = record { to = ι ; cong = ι-cong }
    where
      ι-cong :  {x y}  [ A ×ₛ (-setoid {c}) ][ x  y ]  [ A ][ ι x  ι y ]
      ι-cong :  {x y}  [ A ×ₛ (-setoid {c}) ][ x  y ]  [ A ][ ι x  ι y ]
      force (ι-cong {x} {y} x≡y) = ι-cong x≡y
      ι-cong {x , zero} {y , zero} (x≡y , n≡m) = now x≡y
      ι-cong {x , suc n} {y , suc m} (x≡y , n≡m) = later (ι-cong (x≡y , suc-inj n≡m))

  delta-prop₁  : {x y : Delay  A } (x≲y : [ A ][ x  y ])  [ A ][ liftF proj₁ (delta x≲y)  x ]
  delta-prop : {x y : Delay  A } (x≲y : [ A ][ x ≲′ y ])  [ A ][ liftF proj₁ (force (delta x≲y))  x ]
  delta-prop₁ {.(now _)} {x} (↓≲ x↓a) = now (≡-refl A)
  delta-prop₁ {.(later _)} {.(later _)} (later≲ x≲y) = later (delta-prop x≲y)
  force (delta-prop x≲y) = delta-prop₁ (force≲ x≲y)

  delta-prop₂ : {x y : Delay  A } (x≲y : [ A ][ x  y ])  [ A ][ μ {A} (liftF ι (delta x≲y))  y ]
  delta-prop : {x y : Delay  A } (x≲y : [ A ][ x ≲′ y ])  [ A ][ μ {A} (liftF ι (force (delta x≲y)))  y ]
  delta-prop₂ (↓≲ x↓a)     = -sym A (ι↓ x↓a)
    where
      ι↓ :  {x : Delay  A }{a :  A }  (x↓a : [ A ][ x  a ])  [ A ][ x  ι (a , delta₀ x↓a) ]
      ι↓ {.(now _)} {a} (now↓ x≡y)     = now x≡y
      ι↓ {.(later _)} {a} (later↓ x↓a) = later (record { force = ι↓ x↓a })
  delta-prop₂ (later≲ x≲y) = later (delta-prop x≲y)
  force (delta-prop x≲y) = delta-prop₂ (force≲ x≲y)