Library contrib
Require Import List.
Require Import ZArith.
Require Import List.
Require Import EqNat.
Require Import Classical.
Require Import util.
Require Import util_tactic.
Require Import heap.
Require Import heap_tactic.
Require Import bipl.
Require Import axiomatic.
Open Local Scope heap_scope.
Open Local Scope sep_scope.
Load contrib_tactics.
Lemma singleton_equal: forall s h1 h2 e1 e2 e3 e4,
(e1 |-> e2) s h1 ->
(e3 |-> e4) s h2 ->
eval e1 s = eval e3 s ->
eval e2 s = eval e4 s ->
h1 = h2.
Lemma semax_lookup_simple: forall R e e' x s h,
((e |-> e') ** TT) s h ->
(update_store2 x e' R) s h ->
exists e0, ((e |-> e0) ** ((e |-> e0) -* (update_store2 x e0 R))) s h.
Lemma update_heap2_sep_con: forall x v P Q,
(update_heap2 x v P) ** Q ==> update_heap2 x v (P**Q).
Lemma update_heap2_sep_con': forall P Q x1 x2 v1 v2,
(update_heap2 x1 v1 P) ** (update_heap2 x2 v2 Q) ==>
update_heap2 x1 v1 (update_heap2 x2 v2 (P**Q)).
Lemma semax_assign' :forall Q P x e,
(P ==> update_store2 x e Q) ->
{{ P }} x <- e {{ Q }}.
Lemma semax_assign'' : forall R P Q x e c,
(P ==> update_store2 x e R) ->
{{ R }} c {{ Q }} ->
{{ P }} x <- e ; c {{ Q }}.
Lemma semax_lookup' : forall Q P x e,
(P ==> lookup2 x e Q) ->
{{ P }} x <-* e {{ Q }}.
Lemma semax_lookup'' : forall R P Q x e c,
(P ==> lookup2 x e R) ->
{{ R }} c {{ Q }} ->
{{ P }} x <-* e ; c {{ Q }}.
Lemma semax_lookup_backwards' : forall P Q x e,
(P ==> (fun s h => exists e0, ((e |-> e0) ** ((e |-> e0) -* (update_store2 x e0 Q))) s h)) ->
{{ P }} x <-* e {{ Q }}.
Lemma semax_lookup_backwards'' : forall P Q R x e c,
(P ==> (fun s h =>
exists e0, ((e |-> e0) ** ((e |-> e0) -* (update_store2 x e0 R))) s h)) ->
{{ R }} c {{ Q }} ->
{{ P }} x <-* e; c {{ Q }}.
Lemma cell_read: forall e v Q s h,
((e |-> v) ** TT) s h /\ Q s h ->
((e |-> v) ** ((e |-> v) -* Q)) s h.
Lemma semax_lookup_backwards''2 : forall P Q R x e c,
(P ==> (fun s h => exists e0, ((e |-> e0) ** TT) s h /\ (update_store2 x e0 R) s h)) ->
{{ R }} c {{ Q }} ->
{{ P }} x <-* e; c {{ Q }}.
Lemma semax_mutation' : forall Q P e1 e2,
(P ==> update_heap2 e1 e2 Q) ->
{{ P }} e1 *<- e2 {{ Q }}.
Lemma semax_mutation'' : forall R P Q c e1 e2,
(P ==> update_heap2 e1 e2 R) ->
{{ R }} c {{ Q }} ->
{{ P }} e1 *<- e2 ; c {{ Q }}.
Lemma semax_mutation_backwards' : forall Q P e1 e2,
(P ==> (fun s h => exists e'', ((e1 |-> e'') ** ((e1 |-> e2) -* Q)) s h)) ->
{{ P }} e1 *<- e2 {{ Q }}.
Lemma semax_mutation_backwards'' : forall Q P R e1 e2 c,
(P ==> (fun s h => exists e'', ((e1 |-> e'') ** ((e1 |-> e2) -* R)) s h)) ->
{{ R }} c {{ Q }} ->
{{ P }} e1 *<- e2 ; c {{ Q }}.
Lemma semax_while' : forall b c I P Q,
(P ==> I) ->
{{ fun s h => I s h /\ eval_b b s = true }} c {{ I }} ->
((fun s h => I s h /\ eval_b b s = false) ==> Q) ->
{{ P }} while b c {{ Q }}.
Lemma semax_while'': forall b c I P Q c1,
(P ==> I) ->
{{ fun s h => I s h /\ eval_b b s = true }} c {{ I }} ->
{{ fun s h => I s h /\ eval_b b s = false }} c1 {{ Q }} ->
{{ P }} while b c; c1 {{ Q }}.
Lemma semax_skip': forall P Q,
(P ==> Q) ->
{{ P }} skip {{ Q }}.
Lemma apply_triple: forall (P P' Q Q': assert) c,
{{P'}} c {{Q'}} ->
(P ==> P') ->
(Q' ==> Q) ->
{{P}} c {{Q}}.
Lemma apply_triple': forall (P P' Q Q': assert) c c',
{{P'}} c {{Q'}} ->
(P ==> P') ->
{{Q'}} c' {{Q}} ->
{{P}} c; c' {{Q}}.
Lemma semax_and' : forall c P Q, {{ P }} c {{ Q }} ->
forall P' Q' d, {{ P' }} d {{ Q' }} ->
c = d ->
{{ And P P' }} c {{ And Q Q' }}.
Lemma semax_and : forall c P Q P' Q',
{{ P }} c {{ Q }} ->
{{ P' }} c {{ Q' }} ->
{{ And P P' }} c {{ And Q Q' }}.
Lemma forward_reasoning : forall x e P,
inde (x::nil) P ->
~ In x (expr_var e) ->
{{ P }} x <- e {{ fun s h => P s h /\ eval (var_e x) s = eval e s }}.
Lemma frame_rule': forall P c Q,
{{ P }} c {{ Q }} ->
forall R1 R2,
inde (modified_cmd_var c) R1 ->
(R1 ==> R2) ->
{{ P ** R1 }} c {{ Q ** R2 }}.
Lemma frame_rule'': forall c1 c2 P1 P2 Q1 Q2,
{{ P1 }} c1 {{ Q1 }} ->
{{ P2 }} c2 {{ Q2 }} ->
inde (modified_cmd_var c1) P2 ->
inde (modified_cmd_var c2) Q1 ->
{{ P1 ** P2 }} c1 ; c2 {{ Q1 ** Q2 }}.
Lemma semax_mutation_frame_rule : forall x v c P Q v',
{{ P }} c {{ Q }} ->
inde (modified_cmd_var c) (x |-> v) ->
{{ (x |-> v') ** P }} x *<- v ; c {{ (x |-> v) ** Q }}.
Fixpoint Array (p:loc) (size:nat) {struct size} : assert :=
match size with
O => sep.emp
| S n => (fun s h => exists y, (nat_e p |-> int_e y) s h) **
Array (p + 1) n
end.
Lemma Array_inde_store : forall m n s h,
Array n m s h -> forall s', Array n m s' h.
Ltac Array_equiv :=
match goal with
| id: Array ?adr1 ?size1 ?s1 ?h |- Array ?adr2 ?size2 ?s2 ?h =>
assert (Array_equivA1: adr2 = adr1); [omega |
assert (Array_equivA2: size2 = size1); [omega |
((rewrite Array_equivA1) || idtac); ((rewrite Array_equivA2) || idtac);
eapply (Array_inde_store); apply id
]
]
end.
Lemma Array_inde: forall l adr size,
inde l (Array adr size).
Lemma Array_inde_list: forall l st sz,
inde l (Array st sz).
Lemma Array_concat_split: forall size1 size2 adr,
Array adr (size1+size2) <==> (Array adr size1 ** Array (adr+size1) size2).
Ltac TArray_concat_split_r size1 size2:=
match goal with
| |- Array ?adr ?size ?s ?h =>
generalize (Array_concat_split size1 size2 adr s h); intro TArray_concat_split_rH1; inversion_clear TArray_concat_split_rH1 as [TArray_concat_split_rH11 TArray_concat_split_rH12];
assert (TArray_concat_split_rA1: size1 + size2 = size); [omega | (rewrite <- TArray_concat_split_rA1 || idtac); clear TArray_concat_split_rA1; apply TArray_concat_split_rH12; clear TArray_concat_split_rH11 TArray_concat_split_rH12]
end.
Ltac TArray_concat_split_l_l sz id:=
match goal with
| id: Array ?adr ?size ?s ?h |- _ =>
assert (TArray_concat_split_l_lA1: size = sz + (size - sz));
[omega |
rewrite TArray_concat_split_l_lA1 in id; clear TArray_concat_split_l_lA1;
generalize (Array_concat_split sz (size - sz) adr s h); intro TArray_concat_split_l_lH1;
inversion_clear TArray_concat_split_l_lH1 as [TArray_concat_split_l_lH2 TArray_concat_split_l_lH3]; generalize (TArray_concat_split_l_lH2 id); clear TArray_concat_split_l_lH3; clear TArray_concat_split_l_lH2; intro
]
end.
Ltac TArray_concat_split_l_r sz id:=
match goal with
| id: Array ?adr ?size ?s ?h |- _ =>
assert (TArray_concat_split_l_lA1: size = (size - sz) + sz);
[omega |
rewrite TArray_concat_split_l_lA1 in id; clear TArray_concat_split_l_lA1;
generalize (Array_concat_split (size - sz) sz adr s h); intro TArray_concat_split_l_lH1;
inversion_clear TArray_concat_split_l_lH1 as [TArray_concat_split_l_lH2 TArray_concat_split_l_lH3]; generalize (TArray_concat_split_l_lH2 id); clear TArray_concat_split_l_lH3; clear TArray_concat_split_l_lH2; intro
]
end.
Lemma Array_split : forall size' size,
size' <= size ->
forall adr s h, adr > 0 -> Array adr size s h ->
(Array adr size' ** Array (adr+size') (size - size')) s h.
Lemma Array_disjoint : forall size adr s h, adr > 0 ->
Array adr size s h ->
forall adr',
adr' >= adr+size ->
forall z,
h # (heap.singleton adr' z).
Lemma Array_concat : forall size adr s h, adr > 0 ->
Array adr size s h ->
forall z,
Array adr (size + 1) s (h +++ (heap.singleton (adr+size) z)).
Lemma mapstos_to_array: forall l x x' sz s h,
(x |--> l) s h ->
eval x s = Z_of_nat x' ->
sz = length l ->
Array x' sz s h.
Fixpoint ArrayI (x:loc) (l:list Z) {struct l} : assert :=
match l with
nil => sep.emp
| hd::tl =>
(fun s h => (nat_e x |-> int_e hd) s h) **
ArrayI (x + 1) tl
end.
Lemma ArrayI_inde_store: forall l n s h,
ArrayI n l s h -> forall s', ArrayI n l s' h.
Lemma ArrayI_disjoint_heap: forall lst a h' e e' s h,
ArrayI a lst s h ->
(nat_e e |-> e') s h' ->
(e < a \/ e >= a + length lst) ->
h # h'.
Lemma Data_destructive_update_inv2: forall lx x h1 h2 e1 x1 s h,
(ArrayI x lx ** TT) s h ->
h1 # h2 ->
h = h1 +++ h2 ->
(nat_e e1 |-> x1) s h1 ->
(e1 < x \/ e1 >= x + length lx) ->
(ArrayI x lx ** TT) s h2.
Lemma Data_destructive_update_inv: forall lx x h1 h2 e1 x1 x2 h'1 h' s h,
(ArrayI x lx ** TT) s h ->
h1 # h2 ->
h = h1 +++ h2 ->
(nat_e e1 |-> x1) s h1 ->
(nat_e e1 |-> x2) s h'1 ->
h'1 # h2 ->
h' = h'1 +++ h2 ->
(e1 < x \/ e1 >= x + length lx) ->
(ArrayI x lx ** TT) s h'.