Library schorr_waite
Load seplog_header.
Notation " s '|b=' b " := (eval_b b s = true) (right associativity, at level 80).
Notation " s '|b!=' b " := (eval_b b s = false) (right associativity, at level 80).
Ltac Step R :=
match goal with
| id: {{?P'}} ?c {{?Q'}} |- {{?P}} ?c {{?Q}} => eapply apply_triple; [apply id | idtac | idtac]
| id: {{?P'}} ?c {{?Q'}} |- {{?P}} ?c; ?c' {{?Q}} => eapply apply_triple'; [apply id | idtac | idtac]
| |- {{?P}} ?x <- ?e {{?Q}} => eapply semax_assign'; red
| |- {{?P}} ?x <- ?e ; ?c {{?Q}} => eapply semax_assign'' with R; [red; intros | idtac]
| |- {{?P}} ?x <-* ?e {{?Q}} => eapply semax_lookup_backwards'
| |- {{?P}} ?x <-* ?e ; ?c {{?Q}} => eapply semax_lookup_backwards'' with R; [red; intros | idtac]
| |- {{?P}} ?e1 *<- ?e2 {{?Q}} => eapply semax_mutation_backwards'
| |- {{?P}} ?e1 *<- ?e2 ; ?c {{?Q}} => eapply semax_mutation_backwards'' with R; [red; intros | idtac]
| |- {{?P}} while ?b ?c1 {{?Q}} => eapply semax_while' with R
| |- {{?P}} while ?b ?c1 ; ?c2 {{?Q}} => eapply semax_while'' with R; [red; intros | idtac | idtac]
| |- {{?P}} skip {{?Q}} => eapply semax_skip'
| |- {{?P}} ifte ?b thendo ?c1 elsedo ?c2 {{?Q}} => eapply semax_ifte
| |- {{?P}} (ifte ?b thendo ?c1 elsedo ?c2); ?c' {{?Q}} => apply semax_seq with R; [eapply semax_ifte; [idtac| idtac] | idtac]
| |- {{?P}} ?c1; ?c2 {{?Q}} => apply semax_seq with R; [idtac| idtac]
end.
Ltac Decompose_hyp :=
match goal with
| id: ?P1 /\ ?P2 |- _ => (decompose [and] id; clear id); Decompose_hyp
| id: (?P1 ** ?P2) ?s ?h |- _ =>
(let x:=fresh in (
inversion_clear id as [x X];
let y:= fresh in (
inversion_clear X as [y Y];
decompose [and] Y; clear Y
)
)); Decompose_hyp
| id: ex ?P |- _ => (inversion_clear id); Decompose_hyp
| _ => idtac
end.
Lemma expr_store_update_rewrite: forall e x p s,
eval (expr_rewrite e (var_e x) (int_e p)) s = eval e (store.update x p s).
Lemma expr_b_store_update_rewrite: forall b x p s,
eval_b (expr_b_rewrite b (var_e x) (int_e p)) s = eval_b b (store.update x p s).
Lemma mapsto_store_update_rewrite: forall e1 e2 x p s h,
((expr_rewrite e1 (var_e x) (int_e p)) |-> (expr_rewrite e2 (var_e x) (int_e p))) s h ->
(e1 |-> e2) (store.update x p s) h.
Definition l := 0%Z.
Definition r := 1%Z.
Definition c := 2%Z.
Definition m := 3%Z.
Hint Unfold l.
Hint Unfold r.
Hint Unfold c.
Hint Unfold m.
Ltac Unfolds_fields :=
unfold l; unfold r; unfold c; unfold m.
Definition SW (t p q:var.v) (root: nat):=
(t <- (nat_e root));
(p <- (nat_e 0));
(q <-* (t -.> m));
while (((var_e p) =/= (nat_e 0)) ||| (((var_e t) =/= (nat_e 0)) &&& ((var_e q) =/= (nat_e 1)))) (
(ifte (((var_e t) == (nat_e 0)) ||| ((var_e q) == (nat_e 1))) thendo
(
(q <-* (p -.> c));
ifte ((var_e q) == (nat_e 1)) thendo (
(q <- (var_e t)); (t <- (var_e p)); (p <-* (p -.> r)); ((t -.>r) *<- (var_e q))
) elsedo (
(q <- (var_e t)); (t <-* (p -.> r)); ((p -.> r) *<- (p -.> l)); ((p -.> l) *<- (var_e q)); ((p -.> c) *<- (nat_e 1))
)
) elsedo (
(q <- (var_e p)); (p <- (var_e t)); (t <-* (t -.> l)); ((p -.> l) *<- (var_e q)); ((p -.> m) *<- (nat_e 1)); ((p -.> c) *<- (nat_e 0))
)
); (q <-* (t -.> m))
).
Inductive Graph: list (nat * nat * nat) -> list (nat * nat * nat) -> store.s -> heap.h -> Prop :=
empty_N: forall s h, sep.emp s h -> Graph nil nil s h
| empty_E: forall n l s h1 h2 h c m,
n <> 0 ->
h1 # h2 ->
h = h1 +++ h2 ->
((nat_e n) |--> ((nat_e 0)::(nat_e 0)::(nat_e c)::(nat_e m)::nil)) s h1 ->
Graph l nil s h2 ->
Graph ((n,c,m)::l) nil s h
| elim_E: forall ln n1 n2 n3 le s h h1 h2 c m,
(n2 = 0 \/ In (n2,c,m) ln) ->
(n3 = 0 \/ In (n3,c,m) ln) ->
h1 # h2 ->
h = h1 +++ h2 ->
((nat_e n1)|-->((nat_e n2)::(nat_e n3)::(nat_e c)::(nat_e m)::nil)) s h1 ->
(((nat_e n1)|-->((nat_e 0)::(nat_e 0)::(nat_e c)::(nat_e m)::nil)) -* (Graph ln (le))) s h2 ->
Graph ln ((n1, n2, n3)::le) s h.
Lemma Graph_inde_store: forall N E s s' h,
Graph N E s h -> Graph N E s' h.
Lemma Graph_less_edge: forall N E s h,
Graph N E s h -> exists h', Graph N nil s h'.
Lemma Graph_node_unique_store: forall N E s h n m c,
Graph N E s h ->
Definition SW_precond (N: list nat) (E: list (nat * nat * nat)) (root: nat): assert :=
fun s h => Graph N E s h /\ In root N.
Definition SW_postcond (N: list nat) (E: list (nat * nat * nat)) : assert :=
Graph N E.
Definition t := 0.
Definition p := 1.
Definition q := 2.
Hint Unfold t.
Hint Unfold p.
Hint Unfold q.
Definition SW_specif : Prop :=
forall root N E,
{{SW_precond N E root}}
SW t p q root
{{SW_postcond N E}}.
Lemma SW_verif : SW_specif.