Prove exercise 4.26 from the textbook

Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
Danila Fedorin 2024-03-06 00:35:29 -08:00
parent fa0282ff6f
commit 48983c55b1

View File

@ -9,7 +9,7 @@ module Lattice.Map {a b : Level} (A : Set a) (B : Set b)
(≡-dec-A : Decidable (_≡_ {a} {A})) (≡-dec-A : Decidable (_≡_ {a} {A}))
(lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where (lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
import Data.List.Membership.Propositional as MemProp open import Data.List.Membership.Propositional as MemProp using () renaming (_∈_ to _∈ˡ_)
open import Relation.Nullary using (¬_; Dec; yes; no) open import Relation.Nullary using (¬_; Dec; yes; no)
open import Data.Nat using () open import Data.Nat using ()
@ -27,6 +27,7 @@ open IsLattice lB using () renaming
; ⊔-idemp to ⊔₂-idemp; ⊔-comm to ⊔₂-comm; ⊔-assoc to ⊔₂-assoc ; ⊔-idemp to ⊔₂-idemp; ⊔-comm to ⊔₂-comm; ⊔-assoc to ⊔₂-assoc
; ⊓-idemp to ⊓₂-idemp; ⊓-comm to ⊓₂-comm; ⊓-assoc to ⊓₂-assoc ; ⊓-idemp to ⊓₂-idemp; ⊓-comm to ⊓₂-comm; ⊓-assoc to ⊓₂-assoc
; absorb-⊔-⊓ to absorb-⊔₂-⊓₂; absorb-⊓-⊔ to absorb-⊓₂-⊔₂ ; absorb-⊔-⊓ to absorb-⊔₂-⊓₂; absorb-⊓-⊔ to absorb-⊓₂-⊔₂
; _≼_ to _≼₂_
) )
private module ImplKeys where private module ImplKeys where
@ -55,6 +56,15 @@ private module _ where
ListAB-functional {l = _ xs } (push _ uxs) (there k,v∈xs) (there k,v'∈xs) = ListAB-functional {l = _ xs } (push _ uxs) (there k,v∈xs) (there k,v'∈xs) =
ListAB-functional uxs k,v∈xs k,v'∈xs ListAB-functional uxs k,v∈xs k,v'∈xs
k∈-dec : (k : A) (l : List A) Dec (k l)
k∈-dec k [] = no (λ ())
k∈-dec k (x xs)
with (≡-dec-A k x)
... | yes refl = yes (here refl)
... | no k≢x with (k∈-dec k xs)
... | yes k∈xs = yes (there k∈xs)
... | no k∉xs = no (λ { (here k≡x) k≢x k≡x; (there k∈xs) k∉xs k∈xs })
∈k-dec : (k : A) (l : List (A × B)) Dec (k keys l) ∈k-dec : (k : A) (l : List (A × B)) Dec (k keys l)
∈k-dec k [] = no (λ ()) ∈k-dec k [] = no (λ ())
∈k-dec k ((k' , v) xs) ∈k-dec k ((k' , v) xs)
@ -463,7 +473,6 @@ private module ImplInsert (f : B → B → B) where
intersect-combines u₁ u₂ k,v₁∈l₁ k,v₂∈l₂ = intersect-combines u₁ u₂ k,v₁∈l₁ k,v₂∈l₂ =
restrict-preserves-∈₂ (∈-cong proj₁ k,v₁∈l₁) (updates-combine u₁ u₂ k,v₁∈l₁ k,v₂∈l₂) restrict-preserves-∈₂ (∈-cong proj₁ k,v₁∈l₁) (updates-combine u₁ u₂ k,v₁∈l₁ k,v₂∈l₂)
Map : Set (a ⊔ℓ b) Map : Set (a ⊔ℓ b)
Map = Σ (List (A × B)) (λ l Unique (ImplKeys.keys l)) Map = Σ (List (A × B)) (λ l Unique (ImplKeys.keys l))
@ -532,6 +541,7 @@ open ImplInsert _⊔₂_ using
; union-preserves-∈₁ ; union-preserves-∈₁
; union-preserves-∈₂ ; union-preserves-∈₂
; union-preserves-∉ ; union-preserves-∉
; union-preserves-∈k₁
) )
⊔-combines : {k : A} {v₁ v₂ : B} {m₁ m₂ : Map} (k , v₁) m₁ (k , v₂) m₂ (k , v₁ ⊔₂ v₂) (m₁ m₂) ⊔-combines : {k : A} {v₁ v₂ : B} {m₁ m₂ : Map} (k , v₁) m₁ (k , v₂) m₂ (k , v₁ ⊔₂ v₂) (m₁ m₂)
@ -582,6 +592,11 @@ Expr-Provenance k (e₁ ∩ e₂) k∈ke₁e₂
... | no k∉ke₁ | yes k∈ke₂ = ⊥-elim (intersect-preserves-∉₁ {l₂ = proj₁ e₂ } k∉ke₁ k∈ke₁e₂) ... | no k∉ke₁ | yes k∈ke₂ = ⊥-elim (intersect-preserves-∉₁ {l₂ = proj₁ e₂ } k∉ke₁ k∈ke₁e₂)
... | no k∉ke₁ | no k∉ke₂ = ⊥-elim (intersect-preserves-∉₂ {l₁ = proj₁ e₁ } k∉ke₂ k∈ke₁e₂) ... | no k∉ke₁ | no k∉ke₂ = ⊥-elim (intersect-preserves-∉₂ {l₁ = proj₁ e₁ } k∉ke₂ k∈ke₁e₂)
Expr-Provenance-≡ : (k : A) (v : B) (e : Expr) (k , v) e Provenance k v e
Expr-Provenance-≡ k v e k,v∈e
with (v' , (p , k,v'∈e)) Expr-Provenance k e (forget {m = e } k,v∈e)
rewrite Map-functional {m = e } k,v∈e k,v'∈e = p
module _ (≈₂-dec : IsDecidable _≈₂_) where module _ (≈₂-dec : IsDecidable _≈₂_) where
private module _ where private module _ where
data SubsetInfo (m₁ m₂ : Map) : Set (a ⊔ℓ b) where data SubsetInfo (m₁ m₂ : Map) : Set (a ⊔ℓ b) where
@ -904,3 +919,184 @@ lattice = record
⊓-equal-keys : {m₁ m₂ : Map} keys m₁ keys m₂ keys m₁ keys (m₁ m₂) ⊓-equal-keys : {m₁ m₂ : Map} keys m₁ keys m₂ keys m₁ keys (m₁ m₂)
⊓-equal-keys km₁≡km₂ = intersect-equal-keys km₁≡km₂ ⊓-equal-keys km₁≡km₂ = intersect-equal-keys km₁≡km₂
private module _ where
open MemProp using (_∈_)
transform : List (A × B) List A (A B) List (A × B)
transform [] _ _ = []
transform ((k , v) xs) ks f
with k∈-dec k ks
... | yes _ = (k , f k) transform xs ks f
... | no _ = (k , v) transform xs ks f
transform-keys-≡ : (l : List (A × B)) (ks : List A) (f : A B)
ImplKeys.keys l ImplKeys.keys (transform l ks f)
transform-keys-≡ [] ks f = refl
transform-keys-≡ ((k , v) xs) ks f
with k∈-dec k ks
... | yes _ rewrite transform-keys-≡ xs ks f = refl
... | no _ rewrite transform-keys-≡ xs ks f = refl
transform-∉k-forward : {l : List (A × B)} (ks : List A) (f : A B) {k : A}
¬ k ∈ˡ ImplKeys.keys l ¬ k ∈ˡ ImplKeys.keys (transform l ks f)
transform-∉k-forward {l} ks f k∉kl rewrite transform-keys-≡ l ks f = k∉kl
transform-∈k-forward : {l : List (A × B)} (ks : List A) (f : A B) {k : A}
k ∈ˡ ImplKeys.keys l k ∈ˡ ImplKeys.keys (transform l ks f)
transform-∈k-forward {l} ks f k∈kl rewrite transform-keys-≡ l ks f = k∈kl
transform-∈k-backward : {l : List (A × B)} (ks : List A) (f : A B) {k : A}
k ∈ˡ ImplKeys.keys (transform l ks f) k ∈ˡ ImplKeys.keys l
transform-∈k-backward {l} ks f k∈kt rewrite transform-keys-≡ l ks f = k∈kt
transform-k∈ks : (l : List (A × B)) {ks : List A} (f : A B) {k : A}
k ∈ˡ ks k ∈ˡ ImplKeys.keys (transform l ks f) (k , f k) ∈ˡ transform l ks f
transform-k∈ks [] _ _ ()
transform-k∈ks ((k' , v) xs) {ks} f {k} k∈ks k∈kl
with k∈-dec k' ks | k∈kl
... | yes _ | here refl = here refl
... | no k∉ks | here refl = ⊥-elim (k∉ks k∈ks)
... | yes _ | there k∈kxs = there (transform-k∈ks xs f k∈ks k∈kxs)
... | no _ | there k∈kxs = there (transform-k∈ks xs f k∈ks k∈kxs)
transform-k∉ks-forward : {l : List (A × B)} {ks : List A} (f : A B) {k : A} {v : B}
¬ k ∈ˡ ks (k , v) ∈ˡ l (k , v) ∈ˡ transform l ks f
transform-k∉ks-forward {[]} _ _ ()
transform-k∉ks-forward {(k' , v') xs} {ks} f {k} {v} k∉ks k,v∈l
with k∈-dec k' ks | k,v∈l
... | yes k∈ks | here refl = ⊥-elim (k∉ks k∈ks)
... | no k∉ks | here refl = here refl
... | yes _ | there k,v∈xs = there (transform-k∉ks-forward f k∉ks k,v∈xs)
... | no _ | there k,v∈xs = there (transform-k∉ks-forward f k∉ks k,v∈xs)
transform-k∉ks-backward : {l : List (A × B)} {ks : List A} (f : A B) {k : A} {v : B}
¬ k ∈ˡ ks (k , v) ∈ˡ transform l ks f (k , v) ∈ˡ l
transform-k∉ks-backward {[]} _ _ ()
transform-k∉ks-backward {(k' , v') xs} {ks} f {k} {v} k∉ks k,v∈tl
with k∈-dec k' ks | k,v∈tl
... | yes k∈ks | here refl = ⊥-elim (k∉ks k∈ks)
... | no k∉ks | here refl = here refl
... | yes _ | there k,v∈txs = there (transform-k∉ks-backward f k∉ks k,v∈txs)
... | no _ | there k,v∈txs = there (transform-k∉ks-backward f k∉ks k,v∈txs)
_updating_via_ : Map List A (A B) Map
_updating_via_ (kvs , uks) ks f =
( transform kvs ks f
, subst Unique (transform-keys-≡ kvs ks f) uks
)
updating-via-∉k-forward : (m : Map) (ks : List A) (f : A B) {k : A}
¬ k ∈k m ¬ k ∈k (m updating ks via f)
updating-via-∉k-forward m = transform-∉k-forward
updating-via-∈k-forward : (m : Map) (ks : List A) (f : A B) {k : A}
k ∈k m k ∈k (m updating ks via f)
updating-via-∈k-forward m = transform-∈k-forward
updating-via-∈k-backward : (m : Map) (ks : List A) (f : A B) {k : A}
k ∈k (m updating ks via f) k ∈k m
updating-via-∈k-backward m = transform-∈k-backward
updating-via-k∈ks : (m : Map) {ks : List A} (f : A B) {k : A}
k ∈ˡ ks k ∈k (m updating ks via f) (k , f k) (m updating ks via f)
updating-via-k∈ks m = transform-k∈ks (proj₁ m)
updating-via-k∈ks-forward : (m : Map) {ks : List A} (f : A B) {k : A}
k ∈ˡ ks k ∈k m (k , f k) (m updating ks via f)
updating-via-k∈ks-forward m {ks} f k∈ks k∈km rewrite transform-keys-≡ (proj₁ m) ks f = transform-k∈ks (proj₁ m) f k∈ks k∈km
updating-via-k∈ks-≡ : (m : Map) {ks : List A} (f : A B) {k : A} {v : B}
k ∈ˡ ks (k , v) (m updating ks via f) v f k
updating-via-k∈ks-≡ m {ks} f k∈ks k,v∈um
with updating-via-k∈ks m f k∈ks (forget {m = (m updating ks via f)} k,v∈um)
... | k,fk∈um = Map-functional {m = (m updating ks via f)} k,v∈um k,fk∈um
updating-via-k∉ks-forward : (m : Map) {ks : List A} (f : A B) {k : A} {v : B}
¬ k ∈ˡ ks (k , v) m (k , v) (m updating ks via f)
updating-via-k∉ks-forward m = transform-k∉ks-forward
updating-via-k∉ks-backward : (m : Map) {ks : List A} (f : A B) {k : A} {v : B}
¬ k ∈ˡ ks (k , v) (m updating ks via f) (k , v) m
updating-via-k∉ks-backward m = transform-k∉ks-backward
module _ {l} {L : Set l}
{_≈ˡ_ : L L Set l} {_⊔ˡ_ : L L L} {_⊓ˡ_ : L L L}
(lL : IsLattice L _≈ˡ_ _⊔ˡ_ _⊓ˡ_) where
open IsLattice isLattice using (_≼_)
open IsLattice lL using () renaming (_≼_ to _≼ˡ_)
module _ (f : L Map) (f-Monotonic : Monotonic _≼ˡ_ _≼_ f)
(g : A L B) (g-Monotonicʳ : k Monotonic _≼ˡ_ _≼₂_ (g k))
(ks : List A) where
updater : L A B
updater l k = g k l
f' : L Map
f' l = (f l) updating ks via (updater l)
f'-Monotonic : Monotonic _≼ˡ_ _≼_ f'
f'-Monotonic {l₁} {l₂} l₁≼l₂ = (f'l₁f'l₂⊆f'l₂ , f'l₂⊆f'l₁f'l₂)
where
fl₁fl₂⊆fl₂ = proj₁ (f-Monotonic l₁≼l₂)
fl₂⊆fl₁fl₂ = proj₂ (f-Monotonic l₁≼l₂)
f'l₁f'l₂⊆f'l₂ : ((f' l₁) (f' l₂)) f' l₂
f'l₁f'l₂⊆f'l₂ k v k,v∈f'l₁f'l₂
with Expr-Provenance-≡ k v ((` (f' l₁)) (` (f' l₂))) k,v∈f'l₁f'l₂
... | in (single k,v∈f'l₁) k∉kf'l₂ =
let
k∈kfl₁ = updating-via-∈k-backward (f l₁) ks (updater l₁) (forget {m = f' l₁} k,v∈f'l₁)
k∈kfl₁fl₂ = union-preserves-∈k₁ {l₁ = proj₁ (f l₁)} {l₂ = proj₁ (f l₂)} k∈kfl₁
(v' , k,v'∈fl₁l₂) = locate {m = (f l₁ f l₂)} k∈kfl₁fl₂
(v'' , (v'≈v'' , k,v''∈fl₂)) = fl₁fl₂⊆fl₂ k v' k,v'∈fl₁l₂
k∈kf'l₂ = updating-via-∈k-forward (f l₂) ks (updater l₂) (forget {m = f l₂} k,v''∈fl₂)
in
⊥-elim (k∉kf'l₂ k∈kf'l₂)
... | in k∉kf'l₁ (single k,v'∈f'l₂) =
(v , (IsLattice.≈-refl lB , k,v'∈f'l₂))
... | bothᵘ (single {v₁} k,v₁∈f'l₁) (single {v₂} k,v₂∈f'l₂)
with k∈-dec k ks
... | yes k∈ks
with refl updating-via-k∈ks-≡ (f l₁) (updater l₁) k∈ks k,v₁∈f'l₁
with refl updating-via-k∈ks-≡ (f l₂) (updater l₂) k∈ks k,v₂∈f'l₂ =
(updater l₂ k , (g-Monotonicʳ k l₁≼l₂ , k,v₂∈f'l₂))
... | no k∉ks =
let
k,v₁∈fl₁ = updating-via-k∉ks-backward (f l₁) (updater l₁) k∉ks k,v₁∈f'l₁
k,v₂∈fl₂ = updating-via-k∉ks-backward (f l₂) (updater l₂) k∉ks k,v₂∈f'l₂
k,v₁v₂∈fl₁fl₂ = ⊔-combines {m₁ = f l₁} {m₂ = f l₂} k,v₁∈fl₁ k,v₂∈fl₂
(v' , (v'≈v₁v₂ , k,v'∈fl₂)) = fl₁fl₂⊆fl₂ k _ k,v₁v₂∈fl₁fl₂
k,v'∈f'l₂ = updating-via-k∉ks-forward (f l₂) (updater l₂) k∉ks k,v'∈fl₂
in
(v' , (v'≈v₁v₂ , k,v'∈f'l₂))
f'l₂⊆f'l₁f'l₂ : f' l₂ ((f' l₁) (f' l₂))
f'l₂⊆f'l₁f'l₂ k v k,v∈f'l₂
with k∈kfl₂ updating-via-∈k-backward (f l₂) ks (updater l₂) (forget {m = f' l₂} k,v∈f'l₂)
with (v' , k,v'∈fl₂) locate {m = f l₂} k∈kfl₂
with (v'' , (v'≈v'' , k,v''∈fl₁fl₂)) fl₂⊆fl₁fl₂ k v' k,v'∈fl₂
with Expr-Provenance-≡ k v'' ((` (f l₁)) (` (f l₂))) k,v''∈fl₁fl₂
... | in (single k,v''∈fl₁) k∉kfl₂ = ⊥-elim (k∉kfl₂ k∈kfl₂)
... | in k∉kfl₁ (single k,v''∈fl₂) =
let
k∉kf'l₁ = updating-via-∉k-forward (f l₁) ks (updater l₁) k∉kfl₁
in
(v , (IsLattice.≈-refl lB , union-preserves-∈₂ k∉kf'l₁ k,v∈f'l₂))
... | bothᵘ (single {v₁} k,v₁∈fl₁) (single {v₂} k,v₂∈fl₂)
with k∈-dec k ks
... | yes k∈ks with refl updating-via-k∈ks-≡ (f l₂) (updater l₂) k∈ks k,v∈f'l₂ =
let
k,uv₁∈f'l₁ = updating-via-k∈ks-forward (f l₁) (updater l₁) k∈ks (forget {m = f l₁} k,v₁∈fl₁)
k,uv₂∈f'l₂ = updating-via-k∈ks-forward (f l₂) (updater l₂) k∈ks (forget {m = f l₂} k,v₂∈fl₂)
k,uv₁uv₂∈f'l₁f'l₂ = ⊔-combines {m₁ = f' l₁} {m₂ = f' l₂} k,uv₁∈f'l₁ k,uv₂∈f'l₂
in
(updater l₁ k ⊔₂ updater l₂ k , (IsLattice.≈-sym lB (g-Monotonicʳ k l₁≼l₂) , k,uv₁uv₂∈f'l₁f'l₂))
... | no k∉ks
with k,v₁∈f'l₁ updating-via-k∉ks-forward (f l₁) (updater l₁) k∉ks k,v₁∈fl₁
with k,v₂∈f'l₂ updating-via-k∉ks-forward (f l₂) (updater l₂) k∉ks k,v₂∈fl₂
with k,v₁v₂∈f'l₁f'l₂ ⊔-combines {m₁ = f' l₁} {m₂ = f' l₂} k,v₁∈f'l₁ k,v₂∈f'l₂
with refl Map-functional {m = f' l₂} k,v∈f'l₂ k,v₂∈f'l₂
with refl Map-functional {m = f l₂} k,v'∈fl₂ k,v₂∈fl₂ =
(v₁ ⊔₂ v , (v'≈v'' , k,v₁v₂∈f'l₁f'l₂))