Some progress

This commit is contained in:
Leon Vatthauer 2024-01-09 13:34:53 +01:00
parent 3679811f2c
commit 9beebd009d
Signed by: leonv
SSH key fingerprint: SHA256:G4+ddwoZmhLPRB1agvXzZMXIzkVJ36dUYZXf5NxT+u8
3 changed files with 152 additions and 19 deletions

View file

@ -36,7 +36,7 @@ module Category.Ambient.Setoids {} where
≋-trans : ∀ {A B : Setoid } {f g h : A ⟶ B} → f ≋ g → g ≋ h → f ≋ h
≋-trans {A} {B} {f} {g} {h} = IsEquivalence.trans (Setoid.isEquivalence (A ⇨ B)) {f} {g} {h}
-- we define ourselves, instead of importing it, to avoid lifting the universe lifting (builtin Nats are defined on Set₀)
-- we define ourselves, instead of importing it, to avoid lifting the universe levels (builtin Nats are defined on Set₀)
data : Set where
zero :
suc :

View file

@ -13,7 +13,9 @@ open import Categories.Monad
open import Categories.Category.Instance.Setoids
open import Categories.NaturalTransformation hiding (id)
open import Data.Product
open import Data.Nat using (; suc; zero)
open import Data.Product.Relation.Binary.Pointwise.NonDependent
-- open import Data.Nat using (; suc; zero)
open import Category.Ambient.Setoids
module Monad.Instance.K.Instance.D {c } where
@ -204,6 +206,15 @@ module DelayMonad where
∼⇒≈ {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))
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)
@ -214,7 +225,6 @@ module DelayMonad where
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))
@ -227,11 +237,14 @@ module DelayMonad where
liftFₛ : ∀ {A B : Setoid c (c ⊔ )} → A ⟶ B → Delayₛ A ⟶ Delayₛ B
liftFₛ {A} {B} f = record { _⟨$⟩_ = liftF < f > ; cong = lift-cong f }
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))
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 : Setoid c (c ⊔ )} → A ⟶ B → Delayₛ' A ⟶ Delayₛ' B
liftFₛ {A} {B} f = record { _⟨$⟩_ = 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)
-- this needs polymorphic universe levels
_≋_ : ∀ {c' '} {A B : Setoid c' '} → A ⟶ B → A ⟶ B → Set (c' ⊔ ')
@ -316,6 +329,15 @@ module DelayMonad where
μₛ : ∀ (A : Setoid c (c ⊔ )) → Delayₛ (Delayₛ A) ⟶ Delayₛ A
μₛ A = record { _⟨$⟩_ = μ {A} ; cong = μ-cong A }
μₛ∼ : ∀ (A : Setoid c (c ⊔ )) → Delayₛ' (Delayₛ' A) ⟶ Delayₛ' A
μₛ∼ A = record { _⟨$⟩_ = μ {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 y : Delay (Delay A )} → [ Delayₛ A ][ x ≈′ y ] → [ B ][ (μₛ B ∘ liftFₛ (liftFₛ f)) ⟨$⟩ x ≈′ (liftFₛ f ∘ μₛ A) ⟨$⟩ y ]
force≈ (μ-natural {A} {B} f {x} {y} x≈y) = μ-natural f (force≈ x≈y)
@ -367,6 +389,8 @@ module DelayMonad where
; 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 ]
@ -385,6 +409,23 @@ module extra {A : Setoid c (c ⊔ )} where
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))
@ -398,22 +439,51 @@ module extra {A : Setoid c (c ⊔ )} where
force≲ (≈′→≲′ x≈y) = ≈→≲ (force≈ x≈y)
delta₀ : {x : Delay A } {a : A } → (x↓a : [ A ][ x ↓ a ]) →
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 × )
delta : {x y : Delay A } → [ A ][ x ≲′ y ] → Delay ( 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 × → Delay A
ι : A × → Delay A
ι : 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 { _⟨$⟩_ = ι ; 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)
```

View file

@ -15,6 +15,8 @@ open import Categories.FreeObjects.Free
open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_)
open import Categories.Category.Instance.Properties.Setoids.Choice using ()
open import Data.Product.Relation.Binary.Pointwise.NonDependent
open import Data.Product
-- open import Data.Nat
```
-->
@ -31,6 +33,7 @@ module Monad.Instance.Setoids.K' { : Level} where
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)
@ -168,20 +171,80 @@ module Monad.Instance.Setoids.K' { : Level} where
open Elgot-Algebra using () 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 = record { _⟨$⟩_ = < (B Elgot-Algebra.#) helper > ; cong = λ {x} {y} x≈y → {! !} } ; preserves = {! !} }
delay-lift {A} {B} f = record { h = record { _⟨$⟩_ = ((B Elgot-Algebra.#) helper) ._⟨$⟩_ ; cong = helper#≈-cong } ; preserves = {! !} }
where
-- (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 } → (xy : [ 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 { _⟨$⟩_ = helper₁ ; cong = strong-cong }
helper = record { _⟨$⟩_ = helper₁ ; cong = helper₁-cong}
helper#-cong : {x y : Delay A } → (xy : [ A ][ x y ]) → [ ⟦ B ⟧ ][ (B Elgot-Algebra.#) helper ⟨$⟩ x ≡ (B Elgot-Algebra.#) helper ⟨$⟩ y ]
helper#-cong {x} {y} xy = cong ((B Elgot-Algebra.#) helper) xy
helper#≈-cong : {x y : Delay A } → (x≈y : [ A ][ x ≈ y ]) → [ ⟦ B ⟧ ][ (B Elgot-Algebra.#) helper ⟨$⟩ x ≡ (B Elgot-Algebra.#) helper ⟨$⟩ y ]
-- key special case
helper#≈-cong' : {z : Delay ( A × )} → [ ⟦ B ⟧ ][ (B Elgot-Algebra.#) helper ⟨$⟩ liftF proj₁ z ≡ (B Elgot-Algebra.#) 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
strong-cong : ∀ {x y : Delay A } → [ A ][ x y ] → [ ⟦ B ⟧ ⊎ₛ Delayₛ' A ][ helper₁ x ≡ helper₁ y ]
strong-cong {.(now _)} {.(now _)} (now x≡y) = cong (inj₁ₛ {_} {_} {_} {_} {⟦ B ⟧} {Delayₛ' A}) (cong f x≡y)
strong-cong {.(later _)} {.(later _)} (later xy) = cong (inj₂ₛ {_} {_} {_} {_} {⟦ B ⟧} {Delayₛ' A}) (force xy)
ineq₁ = ≈→≲ {A} x≈y
ineq₂ = ≈→≲ {A} (≈-sym A x≈y)
z₁ = delta {A} ineq₁
z₂ = delta {A} ineq₂
helper#≈-cong' {z} = ≡-trans ⟦ B ⟧ (≡-sym ⟦ B ⟧ eq₁) eq₂
where
helper₁' : Delay ( A × {}) → ⟦ B ⟧ ⊎ Delay ( A × )
helper₁' (now (x , zero)) = inj₁ (< f > x)
helper₁' (now (x , suc n)) with helper₁' (now (x , n))
... | inj₁ r = inj₁ r
... | inj₂ y = inj₂ (later (record { force = y }))
helper₁' (later y) = inj₂ (force y)
helper₁-cong' : {x y : Delay ( A × {})} → (xy : [ 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)) with helper₁' (now (x , n)) | helper₁' (now (y , n)) | helper₁-cong' {now (x , n)} (now (x≡y , ≣-refl))
... | inj₁ r | inj₁ s | inj₁ r≡s = inj₁ r≡s
... | inj₂ x' | inj₂ y' | inj₂ x'y' = inj₂ (later (record { force = x'y' }))
helper₁-cong' (later x≡y) = inj₂ (force x≡y)
helper' : Delayₛ' (A ×ₛ -setoid) ⟶ ⟦ B ⟧ ⊎ₛ Delayₛ' (A ×ₛ -setoid)
helper' = record { _⟨$⟩_ = helper₁' ; cong = helper₁-cong'}
-- Should follow by compositionality + fixpoint
eq₁ : [ ⟦ B ⟧ ][ (B Elgot-Algebra.#) helper' ⟨$⟩ z ≡ (B Elgot-Algebra.#) helper ⟨$⟩ liftF proj₁ z ]
eq₁ = {!!}
eq : ∀ {x y} → [ A ×ₛ -setoid ][ x y ] → [ ⟦ B ⟧ ⊎ₛ Delayₛ' A ][ [ inj₁ , inj₂ ∘′ μ ∘′ (liftF ι) ] (helper₁' x) ≡ (helper₁ ∘′ μ ∘′ (liftF ι)) y ]
eq {now (x , zero)} {now (y , zero)} (now (x≡y , _)) = cong (inj₁ₛ {_} {_} {_} {_} {⟦ B ⟧} {Delayₛ' A}) (cong f x≡y)
-- eq {now (x , suc n)} {now (y , suc m)} (now (x≡y , n≡m)) with helper₁' (now (x , n))
-- ... | inj₁ r = {! !} -- <- contradiction!
-- ... | inj₂ r = {! !}
eq {now (x , suc n)} {now (y , suc m)} (now (x≡y , n≡m)) = {! !}
-- TODO induction, by a recursive call to eq, we will get info about helper₁' (now (x , n))
eq {.(later _)} {.(later _)} (later x≡y) = cong (inj₂ₛ {_} {_} {_} {_} {⟦ B ⟧} {Delayₛ' A}) (cong (μₛ∼ A) (cong (liftFₛ ιₛ') (force x≡y)))
-- Should follow by uniformity
eq₂ : [ ⟦ B ⟧ ][ (B Elgot-Algebra.#) helper' ⟨$⟩ z ≡ (B Elgot-Algebra.#) helper ⟨$⟩ μ {A} (liftF (ι {A}) z)]
eq₂ = Elgot-Algebra.#-Uniformity B {Delayₛ' (A ×ₛ -setoid {})} {Delayₛ' A} {helper'} {helper} {μₛ∼ A ∘ liftFₛ ιₛ'} eq
(-refl (A ×ₛ -setoid))
<<_>> = Elgot-Algebra-Morphism.h