Add more properties of update(s) and start work on provenance

Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
Danila Fedorin 2023-08-05 00:02:50 -07:00
parent 66dfe14207
commit 1780cdbda4

116
Map.agda
View File

@ -283,6 +283,9 @@ private module ImplInsert (f : B → B → B) where
... | yes _ = (k' , f v v') xs ... | yes _ = (k' , f v v') xs
... | no _ = (k' , v') update k v xs ... | no _ = (k' , v') update k v xs
updates : List (A × B) List (A × B) List (A × B)
updates l₁ l₂ = foldr update l₂ l₁
restrict : List (A × B) List (A × B) List (A × B) restrict : List (A × B) List (A × B) List (A × B)
restrict l [] = [] restrict l [] = []
restrict l ((k' , v') xs) with ∈k-dec k' l restrict l ((k' , v') xs) with ∈k-dec k' l
@ -290,7 +293,7 @@ private module ImplInsert (f : B → B → B) where
... | no _ = restrict l xs ... | no _ = restrict l xs
intersect : List (A × B) List (A × B) List (A × B) intersect : List (A × B) List (A × B) List (A × B)
intersect l₁ l₂ = restrict l₁ (foldr update l₂ l₁) intersect l₁ l₂ = restrict l₁ (updates l₁ l₂)
update-keys : {k : A} {v : B} {l : List (A × B)} update-keys : {k : A} {v : B} {l : List (A × B)}
keys l keys (update k v l) keys l keys (update k v l)
@ -305,7 +308,7 @@ private module ImplInsert (f : B → B → B) where
update-preserves-Unique {k} {v} {l} u rewrite update-keys {k} {v} {l} = u update-preserves-Unique {k} {v} {l} u rewrite update-keys {k} {v} {l} = u
updates-preserve-Unique : {l₁ l₂ : List (A × B)} updates-preserve-Unique : {l₁ l₂ : List (A × B)}
Unique (keys l₂) Unique (keys (foldr update l₂ l₁)) Unique (keys l₂) Unique (keys (updates l₁ l₂))
updates-preserve-Unique {[]} u = u updates-preserve-Unique {[]} u = u
updates-preserve-Unique {(k , v) xs} u = update-preserves-Unique (updates-preserve-Unique {xs} u) updates-preserve-Unique {(k , v) xs} u = update-preserves-Unique (updates-preserve-Unique {xs} u)
@ -329,21 +332,64 @@ private module ImplInsert (f : B → B → B) where
Unique (keys l₂) Unique (keys (intersect l₁ l₂)) Unique (keys l₂) Unique (keys (intersect l₁ l₂))
intersect-preserves-Unique {l₁} u = restrict-preserves-Unique (updates-preserve-Unique {l₁} u) intersect-preserves-Unique {l₁} u = restrict-preserves-Unique (updates-preserve-Unique {l₁} u)
restrict-needs-both : {k : A} {v : B} {l₁ l₂ : List (A × B)} restrict-needs-both : {k : A} {l₁ l₂ : List (A × B)}
(k , v) restrict l₁ l₂ (k ∈k l₁ × (k , v) l₂) k k restrict l₁ l₂ (k ∈k l₁ × k k l₂)
restrict-needs-both {k} {v} {l₁} {[]} () restrict-needs-both {k} {l₁} {[]} ()
restrict-needs-both {k} {v} {l₁} {(k' , v') xs} k,v∈l₁l₂ restrict-needs-both {k} {l₁} {(k' , _) xs} k∈l₁l₂
with ∈k-dec k' l₁ | k,v∈l₁l₂ with ∈k-dec k' l₁ | k∈l₁l₂
... | yes k'∈kl₁ | here k,v≡k',v' ... | yes k'∈kl₁ | here k≡k'
rewrite cong proj₁ k,v≡k',v' rewrite cong proj₂ k,v≡k',v' = rewrite k≡k' =
(k'∈kl₁ , here refl) (k'∈kl₁ , here refl)
... | yes _ | there k,v∈l₁xs = ... | yes _ | there k∈l₁xs =
let (k∈kl₁ , k,v∈xs) = restrict-needs-both k,v∈l₁xs let (k∈kl₁ , k∈kxs) = restrict-needs-both k∈l₁xs
in (k∈kl₁ , there k,v∈xs) in (k∈kl₁ , there k∈kxs)
... | no k'∉kl₁ | k,v∈l₁xs = ... | no k'∉kl₁ | k∈l₁xs =
let (k∈kl₁ , k,v∈xs) = restrict-needs-both k,v∈l₁xs let (k∈kl₁ , k∈kxs) = restrict-needs-both k∈l₁xs
in (k∈kl₁ , there k,v∈xs) in (k∈kl₁ , there k∈kxs)
update-preserves-∈ : {k k' : A} {v v' : B} {l : List (A × B)}
¬ k k' (k , v) l (k , v) update k' v' l
update-preserves-∈ {k} {k'} {v} {v'} {(k'' , v'') xs} k≢k' (here k,v≡k'',v'')
rewrite cong proj₁ k,v≡k'',v'' rewrite cong proj₂ k,v≡k'',v''
with ≡-dec-A k' k''
... | yes k'≡k'' = absurd (k≢k' (sym k'≡k''))
... | no _ = here refl
update-preserves-∈ {k} {k'} {v} {v'} {(k'' , v'') xs} k≢k' (there k,v∈xs)
with ≡-dec-A k' k''
... | yes _ = there k,v∈xs
... | no _ = there (update-preserves-∈ k≢k' k,v∈xs)
updates-preserve-∈₂ : {k : A} {v : B} {l₁ l₂ : List (A × B)}
¬ k ∈k l₁ (k , v) l₂ (k , v) updates l₁ l₂
updates-preserve-∈₂ {k} {v} {[]} {l₂} _ k,v∈l₂ = k,v∈l₂
updates-preserve-∈₂ {k} {v} {(k' , v') xs} {l₂} k∉kl₁ k,v∈l₂ =
update-preserves-∈ (λ k≡k' k∉kl₁ (here k≡k')) (updates-preserve-∈₂ (λ k∈kxs k∉kl₁ (there k∈kxs)) k,v∈l₂)
update-combines : {k : A} {v v' : B} {l : List (A × B)}
Unique (keys l) (k , v) l (k , f v' v) update k v' l
update-combines {k} {v} {v'} {(k' , v'') xs} _ (here k,v=k',v'')
rewrite cong proj₁ k,v=k',v'' rewrite cong proj₂ k,v=k',v''
with ≡-dec-A k' k'
... | yes _ = here refl
... | no k'≢k' = absurd (k'≢k' refl)
update-combines {k} {v} {v'} {(k' , v'') xs} (push k'≢xs uxs) (there k,v∈xs)
with ≡-dec-A k k'
... | yes k≡k' rewrite k≡k' = absurd (All¬-¬Any k'≢xs (∈-cong proj₁ k,v∈xs))
... | no _ = there (update-combines uxs k,v∈xs)
updates-combine : {k : A} {v₁ v₂ : B} {l₁ l₂ : List (A × B)}
Unique (keys l₁) Unique (keys l₂)
(k , v₁) l₁ (k , v₂) l₂ (k , f v₁ v₂) updates l₁ l₂
updates-combine {k} {v₁} {v₂} {(k' , v') xs} {l₂} (push k'≢xs uxs₁) u₂ (here k,v₁≡k',v') k,v₂∈l₂
rewrite cong proj₁ k,v₁≡k',v' rewrite cong proj₂ k,v₁≡k',v' =
update-combines (updates-preserve-Unique {l₁ = xs} u₂) (updates-preserve-∈₂ (All¬-¬Any k'≢xs) k,v₂∈l₂)
updates-combine {k} {v₁} {v₂} {(k' , v') xs} {l₂} (push k'≢xs uxs₁) u₂ (there k,v₁∈xs) k,v₂∈l₂ =
update-preserves-∈ k≢k' (updates-combine uxs₁ u₂ k,v₁∈xs k,v₂∈l₂)
where
k≢k' : ¬ k k'
k≢k' with ≡-dec-A k k'
... | yes k≡k' rewrite k≡k' = absurd (All¬-¬Any k'≢xs (∈-cong proj₁ k,v₁∈xs))
... | no k≢k' = k≢k'
Map : Set (a b) Map : Set (a b)
@ -361,6 +407,7 @@ Map-functional {m = (l , ul)} k,v∈m k,v'∈m = ListAB-functional ul k,v∈m k,
data Expr : Set (a b) where data Expr : Set (a b) where
`_ : Map Expr `_ : Map Expr
__ : Expr Expr Expr __ : Expr Expr Expr
_∩_ : Expr Expr Expr
module _ (f : B B B) where module _ (f : B B B) where
open ImplInsert f renaming open ImplInsert f renaming
@ -375,15 +422,29 @@ module _ (f : B → B → B) where
intersect : Map Map Map intersect : Map Map Map
intersect (kvs₁ , _) (kvs₂ , uks₂) = (intersect-impl kvs₁ kvs₂ , intersect-preserves-Unique {kvs₁} {kvs₂} uks₂) intersect (kvs₁ , _) (kvs₂ , uks₂) = (intersect-impl kvs₁ kvs₂ , intersect-preserves-Unique {kvs₁} {kvs₂} uks₂)
module _ (fUnion : B B B) (fIntersect : B B B) where
open ImplInsert fUnion using
( union-combines
; union-preserves-∈₁
; union-preserves-∈₂
; union-preserves-∉
)
open ImplInsert fIntersect using
( restrict-needs-both
; updates
)
⟦_⟧ : Expr -> Map ⟦_⟧ : Expr -> Map
` m = m ` m = m
e₁ e₂ = union e₁ e₂ e₁ e₂ = union fUnion e₁ e₂
e₁ e₂ = intersect fIntersect e₁ e₂
data Provenance (k : A) : B Expr Set (a b) where data Provenance (k : A) : B Expr Set (a b) where
single : {v : B} {m : Map} (k , v) m Provenance k v (` m) single : {v : B} {m : Map} (k , v) m Provenance k v (` m)
in₁ : {v : B} {e₁ e₂ : Expr} Provenance k v e₁ ¬ k ∈k e₂ Provenance k v (e₁ e₂) in₁ : {v : B} {e₁ e₂ : Expr} Provenance k v e₁ ¬ k ∈k e₂ Provenance k v (e₁ e₂)
in₂ : {v : B} {e₁ e₂ : Expr} ¬ k ∈k e₁ Provenance k v e₂ Provenance k v (e₁ e₂) in₂ : {v : B} {e₁ e₂ : Expr} ¬ k ∈k e₁ Provenance k v e₂ Provenance k v (e₁ e₂)
bothᵘ : {v₁ v₂ : B} {e₁ e₂ : Expr} Provenance k v₁ e₁ Provenance k v₂ e₂ Provenance k (f v₁ v₂) (e₁ e₂) bothᵘ : {v₁ v₂ : B} {e₁ e₂ : Expr} Provenance k v₁ e₁ Provenance k v₂ e₂ Provenance k (fUnion v₁ v₂) (e₁ e₂)
Expr-Provenance : (k : A) (e : Expr) k ∈k e Σ B (λ v (Provenance k v e × (k , v) e )) Expr-Provenance : (k : A) (e : Expr) k ∈k e Σ B (λ v (Provenance k v e × (k , v) e ))
Expr-Provenance k (` m) k∈km = let (v , k,v∈m) = locate k∈km in (v , (single k,v∈m , k,v∈m)) Expr-Provenance k (` m) k∈km = let (v , k,v∈m) = locate k∈km in (v , (single k,v∈m , k,v∈m))
@ -392,7 +453,7 @@ module _ (f : B → B → B) where
... | yes k∈ke₁ | yes k∈ke₂ = ... | yes k∈ke₁ | yes k∈ke₂ =
let (v₁ , (g₁ , k,v₁∈e₁)) = Expr-Provenance k e₁ k∈ke₁ let (v₁ , (g₁ , k,v₁∈e₁)) = Expr-Provenance k e₁ k∈ke₁
(v₂ , (g₂ , k,v₂∈e₂)) = Expr-Provenance k e₂ k∈ke₂ (v₂ , (g₂ , k,v₂∈e₂)) = Expr-Provenance k e₂ k∈ke₂
in (f v₁ v₂ , (bothᵘ g₁ g₂ , union-combines (proj₂ e₁ ) (proj₂ e₂ ) k,v₁∈e₁ k,v₂∈e₂)) in (fUnion v₁ v₂ , (bothᵘ g₁ g₂ , union-combines (proj₂ e₁ ) (proj₂ e₂ ) k,v₁∈e₁ k,v₂∈e₂))
... | yes k∈ke₁ | no k∉ke₂ = ... | yes k∈ke₁ | no k∉ke₂ =
let (v₁ , (g₁ , k,v₁∈e₁)) = Expr-Provenance k e₁ k∈ke₁ let (v₁ , (g₁ , k,v₁∈e₁)) = Expr-Provenance k e₁ k∈ke₁
in (v₁ , (in g₁ k∉ke₂ , union-preserves-∈₁ (proj₂ e₁ ) k,v₁∈e₁ k∉ke₂)) in (v₁ , (in g₁ k∉ke₂ , union-preserves-∈₁ (proj₂ e₁ ) k,v₁∈e₁ k∉ke₂))
@ -400,6 +461,12 @@ module _ (f : B → B → B) where
let (v₂ , (g₂ , k,v₂∈e₂)) = Expr-Provenance k e₂ k∈ke₂ let (v₂ , (g₂ , k,v₂∈e₂)) = Expr-Provenance k e₂ k∈ke₂
in (v₂ , (in k∉ke₁ g₂ , union-preserves-∈₂ k∉ke₁ k,v₂∈e₂)) in (v₂ , (in k∉ke₁ g₂ , union-preserves-∈₂ k∉ke₁ k,v₂∈e₂))
... | no k∉ke₁ | no k∉ke₂ = absurd (union-preserves-∉ k∉ke₁ k∉ke₂ k∈ke₁e₂) ... | no k∉ke₁ | no k∉ke₂ = absurd (union-preserves-∉ k∉ke₁ k∉ke₂ k∈ke₁e₂)
Expr-Provenance k (e₁ e₂) k∈ke₁e₂
with ∈k-dec k (proj₁ e₁ ) | ∈k-dec k (proj₁ e₂ )
... | yes k∈ke₁ | yes k∈ke₂ = {!!}
... | yes k∈ke₁ | no k∉ke₂ = {!!}
... | no k∉ke₁ | yes k∈ke₂ = {!!}
... | no k∉ke₁ | no k∉ke₂ = {!!}
module _ (_≈_ : B B Set b) where module _ (_≈_ : B B Set b) where
@ -416,13 +483,18 @@ module _ (_≈_ : B → B → Set b) where
(f : B B B) where (f : B B B) where
module I = ImplInsert f module I = ImplInsert f
-- The Provenance type requires both union and intersection functions,
-- but here we're working with union only. Just use the union function
-- for both -- it doesn't matter, since we don't use intersection in
-- these proofs.
module _ (f-idemp : (b : B) f b b b) where module _ (f-idemp : (b : B) f b b b) where
union-idemp : (m : Map) lift (union f m m) m union-idemp : (m : Map) lift (union f m m) m
union-idemp m@(l , u) = (mm-m-subset , m-mm-subset) union-idemp m@(l , u) = (mm-m-subset , m-mm-subset)
where where
mm-m-subset : subset (union f m m) m mm-m-subset : subset (union f m m) m
mm-m-subset k v k,v∈mm mm-m-subset k v k,v∈mm
with Expr-Provenance f k ((` m) (` m)) (∈-cong proj₁ k,v∈mm) with Expr-Provenance f f k ((` m) (` m)) (∈-cong proj₁ k,v∈mm)
... | (_ , (bothᵘ (single {v'} v'∈m) (single {v''} v''∈m) , v'v''∈mm)) ... | (_ , (bothᵘ (single {v'} v'∈m) (single {v''} v''∈m) , v'v''∈mm))
rewrite Map-functional {m = m} v'∈m v''∈m rewrite Map-functional {m = m} v'∈m v''∈m
rewrite Map-functional {m = union f m m} k,v∈mm v'v''∈mm = rewrite Map-functional {m = union f m m} k,v∈mm v'v''∈mm =
@ -439,7 +511,7 @@ module _ (_≈_ : B → B → Set b) where
where where
union-comm-subset : (m₁ m₂ : Map) subset (union f m₁ m₂) (union f m₂ m₁) union-comm-subset : (m₁ m₂ : Map) subset (union f m₁ m₂) (union f m₂ m₁)
union-comm-subset m₁@(l₁ , u₁) m₂@(l₂ , u₂) k v k,v∈m₁m₂ union-comm-subset m₁@(l₁ , u₁) m₂@(l₂ , u₂) k v k,v∈m₁m₂
with Expr-Provenance f k ((` m₁) (` m₂)) (∈-cong proj₁ k,v∈m₁m₂) with Expr-Provenance f f k ((` m₁) (` m₂)) (∈-cong proj₁ k,v∈m₁m₂)
... | (_ , (bothᵘ {v₁} {v₂} (single v₁∈m₁) (single v₂∈m₂) , v₁v₂∈m₁m₂)) ... | (_ , (bothᵘ {v₁} {v₂} (single v₁∈m₁) (single v₂∈m₂) , v₁v₂∈m₁m₂))
rewrite Map-functional {m = union f m₁ m₂} k,v∈m₁m₂ v₁v₂∈m₁m₂ = rewrite Map-functional {m = union f m₁ m₂} k,v∈m₁m₂ v₁v₂∈m₁m₂ =
(f v₂ v₁ , (f-comm v₁ v₂ , I.union-combines u₂ u₁ v₂∈m₂ v₁∈m₁)) (f v₂ v₁ , (f-comm v₁ v₂ , I.union-combines u₂ u₁ v₂∈m₂ v₁∈m₁))
@ -456,7 +528,7 @@ module _ (_≈_ : B → B → Set b) where
where where
union-assoc₁ : subset (union f (union f m₁ m₂) m₃) (union f m₁ (union f m₂ m₃)) union-assoc₁ : subset (union f (union f m₁ m₂) m₃) (union f m₁ (union f m₂ m₃))
union-assoc₁ k v k,v∈m₁₂m₃ union-assoc₁ k v k,v∈m₁₂m₃
with Expr-Provenance f k (((` m₁) (` m₂)) (` m₃)) (∈-cong proj₁ k,v∈m₁₂m₃) with Expr-Provenance f f k (((` m₁) (` m₂)) (` m₃)) (∈-cong proj₁ k,v∈m₁₂m₃)
... | (_ , (in k∉ke₁₂ (single {v₃} v₃∈e₃) , v₃∈m₁₂m₃)) ... | (_ , (in k∉ke₁₂ (single {v₃} v₃∈e₃) , v₃∈m₁₂m₃))
rewrite Map-functional {m = union f (union f m₁ m₂) m₃} k,v∈m₁₂m₃ v₃∈m₁₂m₃ = rewrite Map-functional {m = union f (union f m₁ m₂) m₃} k,v∈m₁₂m₃ v₃∈m₁₂m₃ =
let (k∉ke₁ , k∉ke₂) = I.∉-union-∉-either {l₁ = l₁} {l₂ = l₂} k∉ke₁₂ let (k∉ke₁ , k∉ke₂) = I.∉-union-∉-either {l₁ = l₁} {l₂ = l₂} k∉ke₁₂
@ -482,7 +554,7 @@ module _ (_≈_ : B → B → Set b) where
union-assoc₂ : subset (union f m₁ (union f m₂ m₃)) (union f (union f m₁ m₂) m₃) union-assoc₂ : subset (union f m₁ (union f m₂ m₃)) (union f (union f m₁ m₂) m₃)
union-assoc₂ k v k,v∈m₁m₂₃ union-assoc₂ k v k,v∈m₁m₂₃
with Expr-Provenance f k ((` m₁) ((` m₂) (` m₃))) (∈-cong proj₁ k,v∈m₁m₂₃) with Expr-Provenance f f k ((` m₁) ((` m₂) (` m₃))) (∈-cong proj₁ k,v∈m₁m₂₃)
... | (_ , (in k∉ke₁ (in k∉ke₂ (single {v₃} v₃∈e₃)) , v₃∈m₁m₂₃)) ... | (_ , (in k∉ke₁ (in k∉ke₂ (single {v₃} v₃∈e₃)) , v₃∈m₁m₂₃))
rewrite Map-functional {m = union f m₁ (union f m₂ m₃)} k,v∈m₁m₂₃ v₃∈m₁m₂₃ = rewrite Map-functional {m = union f m₁ (union f m₂ m₃)} k,v∈m₁m₂₃ v₃∈m₁m₂₃ =
(v₃ , (≈-refl , I.union-preserves-∈₂ (I.union-preserves-∉ k∉ke₁ k∉ke₂) v₃∈e₃)) (v₃ , (≈-refl , I.union-preserves-∈₂ (I.union-preserves-∉ k∉ke₁ k∉ke₂) v₃∈e₃))