# The delay monad on the category of setoids is an instance of K
module Monad.Instance.Setoids.K {ℓ : Level} where open _⟶_ using (cong) open import Category.Ambient.Setoids open Ambient (setoidAmbient {ℓ}) using (cocartesian; distributive) open import Monad.Instance.Setoids.Delay {ℓ} {ℓ} open import Monad.Instance.K (setoidAmbient {ℓ}) open import Algebra.Elgot cocartesian open import Algebra.Elgot.Free cocartesian open import Algebra.Elgot.Stable distributive open import Category.Construction.ElgotAlgebras {C = Setoids ℓ ℓ} cocartesian open import Monad.PreElgot (setoidAmbient {ℓ}) open Bisimilarity renaming (_≈_ to [_][_≈_]; _≈′_ to [_][_≈′_]; _∼_ to [_][_∼_]; _∼′_ to [_][_∼′_]; _↓_ to [_][_↓_]) open DelayMonad open extra open Setoid using () renaming (Carrier to ∣_∣; _≈_ to [_][_≡_]) open eq using () renaming (refl to ≡-refl; sym to ≡-sym; trans to ≡-trans) ≡→≡ : ∀ {A : Setoid ℓ ℓ} {x y : ∣ A ∣} → x ≡ y → [ A ][ x ≡ y ] ≡→≡ {A} {x} {y} eq rewrite eq = ≡-refl A iter : ∀ {A X : Setoid ℓ ℓ} → (∣ X ∣ → (Delay ∣ A ∣ ⊎ ∣ X ∣)) → ∣ X ∣ → Delay ∣ A ∣ iter′ : ∀ {A X : Setoid ℓ ℓ} → (∣ X ∣ → (Delay ∣ A ∣ ⊎ ∣ X ∣)) → ∣ X ∣ → Delay′ ∣ A ∣ force (iter′ {A} {X} f x) = iter {A} {X} f x iter {A} {X} f x with f x ... | inj₁ a = a ... | inj₂ b = later (iter′ {A} {X} f b) conflict : ∀ {ℓ''} (X Y : Setoid ℓ ℓ) {Z : Set ℓ''} {x : ∣ X ∣} {y : ∣ Y ∣} → [ X ⊎ₛ Y ][ inj₁ x ≡ inj₂ y ] → Z conflict X Y () inj₁-helper : ∀ {X Y : Setoid ℓ ℓ} (f : X ⟶ (Y ⊎ₛ X)) {x y : ∣ X ∣} {a b : ∣ Y ∣} → [ X ][ x ≡ y ] → f ⟨$⟩ x ≡ inj₁ a → f ⟨$⟩ y ≡ inj₁ b → [ Y ][ a ≡ b ] inj₁-helper {X} {Y} f {x} {y} {a} {b} x≡y fi₁ fi₂ = drop-inj₁ {x = a} {y = b} helper where helper : [ Y ⊎ₛ X ][ inj₁ a ≡ inj₁ b ] helper rewrite (≣-sym fi₁) | (≣-sym fi₂) = cong f x≡y inj₂-helper : ∀ {X Y : Setoid ℓ ℓ} (f : X ⟶ (Y ⊎ₛ X)) {x y : ∣ X ∣} {a b : ∣ X ∣} → [ X ][ x ≡ y ] → f ⟨$⟩ x ≡ inj₂ a → f ⟨$⟩ y ≡ inj₂ b → [ X ][ a ≡ b ] inj₂-helper {X} {Y} f {x} {y} {a} {b} x≡y fi₁ fi₂ = drop-inj₂ {x = a} {y = b} helper where helper : [ Y ⊎ₛ X ][ inj₂ a ≡ inj₂ b ] helper rewrite (≣-sym fi₁) | (≣-sym fi₂) = cong f x≡y absurd-helper : ∀ {ℓ'} {X Y : Setoid ℓ ℓ} {A : Set ℓ'} (f : X ⟶ (Y ⊎ₛ X)) {x y : ∣ X ∣} {a : ∣ Y ∣} {b : ∣ X ∣} → [ X ][ x ≡ y ] → f ⟨$⟩ x ≡ inj₁ a → f ⟨$⟩ y ≡ inj₂ b → A absurd-helper {ℓ'} {X} {Y} {A} f {x} {y} {a} {b} x≡y fi₁ fi₂ = conflict Y X helper where helper : [ Y ⊎ₛ X ][ inj₁ a ≡ inj₂ b ] helper rewrite (≣-sym fi₁) | (≣-sym fi₂) = cong f x≡y iter-cong∼ : ∀ {A X : Setoid ℓ ℓ} (f : X ⟶ (Delayₛ∼ A ⊎ₛ X)) {x y : ∣ X ∣} → [ X ][ x ≡ y ] → [ A ][ (iter {A} {X} < f > x) ∼ (iter {A} {X} < f > y) ] iter-cong∼′ : ∀ {A X : Setoid ℓ ℓ} (f : X ⟶ (Delayₛ∼ A ⊎ₛ X)) {x y : ∣ X ∣} → [ X ][ x ≡ y ] → [ A ][ (iter {A} {X} < f > x) ∼′ (iter {A} {X} < f > y) ] force∼ (iter-cong∼′ {A} {X} f {x} {y} x≡y) = iter-cong∼ f x≡y iter-cong∼ {A} {X} f {x} {y} x≡y with < f > x in eqx | < f > y in eqy ... | inj₁ a | inj₁ b = inj₁-helper f x≡y eqx eqy ... | inj₁ a | inj₂ b = absurd-helper f x≡y eqx eqy ... | inj₂ a | inj₁ b = absurd-helper f (≡-sym X x≡y) eqy eqx ... | inj₂ a | inj₂ b = later∼ (iter-cong∼′ {A} {X} f (inj₂-helper f x≡y eqx eqy)) iter-cong : ∀ {A X : Setoid ℓ ℓ} (f : X ⟶ (Delayₛ A ⊎ₛ X)) {x y : ∣ X ∣} → [ X ][ x ≡ y ] → [ A ][ (iter {A} {X} < f > x) ≈ (iter {A} {X} < f > y) ] iter-cong′ : ∀ {A X : Setoid ℓ ℓ} (f : X ⟶ (Delayₛ A ⊎ₛ X)) {x y : ∣ X ∣} → [ X ][ x ≡ y ] → [ A ][ (iter {A} {X} < f > x) ≈′ (iter {A} {X} < f > y) ] force≈ (iter-cong′ {A} {X} f {x} {y} x≡y) = iter-cong f x≡y iter-cong {A} {X} f {x} {y} x≡y with < f > x in eqx | < f > y in eqy ... | inj₁ a | inj₁ b = inj₁-helper f x≡y eqx eqy ... | inj₁ a | inj₂ b = absurd-helper f x≡y eqx eqy ... | inj₂ a | inj₁ b = absurd-helper f (≡-sym X x≡y) eqy eqx ... | inj₂ a | inj₂ b = later≈ (iter-cong′ {A} {X} f (inj₂-helper f x≡y eqx eqy)) iterₛ∼ : ∀ {A X : Setoid ℓ ℓ} → (X ⟶ (Delayₛ∼ A ⊎ₛ X)) → X ⟶ Delayₛ∼ A iterₛ∼ {A} {X} f = record { to = iter {A} {X} < f > ; cong = iter-cong∼ {A} {X} f } iterₛ : ∀ {A X : Setoid ℓ ℓ} → (X ⟶ (Delayₛ A ⊎ₛ X)) → X ⟶ Delayₛ A iterₛ {A} {X} f = record { to = iter {A} {X} < f > ; cong = iter-cong {A} {X} f } iter-fixpoint : ∀ {A X : Setoid ℓ ℓ} {f : X ⟶ (Delayₛ A ⊎ₛ X)} {x : ∣ X ∣} → [ A ][ iter {A} {X} < f > x ≈ [ Function.id , iter {A} {X} < f > ] (f ⟨$⟩ x) ] iter-fixpoint {A} {X} {f} {x} with < f > x in eqx ... | inj₁ a = ≈-refl A ... | inj₂ a = ≈-trans A (≈-sym A later-self) (≈-refl A) iter-resp-≈ : ∀ {A X : Setoid ℓ ℓ} (f g : X ⟶ (Delayₛ A ⊎ₛ X)) → f ≋ g → ∀ {x y : ∣ X ∣} → [ X ][ x ≡ y ] → [ A ][ iter {A} {X} < f > x ≈ iter {A} {X} < g > y ] iter-resp-≈′ : ∀ {A X : Setoid ℓ ℓ} (f g : X ⟶ (Delayₛ A ⊎ₛ X)) → f ≋ g → ∀ {x y : ∣ X ∣} → [ X ][ x ≡ y ] → [ A ][ iter {A} {X} < f > x ≈′ iter {A} {X} < g > y ] force≈ (iter-resp-≈′ {A} {X} f g f≋g {x} {y} x≡y) = iter-resp-≈ f g f≋g {x} {y} x≡y iter-resp-≈ {A} {X} f g f≋g {x} {y} x≡y with < f > x in eqa | < g > y in eqb ... | inj₁ a | inj₁ b = drop-inj₁ helper where helper : [ Delayₛ A ⊎ₛ X ][ inj₁ a ≡ inj₁ b ] helper rewrite (≣-sym eqb) | (≣-sym eqa) = ≡-trans (Delayₛ A ⊎ₛ X) (cong f x≡y) f≋g ... | inj₁ a | inj₂ b = conflict (Delayₛ A) X helper where helper : [ Delayₛ A ⊎ₛ X ][ inj₁ a ≡ inj₂ b ] helper rewrite (≣-sym eqb) | (≣-sym eqa) = ≡-trans (Delayₛ A ⊎ₛ X) (cong f x≡y) f≋g ... | inj₂ a | inj₁ b = conflict (Delayₛ A) X (≡-sym (Delayₛ A ⊎ₛ X) helper) where helper : [ Delayₛ A ⊎ₛ X ][ inj₂ a ≡ inj₁ b ] helper rewrite (≣-sym eqb) | (≣-sym eqa) = ≡-trans (Delayₛ A ⊎ₛ X) (cong f x≡y) f≋g ... | inj₂ a | inj₂ b = later≈ (iter-resp-≈′ f g f≋g (drop-inj₂ helper)) where helper : [ Delayₛ A ⊎ₛ X ][ inj₂ a ≡ inj₂ b ] helper rewrite (≣-sym eqb) | (≣-sym eqa) = ≡-trans (Delayₛ A ⊎ₛ X) (cong f x≡y) f≋g iter-uni : ∀ {A X Y : Setoid ℓ ℓ} {f : X ⟶ (Delayₛ A ⊎ₛ X)} {g : Y ⟶ (Delayₛ A ⊎ₛ Y)} {h : X ⟶ Y} → ([ inj₁ₛ ∘ (idₛ (Delayₛ A)) , inj₂ₛ ∘ h ]ₛ ∘ f) ≋ (g ∘ h) → ∀ {x : ∣ X ∣} {y : ∣ Y ∣} → [ Y ][ y ≡ h ⟨$⟩ x ] → [ A ][ iter {A} {X} < f > x ≈ (iter {A} {Y} < g >) y ] iter-uni′ : ∀ {A X Y : Setoid ℓ ℓ} {f : X ⟶ (Delayₛ A ⊎ₛ X)} {g : Y ⟶ (Delayₛ A ⊎ₛ Y)} {h : X ⟶ Y} → ([ inj₁ₛ ∘ (idₛ (Delayₛ A)) , inj₂ₛ ∘ h ]ₛ ∘ f) ≋ (g ∘ h) → ∀ {x : ∣ X ∣} {y : ∣ Y ∣} → [ Y ][ y ≡ h ⟨$⟩ x ] → [ A ][ iter {A} {X} < f > x ≈′ (iter {A} {Y} < g >) y ] force≈ (iter-uni′ {A} {X} {Y} {f} {g} {h} eq {x} {y} y≡h$x) = iter-uni {A} {X} {Y} {f} {g} {h} eq {x} {y} y≡h$x iter-uni {A} {X} {Y} {f} {g} {h} eq {x} {y} x≡y with f ⟨$⟩ x in eqx | g ⟨$⟩ (h ⟨$⟩ x) in eqy | g ⟨$⟩ y in eqz | eq {x} ... | inj₁ a | inj₁ b | inj₁ c | inj₁ a≈c = ≈-trans A a≈c (inj₁-helper g (≡-sym Y x≡y) eqy eqz) ... | inj₁ a | inj₁ b | inj₂ c | inj₁ _ = absurd-helper g (≡-sym Y x≡y) eqy eqz ... | inj₂ a | inj₂ b | inj₁ c | inj₂ _ = absurd-helper g x≡y eqz eqy ... | inj₂ a | inj₂ b | inj₂ c | inj₂ req = later≈ (iter-uni′ {f = f} {g = g}{h = h} eq c≡h$a) where c≡h$a : [ Y ][ c ≡ h ⟨$⟩ a ] c≡h$a = ≡-trans Y (drop-inj₂ (≡-trans (Delayₛ A ⊎ₛ Y) (≡-trans (Delayₛ A ⊎ₛ Y) (≡-sym (Delayₛ A ⊎ₛ Y) (≡→≡ {Delayₛ A ⊎ₛ Y} eqz)) (cong g x≡y)) (≡→≡ {Delayₛ A ⊎ₛ Y} eqy))) (≡-sym Y req) iter-folding : ∀ {A X Y : Setoid ℓ ℓ} {f : X ⟶ (Delayₛ A ⊎ₛ X)} {h : Y ⟶ (X ⊎ₛ Y)} {x : ∣ X ⊎ₛ Y ∣} → [ A ][ iter {A} {X ⊎ₛ Y} [ inj₁ ∘f iter {A} {X} < f > , inj₂ ∘f < h > ] x ≈ iter {A} {X ⊎ₛ Y} [ [ inj₁ , inj₂ ∘′ inj₁ ] ∘f < f > , (inj₂ ∘f < h >) ] x ] iter-folding′ : ∀ {A X Y : Setoid ℓ ℓ} {f : X ⟶ (Delayₛ A ⊎ₛ X)} {h : Y ⟶ (X ⊎ₛ Y)} {x : ∣ X ⊎ₛ Y ∣} → [ A ][ iter {A} {X ⊎ₛ Y} [ inj₁ ∘f iter {A} {X} < f > , inj₂ ∘f < h > ] x ≈′ iter {A} {X ⊎ₛ Y} [ [ inj₁ , inj₂ ∘′ inj₁ ] ∘f < f > , (inj₂ ∘f < h >) ] x ] force≈ (iter-folding′ {A} {X} {Y} {f} {h} {x}) = iter-folding {A} {X} {Y} {f} {h} {x} iter-folding {A} {X} {Y} {f} {h} {inj₁ x} with f ⟨$⟩ x in eqa ... | inj₁ a = ≈-refl A ... | inj₂ a = later≈ (helper′ a) where helper : ∀ (b : ∣ X ∣) → [ A ][ iter < f > b ≈ iter [ [ inj₁ , inj₂ ∘′ inj₁ ] ∘′ < f > , inj₂ ∘′ < h > ] (inj₁ b) ] helper′ : ∀ (b : ∣ X ∣) → [ A ][ iter < f > b ≈′ iter [ [ inj₁ , inj₂ ∘′ inj₁ ] ∘′ < f > , inj₂ ∘′ < h > ] (inj₁ b) ] helper b with f ⟨$⟩ b in eqb ... | inj₁ c = ≈-refl A ... | inj₂ c = later≈ (helper′ c) force≈ (helper′ b) = helper b iter-folding {A} {X} {Y} {f} {h} {inj₂ x} with h ⟨$⟩ x in eqa ... | inj₁ a = later≈ (iter-folding′ {A} {X} {Y} {f} {h} {inj₁ a}) ... | inj₂ a = later≈ (iter-folding′ {A} {X} {Y} {f} {h} {inj₂ a}) delay-algebras-on : ∀ {A : Setoid ℓ ℓ} → Elgot-Algebra-on (Delayₛ A) delay-algebras-on {A} = record { _# = iterₛ {A} ; #-Fixpoint = λ {X} {f} → iter-fixpoint {A} {X} {f} ; #-Uniformity = λ {X} {Y} {f} {g} {h} eq {x} → iter-uni {A} {X} {Y} {f} {g} {h} eq {x} {h ⟨$⟩ x} (≡-refl Y) ; #-Folding = λ {X} {Y} {f} {h} {x} → iter-folding {A} {X} {Y} {f} {h} {x} ; #-resp-≈ = λ {X} {f} {g} f≋g {x} → iter-resp-≈ {A} {X} f g f≋g {x} {x} (≡-refl X) } delay-algebras : ∀ (A : Setoid ℓ ℓ) → Elgot-Algebra delay-algebras A = record { A = Delayₛ A ; algebra = delay-algebras-on {A}} open Elgot-Algebra using (#-Fixpoint; #-Uniformity; #-Compositionality; #-resp-≈; #-Diamond) renaming (A to ⟦_⟧) delay-lift : ∀ {A : Setoid ℓ ℓ} {B : Elgot-Algebra} → A ⟶ ⟦ B ⟧ → Elgot-Algebra-Morphism (delay-algebras A) B delay-lift {A} {B} f = record { h = delay-lift' ; preserves = λ {X} {g} {x} → preserves' {X} {g} {x} } where open Elgot-Algebra B using (_#) -- (f + id) ∘ out helper₁ : Delay ∣ A ∣ → ∣ ⟦ B ⟧ ∣ ⊎ Delay ∣ A ∣ helper₁ (now x) = inj₁ (< f > x) helper₁ (later x) = inj₂ (force x) helper₁-cong : {x y : Delay ∣ A ∣} → (x∼y : [ A ][ x ∼ y ]) → [ ⟦ B ⟧ ⊎ₛ Delayₛ∼ A ][ helper₁ x ≡ helper₁ y ] helper₁-cong (now∼ x≡y) = inj₁ (cong f x≡y) helper₁-cong (later∼ x≡y) = inj₂ (force∼ x≡y) -- -- setoid-morphism that preserves strong-bisimilarity helper : (Delayₛ∼ A) ⟶ (⟦ B ⟧ ⊎ₛ Delayₛ∼ A) helper = record { to = helper₁ ; cong = helper₁-cong} helper#∼-cong : {x y : Delay ∣ A ∣} → (x∼y : [ A ][ x ∼ y ]) → [ ⟦ B ⟧ ][ helper # ⟨$⟩ x ≡ helper # ⟨$⟩ y ] helper#∼-cong {x} {y} x∼y = cong (helper #) x∼y helper#≈-cong : {x y : Delay ∣ A ∣} → (x≈y : [ A ][ x ≈ y ]) → [ ⟦ B ⟧ ][ helper # ⟨$⟩ x ≡ helper # ⟨$⟩ y ] -- key special case helper#≈-cong' : {z : Delay (∣ A ∣ × ℕ)} → [ ⟦ B ⟧ ][ helper # ⟨$⟩ liftF proj₁ z ≡ helper # ⟨$⟩ μ {A} (liftF (ι {A}) z) ] helper#≈-cong x≈y = ≡-trans ⟦ B ⟧ (helper#∼-cong (∼-sym A (delta-prop₂ {A} ineq₂))) (≡-trans ⟦ B ⟧ (≡-trans ⟦ B ⟧ (≡-sym ⟦ B ⟧ (helper#≈-cong' {z₂})) (≡-trans ⟦ B ⟧ (helper#∼-cong (∼-trans A (delta-prop₁ (≈→≲ (≈-sym A x≈y))) (∼-sym A (∼-trans A (delta-prop₁ (≈→≲ x≈y)) (race-sym≈ x≈y))))) (helper#≈-cong' {z₁}))) (helper#∼-cong (delta-prop₂ {A} ineq₁))) where ineq₁ = ≈→≲ {A} x≈y ineq₂ = ≈→≲ {A} (≈-sym A x≈y) z₁ = delta {A} ineq₁ z₂ = delta {A} ineq₂ helper#≈-cong' {z} = ≡-trans ⟦ B ⟧ (≡-trans ⟦ B ⟧ (≡-sym ⟦ B ⟧ eq₁) (≡-sym ⟦ B ⟧ eq₀)) eq₂ where outer : A ⟶ (A ×ₛ ℕ-setoid {ℓ}) outer = record { to = λ z → z , zero ; cong = λ {x} {y} x≡y → x≡y , ≣-refl } ι-cancel : ∀ {x} → [ A ][ (ι {A} ∘′ (λ z → z , zero)) x ∼ now x ] ι-cancel = ∼-refl A helper₁' : Delay (∣ A ∣ × ℕ {ℓ}) → ∣ ⟦ B ⟧ ∣ ⊎ Delay (∣ A ∣ × ℕ {ℓ}) helper₁' (now (x , zero)) = inj₁ (< f > x) helper₁' (now (x , suc n)) = inj₂ (< liftFₛ∼ outer > (ι {A} (x , n))) helper₁' (later y) = inj₂ (force y) helper₁-cong' : {x y : Delay (∣ A ∣ × ℕ {ℓ})} → (x∼y : [ A ×ₛ ℕ-setoid ][ x ∼ y ]) → [ ⟦ B ⟧ ⊎ₛ Delayₛ∼ (A ×ₛ ℕ-setoid) ][ helper₁' x ≡ helper₁' y ] helper₁-cong' {now (x , zero)} (now∼ (x≡y , ≣-refl)) = inj₁ (cong f x≡y) helper₁-cong' {now (x , suc n)} {now (y , suc _)} (now∼ (x≡y , ≣-refl)) = inj₂ (cong (liftFₛ∼ outer) (cong ιₛ' (x≡y , ≣-refl))) helper₁-cong' (later∼ x∼y) = inj₂ (force∼ x∼y) helper' : (Delayₛ∼ (A ×ₛ ℕ-setoid)) ⟶ (⟦ B ⟧ ⊎ₛ Delayₛ∼ (A ×ₛ ℕ-setoid)) helper' = record { to = helper₁' ; cong = helper₁-cong'} helper₁'' : Delay (∣ A ∣ × ℕ {ℓ}) → ∣ ⟦ B ⟧ ∣ ⊎ Delay (∣ A ∣ × ℕ {ℓ}) helper₁'' (now (x , _)) = inj₁ (< f > x) helper₁'' (later y) = inj₂ (force y) helper₁-cong'' : {x y : Delay (∣ A ∣ × ℕ {ℓ})} → (x∼y : [ A ×ₛ ℕ-setoid ][ x ∼ y ]) → [ ⟦ B ⟧ ⊎ₛ Delayₛ∼ (A ×ₛ ℕ-setoid) ][ helper₁'' x ≡ helper₁'' y ] helper₁-cong'' {now (x , _)} (now∼ (x≡y , ≣-refl)) = inj₁ (cong f x≡y) helper₁-cong'' (later∼ x∼y) = inj₂ (force∼ x∼y) helper'' : (Delayₛ∼ (A ×ₛ ℕ-setoid)) ⟶ (⟦ B ⟧ ⊎ₛ Delayₛ∼ (A ×ₛ ℕ-setoid)) helper'' = record { to = helper₁'' ; cong = helper₁-cong''} -- Needs #-Diamond eq₀ : ∀ {z} → [ ⟦ B ⟧ ][ helper' # ⟨$⟩ z ≡ helper'' # ⟨$⟩ z ] eq₀ {z} = ≡-trans ⟦ B ⟧ (#-resp-≈ B {Delayₛ∼ (A ×ₛ ℕ-setoid)} {helper'} step₁) (≡-trans ⟦ B ⟧ (#-Diamond B {Delayₛ∼ (A ×ₛ ℕ-setoid)} helper''') (#-resp-≈ B (λ {x} → (step₂ {x})))) where helper₁''' : Delay (∣ A ∣ × ℕ {ℓ}) → ∣ ⟦ B ⟧ ∣ ⊎ (Delay (∣ A ∣ × ℕ {ℓ}) ⊎ Delay (∣ A ∣ × ℕ {ℓ})) helper₁''' (now (x , zero)) = inj₁ (< f > x) helper₁''' (now (x , suc n)) = inj₂ (inj₁ (< liftFₛ∼ outer > (ι {A} (x , n)))) helper₁''' (later y) = inj₂ (inj₂ (force y)) helper₁-cong''' : {x y : Delay (∣ A ∣ × ℕ {ℓ})} → (x∼y : [ A ×ₛ ℕ-setoid ][ x ∼ y ]) → [ ⟦ B ⟧ ⊎ₛ (Delayₛ∼ (A ×ₛ ℕ-setoid) ⊎ₛ Delayₛ∼ (A ×ₛ ℕ-setoid)) ][ helper₁''' x ≡ helper₁''' y ] helper₁-cong''' {now (x , zero)} (now∼ (x≡y , ≣-refl)) = inj₁ (cong f x≡y) helper₁-cong''' {now (x , suc n)} {now (y , suc _)} (now∼ (x≡y , ≣-refl)) = inj₂ (inj₁ (cong (liftFₛ∼ outer) (cong ιₛ' (x≡y , ≣-refl)))) helper₁-cong''' (later∼ x∼y) = inj₂ (inj₂ (force∼ x∼y)) helper''' : (Delayₛ∼ (A ×ₛ ℕ-setoid)) ⟶ (⟦ B ⟧ ⊎ₛ (Delayₛ∼ (A ×ₛ ℕ-setoid) ⊎ₛ Delayₛ∼ (A ×ₛ ℕ-setoid))) helper''' = record { to = helper₁''' ; cong = helper₁-cong''' } step₁ : ∀ {x} → [ ⟦ B ⟧ ⊎ₛ Delayₛ∼ (A ×ₛ ℕ-setoid) ][ helper' ⟨$⟩ x ≡ ([ inj₁ , inj₂ ∘′ [ id , id ] ] ∘′ helper₁''') x ] step₁ {now (a , zero)} = ≡-refl (⟦ B ⟧ ⊎ₛ Delayₛ∼ (A ×ₛ ℕ-setoid)) step₁ {now (a , suc n)} = ≡-refl (⟦ B ⟧ ⊎ₛ Delayₛ∼ (A ×ₛ ℕ-setoid)) step₁ {later x} = ≡-refl (⟦ B ⟧ ⊎ₛ Delayₛ∼ (A ×ₛ ℕ-setoid)) step₂ : ∀ {x} → [ ⟦ B ⟧ ⊎ₛ Delayₛ∼ (A ×ₛ ℕ-setoid) ][ ([ inj₁ , [ inj₁ ∘′ < ([ inj₁ₛ , inj₂ₛ ∘ [ idₛ (Delayₛ∼ (A ×ₛ ℕ-setoid)) , idₛ (Delayₛ∼ (A ×ₛ ℕ-setoid)) ]ₛ ]ₛ ∘ helper''') # > , inj₂ ] ] ∘′ helper₁''') x ≡ helper'' ⟨$⟩ x ] step₂ {now (a , zero)} = ≡-refl (⟦ B ⟧ ⊎ₛ Delayₛ∼ (A ×ₛ ℕ-setoid)) step₂ {now (x , suc n)} = inj₁ (by-induction n) where by-induction : ∀ n → [ ⟦ B ⟧ ][ < ([ inj₁ₛ , inj₂ₛ ∘ [ idₛ (Delayₛ∼ (A ×ₛ ℕ-setoid)) , idₛ (Delayₛ∼ (A ×ₛ ℕ-setoid)) ]ₛ ]ₛ ∘ helper''') # > (< liftFₛ∼ outer > (ι (x , n))) ≡ f ⟨$⟩ x ] by-induction zero = #-Fixpoint B by-induction (suc n) = ≡-trans ⟦ B ⟧ (#-Fixpoint B) (by-induction n) step₂ {later y} = ≡-refl (⟦ B ⟧ ⊎ₛ Delayₛ∼ (A ×ₛ ℕ-setoid)) eq₁ : ∀ {z} → [ ⟦ B ⟧ ][ helper'' # ⟨$⟩ z ≡ helper # ⟨$⟩ liftF proj₁ z ] eq₁ {z} = #-Uniformity B {f = helper''} {g = helper} {h = liftFₛ∼ proj₁ₛ} by-uni where by-uni : ∀ {x} → [ ⟦ B ⟧ ⊎ₛ Delayₛ∼ A ][ [ inj₁ , inj₂ ∘′ (liftF proj₁) ] (helper₁'' x) ≡ (< helper > ∘′ liftF proj₁) x ] by-uni {now (a , b)} = ≡-refl (⟦ B ⟧ ⊎ₛ Delayₛ∼ A) by-uni {later x} = ≡-refl (⟦ B ⟧ ⊎ₛ Delayₛ∼ A) eq : ∀ {x y} → [ A ×ₛ ℕ-setoid ][ x ∼ y ] → [ ⟦ B ⟧ ⊎ₛ Delayₛ∼ A ][ [ inj₁ , inj₂ ∘′ μ ∘′ (liftF ι) ] (helper₁' x) ≡ (helper₁ ∘′ μ ∘′ (liftF ι)) y ] eq {now (x , n)} {now (y , .n)} (now∼ (x∼y , ≣-refl)) = eq' {n} where eq' : ∀ {n} → [ ⟦ B ⟧ ⊎ₛ Delayₛ∼ A ][ [ inj₁ , inj₂ ∘′ μ ∘′ liftF ι ] (helper₁' (now (x , n))) ≡ (helper₁ ∘′ μ {A} ∘′ liftF ι) (now (y , n)) ] eq' {zero} = inj₁ (cong f x∼y) eq' {suc n} = inj₂ (∼-trans A (cong (μₛ∼ A) (∼-sym (Delayₛ∼ A) (lift-comp∼ {f = outer} {g = ιₛ'} {ι (x , n)} (∼-refl A)))) (∼-trans A identityˡ∼ (cong ιₛ' (x∼y , ≣-refl)))) eq (later∼ x∼y) = inj₂ (cong (μₛ∼ A) (cong (liftFₛ∼ ιₛ') (force∼ x∼y))) eq₂ : [ ⟦ B ⟧ ][ helper' # ⟨$⟩ z ≡ helper # ⟨$⟩ μ {A} (liftF (ι {A}) z)] eq₂ = Elgot-Algebra.#-Uniformity B {Delayₛ∼ (A ×ₛ ℕ-setoid {ℓ})} {Delayₛ∼ A} {helper'} {helper} {μₛ∼ A ∘ liftFₛ∼ ιₛ'} (λ {x} → eq {x} {x} (∼-refl (A ×ₛ ℕ-setoid))) delay-lift' = record { to = < helper # > ; cong = helper#≈-cong } -- discretize a setoid ‖_‖ : Setoid ℓ ℓ → Setoid ℓ ℓ ∣_∣ ‖ X ‖ = ∣ X ∣ [_][_≡_] ‖ X ‖ = _≡_ Setoid.isEquivalence ‖ X ‖ = Eq.isEquivalence ‖‖-quote : ∀ {X : Setoid ℓ ℓ} → ‖ X ‖ ⟶ X _⟶_.to ‖‖-quote x = x cong (‖‖-quote {X}) ≣-refl = ≡-refl X disc-dom : ∀ {X : Setoid ℓ ℓ} → X ⟶ (Delayₛ A ⊎ₛ X) → ‖ X ‖ ⟶ (Delayₛ∼ A ⊎ₛ ‖ X ‖) _⟶_.to (disc-dom f) x = f ⟨$⟩ x cong (disc-dom {X} f) {x} {y} x≡y rewrite x≡y = ≡-refl (Delayₛ∼ A ⊎ₛ ‖ X ‖) iter-g-var : ∀ {X : Setoid ℓ ℓ} → (g : X ⟶ (Delayₛ A ⊎ₛ X)) → ∀ {x} → [ A ][ (iter {A} {X} < g >) x ∼ (iterₛ∼ {A} {‖ X ‖} (disc-dom g)) ⟨$⟩ x ] iter-g-var′ : ∀ {X : Setoid ℓ ℓ} → (g : X ⟶ (Delayₛ A ⊎ₛ X)) → ∀ {x} → [ A ][ (iter {A} {X} < g >) x ∼′ (iterₛ∼ {A} {‖ X ‖} (disc-dom g)) ⟨$⟩ x ] force∼ (iter-g-var′ {X} g {x}) = iter-g-var {X} g {x} iter-g-var {X} g {x} with g ⟨$⟩ x ... | inj₁ a = ∼-refl A ... | inj₂ a = later∼ (iter-g-var′ {X} g {a}) preserves' : ∀ {X : Setoid ℓ ℓ} {g : X ⟶ (Delayₛ A ⊎ₛ X)} → ∀ {x} → [ ⟦ B ⟧ ][ (delay-lift' ∘ (iterₛ {A} {X} g)) ⟨$⟩ x ≡ ([ inj₁ₛ ∘ delay-lift' , inj₂ₛ ]ₛ ∘ g) # ⟨$⟩ x ] preserves' {X} {g} {x} = ≡-trans ⟦ B ⟧ step₁ step₂ where step₁ : [ ⟦ B ⟧ ][ (delay-lift' ∘ (iterₛ {A} {X} g)) ⟨$⟩ x ≡ ([ inj₁ₛ ∘ (helper #) , inj₂ₛ ]ₛ ∘ (disc-dom g)) # ⟨$⟩ x ] step₁ = ≡-trans ⟦ B ⟧ (≡-trans ⟦ B ⟧ (helper#∼-cong (iter-g-var g)) (sub-step₁ (disc-dom g) {inj₂ x})) (≡-sym ⟦ B ⟧ (#-Compositionality B {f = helper} {h = disc-dom g})) where sub-step₁ : (g : ‖ X ‖ ⟶ ((Delayₛ∼ A) ⊎ₛ ‖ X ‖)) → ∀ {x} → [ ⟦ B ⟧ ][ ((helper #) ∘ [ idₛ (Delayₛ∼ A) , iterₛ∼ g ]ₛ) ⟨$⟩ x ≡ ([ [ inj₁ₛ , inj₂ₛ ∘ inj₁ₛ ]ₛ ∘ helper , inj₂ₛ ∘ inj₂ₛ ]ₛ ∘ [ inj₁ₛ , g ]ₛ) # ⟨$⟩ x ] sub-step₁ g {u} = ≡-sym ⟦ B ⟧ (#-Uniformity B {h = [ idₛ (Delayₛ∼ A) , iterₛ∼ g ]ₛ} (λ {y} → last-step {y})) where last-step : ∀ {x} → [ ⟦ B ⟧ ⊎ₛ (Delayₛ∼ A) ][ [ inj₁ₛ , inj₂ₛ ∘ [ idₛ (Delayₛ∼ A) , iterₛ∼ g ]ₛ ]ₛ ∘ [ [ inj₁ₛ , inj₂ₛ ∘ inj₁ₛ ]ₛ ∘ helper , inj₂ₛ ∘ inj₂ₛ ]ₛ ∘ [ inj₁ₛ , g ]ₛ ⟨$⟩ x ≡ (helper ∘ [ idₛ (Delayₛ∼ A) , iterₛ∼ g ]ₛ) ⟨$⟩ x ] last-step {inj₁ (now a)} = ≡-refl (⟦ B ⟧ ⊎ₛ (Delayₛ∼ A)) last-step {inj₁ (later a)} = ≡-refl (⟦ B ⟧ ⊎ₛ (Delayₛ∼ A)) last-step {inj₂ a} with g ⟨$⟩ a in eqb ... | inj₁ (now b) = ≡-refl (⟦ B ⟧ ⊎ₛ (Delayₛ∼ A)) ... | inj₁ (later b) = ≡-refl (⟦ B ⟧ ⊎ₛ (Delayₛ∼ A)) ... | inj₂ b = ≡-refl (⟦ B ⟧ ⊎ₛ (Delayₛ∼ A)) step₂ : [ ⟦ B ⟧ ][ ([ inj₁ₛ ∘ (helper #) , inj₂ₛ ]ₛ ∘ (disc-dom g)) # ⟨$⟩ x ≡ ([ inj₁ₛ ∘ delay-lift' , inj₂ₛ ]ₛ ∘ g) # ⟨$⟩ x ] step₂ = #-Uniformity B {h = ‖‖-quote} sub-step₂ where sub-step₂ : ∀ {x} → [ ⟦ B ⟧ ⊎ₛ X ][([ inj₁ₛ , inj₂ₛ ]ₛ ∘ ([ inj₁ₛ ∘ (helper #) , inj₂ₛ ]ₛ ∘ (disc-dom g))) ⟨$⟩ x ≡ ([ inj₁ₛ ∘ delay-lift' , inj₂ₛ ]ₛ ∘ g) ⟨$⟩ x ] sub-step₂ {x} with g ⟨$⟩ x ... | inj₁ y = ≡-refl (⟦ B ⟧ ⊎ₛ X) ... | inj₂ y = ≡-refl (⟦ B ⟧ ⊎ₛ X) open Elgot-Algebra-Morphism using (preserves) renaming (h to <<_>>) freeAlgebra : ∀ (A : Setoid ℓ ℓ) → FreeObject elgotForgetfulF A freeAlgebra A = record { FX = delay-algebras A ; η = ηₛ A ; _* = delay-lift ; *-lift = λ {B} f {x} → Elgot-Algebra.#-Fixpoint B ; *-uniq = λ {B} f g eq {x} → *-uniq' {B} f g (delay-lift f) eq (#-Fixpoint B) {x} } where *-uniq' : ∀ {B : Elgot-Algebra} (f : A ⟶ ⟦ B ⟧) (g h : Elgot-Algebra-Morphism (delay-algebras A) B) → (<< g >> ∘ (ηₛ A)) ≋ f → (<< h >> ∘ (ηₛ A)) ≋ f → << g >> ≋ << h >> *-uniq' {B} f g h eqᵍ eqʰ {x} = ≡-trans ⟦ B ⟧ (cong << g >> iter-id) (≡-trans ⟦ B ⟧ (preserves g {Delayₛ∼ A} {[ inj₁ₛ ∘ D∼⇒D≈ , inj₂ₛ ]ₛ ∘ helper-now} {x = x}) (≡-trans ⟦ B ⟧ (#-resp-≈ B (λ {x} → helper-eq' {x}) {x}) (≡-trans ⟦ B ⟧ (≡-sym ⟦ B ⟧ (preserves h {Delayₛ∼ A} {[ inj₁ₛ ∘ D∼⇒D≈ , inj₂ₛ ]ₛ ∘ helper-now} {x = x})) (≡-sym ⟦ B ⟧ (cong << h >> iter-id))))) where open Elgot-Algebra B using (_#) D∼⇒D≈ : ∀ {A : Setoid ℓ ℓ} → Delayₛ∼ A ⟶ Delayₛ A D∼⇒D≈ {A} = record { to = λ x → x ; cong = λ eq → ∼⇒≈ eq } helper-now₁ : (Delay ∣ A ∣) → (Delay ∣ A ∣ ⊎ (Delay ∣ A ∣)) helper-now₁ (now x) = inj₁ (now x) helper-now₁ (later x) = inj₂ (force x) helper-now : Delayₛ∼ A ⟶ ((Delayₛ∼ A) ⊎ₛ (Delayₛ∼ A)) helper-now = record { to = helper-now₁ ; cong = λ { (now∼ eq) → inj₁ (now∼ eq) ; (later∼ eq) → inj₂ (force∼ eq) } } helper-eq' : ∀ {x} → [ ⟦ B ⟧ ⊎ₛ Delayₛ∼ A ][ ([ inj₁ₛ ∘ << g >> , inj₂ₛ ]ₛ ∘ [ inj₁ₛ ∘ D∼⇒D≈ , inj₂ₛ ]ₛ ∘ helper-now) ⟨$⟩ x ≡ ([ inj₁ₛ ∘ << h >> , inj₂ₛ ]ₛ ∘ [ inj₁ₛ ∘ D∼⇒D≈ , inj₂ₛ ]ₛ ∘ helper-now) ⟨$⟩ x ] helper-eq' {now x} = inj₁ (≡-trans ⟦ B ⟧ eqᵍ (≡-sym ⟦ B ⟧ eqʰ)) helper-eq' {later x} = inj₂ (∼-refl A) iter-id : ∀ {x} → [ A ][ x ≈ iterₛ ([ inj₁ₛ ∘ D∼⇒D≈ , inj₂ₛ ]ₛ ∘ helper-now) ⟨$⟩ x ] iter-id′ : ∀ {x} → [ A ][ x ≈′ iterₛ ([ inj₁ₛ ∘ D∼⇒D≈ , inj₂ₛ ]ₛ ∘ helper-now) ⟨$⟩ x ] force≈ (iter-id′ {x}) = iter-id {x} iter-id {now x} = ≈-refl A iter-id {later x} = later≈ (iter-id′ {force x}) open CartesianClosed (Setoids-CCC ℓ) renaming (exp to setoid-exp) open import Algebra.Elgot.Properties distributive setoid-exp using (free-isStable) delayK : MonadK delayK = record { freealgebras = freeAlgebra ; stable = λ X → free-isStable (freeAlgebra X) }