Library heap
Require Import ZArith.
Require Import EqNat.
Require Import Bool.
Require Import ZArith.
Require Import List.
Require Import Classical.
Require Import util.
Module Type MAP.
Parameter l : Set.
Parameter v : Set.
Parameter h : Set.
Parameter emp : h.
Parameter singleton : l -> v -> h.
Parameter lookup : l -> h -> option v.
Parameter lookup_emp : forall x, lookup x emp = None.
Parameter lookup_singleton : forall n w, lookup n (singleton n w) = Some w.
Parameter extensionality: forall h1 h2, (forall x, lookup x h1 = lookup x h2) -> h1 = h2.
Parameter update : l -> v -> h -> h.
Parameter update_singleton : forall n w w', update n w' (singleton n w) = singleton n w'.
Parameter union : h -> h -> h.
Notation "k '+++' m" := (union k m) (at level 69) : heap_scope.
Open Local Scope heap_scope.
Parameter equal_union_emp : forall h1, h1 +++ emp = h1.
Parameter union_assoc : forall h1 h2 h3, h1 +++ (h2 +++ h3) = (h1 +++ h2) +++ h3.
Parameter lookup_union_or: forall h1 h2 n z,
lookup n (h1 +++ h2) = Some z ->
lookup n h1 = Some z \/ lookup n h2 = Some z.
Parameter lookup_update : forall x y z st, x <> y -> lookup x st = lookup x (singleton y z +++ st).
Parameter lookup_update' : forall x z st, lookup x (singleton x z +++ st) = Some z.
Parameter lookup_update'' : forall x y z st,
x <> y -> lookup x st = lookup x (update y z st).
Parameter disjoint : h -> h -> Prop.
Notation "k '#' m" := (disjoint k m) (at level 79) : heap_scope.
Parameter disjoint_emp : forall h, h # emp.
Parameter disjoint_singleton : forall x y z z', x <> y -> (singleton x z) # (singleton y z').
Parameter disjoint_singleton' : forall x y z z', (singleton x z) # (singleton y z') -> x <> y.
Parameter disjoint_com : forall h0 h1, h0 # h1 -> h1 # h0.
Parameter disjoint_update : forall n z h1 h2, h1 # h2 -> (update n z h1) # h2.
Parameter disjoint_union : forall h1 h2 h0, h0 # h1 /\ h0 # h2 -> h0 # (h1 +++ h2).
Parameter disjoint_union' : forall h1 h2 h0, h0 # (h1 +++ h2) -> h0 # h1 /\ h0 # h2.
Parameter equal_union_disjoint : forall x1 x2 x0, (x1 +++ x2) # x0 -> x1 # x0.
Parameter union_com : forall h1 h2, h1 # h2 -> h1 +++ h2 = h2 +++ h1.
Parameter lookup_union : forall h1 h2, h1 # h2 ->
forall n z, lookup n h1 = Some z ->
lookup n (h1 +++ h2) = Some z.
Parameter lookup_union_R : forall h1 h2, h1 # h2 ->
forall n z, lookup n h2 = Some z ->
lookup n (h1 +++ h2) = Some z.
Parameter equal_update_L : forall h1 h2, h1 # h2 ->
forall n z z1, lookup n h1 = Some z1 ->
update n z (h1 +++ h2) = (update n z h1) +++ h2.
Parameter equal_update_R : forall (h1 h2:h), h1 # h2 ->
forall (n:l) (z z2:v), lookup n h2 = Some z2 ->
update n z (h1 +++ h2) = h1 +++ (update n z h2).
Parameter lookup_exists_partition: forall h a b,
lookup a h = Some b ->
exists h'' , h = (singleton a b) +++ h'' /\ (singleton a b) # h''.
Parameter disjoint_comp : forall h'1 h1 h2 h'2,
h1 # h'1 -> h1 # h2 -> h'1 # h'2 ->
h'1 +++ h'2 = h1 +++ h2 ->
exists h', h'1 # h' /\ h2 = (h' +++ h'1).
Parameter dif : h -> l -> h.
Notation "k '---' l" := (dif k l) (at level 69) : heap_scope.
Parameter dif_singleton_emp: forall a b, (singleton a b) --- a = emp.
Parameter dif_union: forall h1 h2 a, (h1 +++ h2) --- a = (h1 --- a) +++ (h2 --- a).
Parameter dif_disjoint: forall h a b, h # (singleton a b) -> (h --- a) = h.
Parameter dif_disjoint': forall h1 h2 l, h1 # h2 -> (h1 --- l) # (h2 --- l).
Parameter lookup_dif : forall w st, lookup w (st --- w) = None.
Parameter lookup_dif' : forall st x y, x <> y -> lookup x (st --- y) = lookup x st.
End MAP.
Module Type ELEM.
Parameter elem : Set.
End ELEM.
Require OrderedType.
Module Type OrderedTypeExt.
Parameter t : Set.
Parameter eq : t -> t -> Prop.
Parameter lt : t -> t -> Prop.
Axiom eq_refl : forall x : t, eq x x.
Axiom eq_sym : forall x y : t, eq x y -> eq y x.
Axiom eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z.
Axiom lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z.
Axiom lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
Parameter compare : forall x y : t, OrderedType.Compare lt eq x y.
Hint Immediate eq_sym.
Hint Resolve eq_refl eq_trans lt_not_eq lt_trans.
Axiom eq_ext : forall x y : t, eq x y -> x = y.
End OrderedTypeExt.
Module map (A: OrderedTypeExt) (E : ELEM) : MAP
with Definition l := A.t
with Definition v := E.elem.
Definition l := A.t.
Definition v := E.elem.
Notation "a < b" := (A.lt a b) : addr_scope.
Open Local Scope addr_scope.
Delimit Scope addr_scope with addr_scope.
Module otf := OrderedType.OrderedTypeFacts A.
Definition mygt a b := (b < a)%addr_scope.
Notation "a > b" := (mygt a b) : addr_scope.
Definition myle x y := x < y \/ A.eq x y.
Notation "a <= b" := (myle a b) : addr_scope.
Definition myge a b := (b <= a)%addr_scope.
Notation "a >= b" := (myge a b) : addr_scope.
Lemma mytrichotomy : forall x y : l, (x < y \/ x = y \/ x > y)%addr_scope.
Lemma gt_neq : forall a b : l, (a < b -> a <> b)%addr_scope.
Lemma mylt_neq : forall a b : l, (b < a -> a <> b)%addr_scope.
Lemma myleq_neq_lt : forall a b : l, (a <= b -> a <> b -> a < b)%addr_scope.
Lemma myle_ge_eq : forall a b:l, (a <= b -> b <= a -> a = b)%addr_scope.
Lemma mylt_trans2 : forall n m p : l, (n <= m -> m < p -> n < p)%addr_scope.
Lemma myle_eq_lt : forall a b, (a <= b -> a = b \/ a < b)%addr_scope.
Definition addr_lt (x y:l) :=
match A.compare x y with
| OrderedType.LT _ => true
| _ => false
end.
Lemma addr_lt_false' : forall n m : l, (n >= m -> addr_lt n m = false)%addr_scope.
Lemma addr_lt_classic : forall n m : l, addr_lt n m = true \/ addr_lt n m = false.
Lemma addr_lt_false : forall n m : l, addr_lt n m = false -> (n >= m)%addr_scope.
Lemma addr_lt_true : forall n m : l, addr_lt n m = true -> (n < m)%addr_scope.
Lemma addr_lt_true' : forall n m : l, (n < m)%addr_scope -> addr_lt n m = true.
Lemma addr_lt_irrefl : forall n : l, addr_lt n n = false.
Lemma addr_lt_assym : forall n m : l, addr_lt n m = true -> addr_lt m n = false.
Definition beq_addr (x y:l) :=
match (A.compare x y) with
| OrderedType.EQ _ => true
| _ => false
end.
Lemma beq_dif_false_addr : forall n m : l, n <> m -> beq_addr n m = false.
Lemma beq_addr_refl : forall n : l, true = beq_addr n n.
Lemma beq_addr_classic : forall a b : l, beq_addr a b = true \/ beq_addr a b = false.
Lemma beq_addr_true : forall x y : l, beq_addr x y = true -> x = y.
Lemma beq_addr_false : forall a b : l, beq_addr a b = false -> a <> b.
Lemma beq_addr_com : forall n m : l, beq_addr n m = beq_addr m n.
Fixpoint lb (loc:l) (he:list l) {struct he} : Prop :=
match he with
| nil => True
| hd::tl => loc < hd /\ lb loc tl
end.
Lemma lb_notin : forall k a, lb a k -> ~ In a k.
Lemma lb_notin' : forall k a, In a k -> ~ lb a k.
Lemma lb_In_uneq: forall k a c, In a k -> lb c k -> beq_addr a c = false.
Lemma lb_trans: forall k n m , lb m k -> n < m -> lb n k.
Fixpoint set (locs:list l) : Prop :=
match locs with
nil => True
| hd::tl => lb hd tl /\ set tl
end.
Definition dom (k:list (l*v)) := map (fun x => fst x) k.
Inductive h' : Set := mk_h : forall lst, set (dom lst) -> h'.
Definition h := h'.
Definition lst := fun k:h => let (lst, _) := k in lst.
Definition prf := fun k:h =>
let (lst, va) as h return (set (dom (lst h))) := k in va.
Definition set_nil : set nil.
Definition emp := mk_h nil set_nil.
Definition set_singleton : forall n v, set (dom ((n,v)::nil)).
Definition singleton := fun (n:l) (w:v) =>
mk_h ((n, w)::nil) (set_singleton n w).
Fixpoint lookup' (n:l) (k:list (l*v)) {struct k} : option v :=
match k with
nil => None
| hd::tl => match beq_addr n (fst hd) with
true => Some (snd hd)
| false => lookup' n tl
end
end.
Definition lookup n k := lookup' n (lst k).
Lemma lookup_emp : forall x, lookup x emp = None.
Lemma lookup_singleton : forall n w, lookup n (singleton n w) = Some w.
Lemma lookup'_In : forall (k:list (l*v)) (n:l) (z:v),
lookup' n k = Some z -> In n (dom k).
Lemma lookup_not_In: forall h x, ~ In x (dom (lst h)) -> lookup x h = None.
Definition equal (h1: h) (h2: h) := (lst h1) = (lst h2).
Notation "k '===' m" := (equal k m) (at level 79) : heap_scope.
Open Local Scope heap_scope.
Lemma equal_ext: forall h1 h2, h1 === h2 -> h1 = h2.
Lemma equal_ext': forall h1 h2, h1 = h2 -> h1 === h2.
Lemma extensionality: forall h1 h2,
(forall x, lookup x h1 = lookup x h2) -> h1 = h2.
Fixpoint update' (n:l) (w:v) (k:list (l*v)) {struct k} : list (l*v) :=
match k with
nil => k
| hd::tl => match beq_addr n (fst hd) with
true => (fst hd, w)::tl
| false => hd::(update' n w tl)
end
end.
Lemma lb_update : forall (n:l) (k:list (l*v)), lb n (dom k) ->
forall (m:l) (w:v), lb n (dom (update' m w k)).
Lemma set_update: forall (k:list (l*v)), set (dom k) ->
forall (n:l) (w:v), set (dom (update' n w k)).
Definition update (n:l) (w:v) (k:h) := mk_h (update' n w (lst k)) (set_update (lst k) (prf k) n w).
Lemma lookup'_update' : forall (k:list (l*v)) (x y:l) (z:v), x <> y ->
lookup' x k = lookup' x (update' y z k).
Lemma lookup_update'' : forall x y z st,
x <> y -> lookup x st = lookup x (update y z st).
Lemma update_dom : forall (k:h) (n:l) (z:v),
dom (lst (update n z k)) = dom (lst k).
Lemma update_singleton : forall (n:l) (w w':v),
update n w' (singleton n w) = singleton n w'.
Fixpoint add_lst (n:l) (w:v) (lst:list (l*v)) {struct lst}: list (l*v) :=
match lst with
nil => (n,w)::nil
| (n',w')::tl => match addr_lt n' n with
true => (n',w') :: (add_lst n w tl)
| false => match beq_addr n' n with
true => (n',w) :: tl
| false => (n,w) :: (n',w')::tl
end
end
end.
Lemma lb_add_lst : forall (k:list (l*v)) (n:l), lb n (dom k) ->
forall m, n < m -> forall w, lb n (dom (add_lst m w k)).
Lemma add_lst_lb: forall k a b, lb a (dom k) -> add_lst a b k = (a, b)::k.
Lemma set_add : forall lst, set (dom lst) ->
forall n w, set (dom (add_lst n w lst)).
Lemma In_dom_add_lst: forall k a b, In a (dom (add_lst a b k)).
Lemma In_dom_add_lst': forall k x a b, In x (dom k) ->
In x (dom (add_lst a b k)).
Lemma add_lst'_In_dom: forall k x a b, In x (dom (add_lst a b k)) -> x = a \/ In x (dom k).
Lemma In_dom_add_lst2: forall k n w n', n <> n' ->
In n' (dom (add_lst n w k)) -> In n' (dom k).
Lemma add_lst_add_lst': forall k n w' w, add_lst n w (add_lst n w' k) = add_lst n w k.
Lemma add_lst_add_lst: forall k n' w' n w, n' <> n ->
add_lst n w (add_lst n' w' k) = add_lst n' w' (add_lst n w k).
Lemma update_add_lst: forall k n w w', update' n w (add_lst n w' k) = add_lst n w k.
Lemma update_add_lst': forall k n n' w w', n' <> n -> update' n w (add_lst n' w' k) = add_lst n' w' (update' n w k).
Lemma lookup'_add_lst : forall k n w, lookup' n (add_lst n w k) = Some w.
Lemma lookup'_add_lst' : forall k n w n' , n <> n' -> lookup' n' (add_lst n w k) = lookup' n' k.
Fixpoint app' (h1 h2:list (l*v)) {struct h1} : list (l*v) :=
match h1 with
nil => h2
| hd::tl =>
add_lst (fst hd) (snd hd) (app' tl h2)
end.
Lemma set_app' : forall h1 h2, set (dom h1) -> set (dom h2) ->
set (dom (app' h1 h2)).
Definition union h1 h2 :=
mk_h (app' (lst h1) (lst h2)) (set_app' _ _ (prf h1) (prf h2)).
Notation "k '+++' m" := (union k m) (at level 69) : heap_scope.
Lemma app'_nil : forall k, set (dom k) -> app' k nil = k.
Lemma add_lst_app': forall lst1 lst2 n w,
add_lst n w (app' lst1 lst2) = app' (add_lst n w lst1) lst2.
Lemma app'_com : forall l1 l2, set (dom l1) -> set (dom l2) ->
inter (dom l1) (dom l2) nil ->
app' l1 l2 = app' l2 l1.
Lemma update'_app' : forall (l1 l2:list (l*v)) (n:l) (z:v),
In n (dom l1) /\ ~ In n (dom l2) ->
app' (update' n z l1) l2 = update' n z (app' l1 l2).
Lemma lookup'_app'_L : forall l1 l2 n z,
lookup' n l1 = Some z ->
lookup' n (app' l1 l2) = Some z.
Lemma lookup'_app'_R : forall l1 l2 n z,
inter (dom l1) (dom l2) nil ->
lookup' n l2 = Some z ->
lookup' n (app' l1 l2) = Some z.
Lemma In_dom_union : forall k x, In x (dom (lst k)) ->
forall m, In x (dom (lst (k +++ m))).
Lemma equal_union_emp : forall h1, h1 +++ emp = h1.
Lemma In_dom_union' : forall m k x, In x (dom (lst k)) ->
In x (dom (lst (m +++ k))).
Lemma union_assoc : forall h1 h2 h3, h1 +++ (h2 +++ h3) = (h1 +++ h2) +++ h3.
Lemma In_union_dom : forall k m x, In x (dom (lst (k +++ m))) ->
In x (dom (lst k)) \/ In x (dom (lst m)).
Lemma lookup_union_or: forall h1 h2 n z,
lookup n (h1 +++ h2) = Some z ->
lookup n h1 = Some z \/ lookup n h2 = Some z.
Lemma lookup_update : forall x y z st, x <> y -> lookup x st = lookup x (singleton y z +++ st).
Lemma lookup_update' : forall x z st, lookup x (singleton x z +++ st) = Some z.
Definition disjoint h1 h2 := inter (dom (lst h1)) (dom (lst h2)) nil.
Notation "k '#' m" := (disjoint k m) (at level 79) : heap_scope.
Lemma disjoint_emp : forall h, h # emp.
Lemma disjoint_singleton : forall x y z z',
x <> y -> singleton x z # singleton y z'.
Lemma disjoint_singleton' : forall x y z z',
singleton x z # singleton y z' -> x <> y.
Lemma disjoint_com: forall h1 h2, h1 # h2 -> h2 # h1.
Lemma disjoint_update : forall n z h1 h2,
h1 # h2 -> (update n z h1) # h2.
Lemma distributivity : forall h1 h2 h0,
h0 # (h1 +++ h2) <-> h0 # h1 /\ h0 # h2.
Lemma disjoint_union : forall h1 h2 h0,
h0 # h1 /\ h0 # h2 -> h0 # (h1 +++ h2).
Lemma disjoint_union' : forall h1 h2 h0, h0 # (h1 +++ h2) -> h0 # h1 /\ h0 # h2.
Lemma equal_union_disjoint : forall x1 x2 x0, (x1 +++ x2) # x0 -> x1 # x0.
Lemma union_com : forall h1 h2, h1 # h2 -> h1 +++ h2 = h2 +++ h1.
Lemma lookup_union : forall h1 h2, h1 # h2 ->
forall n z, lookup n h1 = Some z ->
lookup n (h1 +++ h2) = Some z.
Lemma lookup_union_R : forall h1 h2, h1 # h2 ->
forall n z, lookup n h2 = Some z ->
lookup n (h1 +++ h2) = Some z.
Lemma equal_update_L : forall (h1 h2:h), h1 # h2 ->
forall (n:l) (z z1:v), lookup n h1 = Some z1 ->
update n z (h1 +++ h2) = (update n z h1) +++ h2.
Lemma equal_update_R : forall (h1 h2:h), h1 # h2 ->
forall (n:l) (z z2:v), lookup n h2 = Some z2 ->
update n z (h1 +++ h2) = h1 +++ (update n z h2).
Lemma lookup_exists_partition : forall h a b,
lookup a h = Some b ->
exists h'' ,
h = (singleton a b) +++ h'' /\ (singleton a b) # h''.
Fixpoint del' (n:l) (k:list (l*v)) {struct k} : list (l*v) :=
match k with
nil => nil
| hd::tl => match beq_addr n (fst hd) with
true => del' n tl
| false => hd :: del' n tl
end
end.
Lemma lb_del' : forall k n, lb n (dom k) ->
forall m, lb n (dom (del' m k)).
Lemma lb_del'' : forall k n, lb n (dom k) -> del' n k = k.
Lemma set_del' : forall k, set (dom k) -> forall n, set (dom (del' n k)).
Definition dif h l := mk_h (del' l (lst h) ) (set_del' (lst h) (prf h) l).
Notation "k '---' m" := (dif k m) (at level 69) : heap_scope.
Lemma dif_singleton_emp: forall a b, (singleton a b) --- a = emp.
Lemma add_lst_del': forall lst n w n', set (dom lst) ->
n' <> n -> add_lst n w (del' n' lst) = del' n' (add_lst n w lst).
Lemma add_lst_del'': forall lst n w, del' n (add_lst n w lst) = del' n lst.
Lemma lookup'_del' : forall w st, lookup' w (del' w st) = None.
Lemma lookup'_del'' : forall st x y, x <> y ->
lookup' x (del' y st) = lookup' x st.
Lemma del'_app' : forall k m n, set (dom m) -> set (dom k) ->
del' n (app' k m) = app' (del' n k) (del' n m).
Lemma dif_not_in_dom: forall h l,
~(In l (dom (lst h))) ->
(h --- l) = h.
Lemma dif_union: forall h1 h2 a,
(h1 +++ h2) --- a = (h1 --- a) +++ (h2 --- a).
Lemma dif_disjoint: forall h a b, h # (singleton a b) -> h --- a = h.
Lemma In_dom_del' : forall k x y, x <> y ->
In x (dom (del' y k)) ->
In x (dom k).
Lemma dif_disjoint': forall h1 h2 l, h1 # h2 -> h1 --- l # h2 --- l.
Lemma compose_union_equiv: forall h a b (s: set (dom ((a,b)::h)))
(s': set (dom h)) ,
mk_h ((a,b) :: h) s === (singleton a b) +++ (mk_h h s').
Lemma lookup_disjoint_None: forall h a b,
h # singleton a b -> lookup a h = None.
Lemma disjoint_comp : forall h'1 h1 h2 h'2,
h1 # h'1 ->
h1 # h2 ->
h'1 # h'2 ->
h'1 +++ h'2 = h1 +++ h2 ->
exists h', h'1 # h' /\ h2 = (h' +++ h'1).
Lemma lookup_dif : forall w st, lookup w (st --- w) = None.
Lemma lookup_dif' : forall st x y, x <> y ->
lookup x (st --- y) = lookup x st.
End map.
Module Loc <: OrderedTypeExt.
Definition t := nat.
Definition eq := @eq nat.
Definition eq_refl := @refl_equal t.
Definition eq_sym := @sym_eq t.
Definition eq_trans := @trans_eq t.
Definition lt := lt.
Lemma lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z.
Lemma lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
Definition compare : forall x y : t, OrderedType.Compare lt eq x y.
Lemma eq_ext : forall x y : t, eq x y -> x = y.
End Loc.
Module integer.
Definition elem := Z.
End integer.
Module heap := map Loc integer.
Definition loc := Loc.t.
Definition val := Z.
Definition val2loc (z:Z) : loc :=
match z with
Z0 => O
| Zpos p => nat_of_P p
| Zneg p => nat_of_P p
end.
Definition loc2val (p:loc) : val := Z_of_nat p.
Lemma Z_of_nat2loc : forall n, val2loc (loc2val n) = n.
Notation "k '---' l" := (heap.dif k l) (at level 69) : heap_scope.
Notation "k '+++' m" := (heap.union k m) (at level 69) : heap_scope.
Notation "k '#' m" := (heap.disjoint k m) (at level 79) : heap_scope.