Library vc
Require Import List.
Require Import ZArith.
Require Import List.
Require Import EqNat.
Require Import Classical.
Require Import util.
Require Import heap.
Require Import bipl.
Require Import axiomatic.
Require Import contrib.
Open Local Scope heap_scope.
Open Local Scope sep_scope.
weakest precondition generator
Inductive cmd' : Type :=
skip': cmd'
| assign_var_e' : var.v -> expr -> cmd'
| assign_var_deref' : var.v -> expr -> cmd'
| assign_deref_expr' : expr -> expr -> cmd'
| malloc' : var.v -> expr -> cmd'
| free' : expr -> cmd'
| while' : expr_b -> assert -> cmd' -> cmd'
| seq' : cmd' -> cmd' -> cmd'
| ifte' : expr_b -> cmd' -> cmd' -> cmd'.
Fixpoint proj_cmd (c':cmd') : cmd :=
match c' with
skip' => skip
| assign_var_e' x e => x <- e
| assign_var_deref' x e => x <-* e
| assign_deref_expr' e f => e *<- f
| malloc' x e => x <-malloc e
| free' e => free e
| while' b Q c'' => while b (proj_cmd c'')
| seq' c1 c2 => (seq (proj_cmd c1) (proj_cmd c2))
| ifte' b c1 c2 => (ifte b thendo (proj_cmd c1) elsedo (proj_cmd c2))
end.
Notation "x <- e" := (assign_var_e' x e) (at level 80) : vc_scope.
Notation "x '<-*' e" := (assign_var_deref' x e) (at level 80) : vc_scope.
Notation "e '*<-' f" := (assign_deref_expr' e f) (at level 80) : vc_scope.
Notation "x '<-malloc' e" := (malloc' x e) (at level 80) : vc_scope.
Notation "c ; d" := (seq' c d) (at level 81, right associativity) : vc_scope.
Notation "'ifte' a 'thendo' x 'elsedo' y" := (ifte' a x y)(at level 80) : vc_scope.
Open Local Scope vc_scope.
Fixpoint wp (c:cmd') (P:assert) {struct c} : assert :=
match c with
skip' => P
| assign_var_e' x e => update_store2 x e P
| assign_var_deref' x e =>
fun s h => exists e0, ((e |-> e0) ** ((e |-> e0) -* (update_store2 x e0 P))) s h
| assign_deref_expr' e f =>
fun s h => exists x, (((e |-> x) ** ((e |-> f) -* P)) s h)
| (malloc' x e) =>
fun s h => forall n, (((nat_e n) |-> e) -* (update_store2 x (nat_e n) P)) s h
| free' e => fun s h =>
exists l, val2loc (eval e s) = l /\
exists v, heap.lookup l h = Some v /\
P s (h --- l)
| while' b Q c' => Q
| seq' c1 c2 => wp c1 (wp c2 P)
| ifte' b c1 c2 => (fun s h =>
( eval_b b s = true -> (wp c1 P) s h) /\
( eval_b b s = false -> (wp c2 P) s h))
end.
Fixpoint vc (c:cmd') (P:assert) {struct c} : assert :=
match c with
skip' => TT
| assign_var_e' x e => TT
| assign_var_deref' x e => TT
| assign_deref_expr' e f => TT
| malloc' x e => TT
| free' e => TT
| while' b Q c' => fun s h =>
(Q s h /\ eval_b b s = false -> P s h) /\
(Q s h /\ eval_b b s = true -> (wp c' Q) s h) /\
((vc c' Q) s h)
| seq' c1 c2 => fun s h =>
vc c1 (wp c2 P) s h /\ (vc c2 P) s h
| ifte' b c1 c2 => fun s h =>
(vc c1 P) s h /\ (vc c2 P) s h
end.
Lemma vc_soundness : forall c' P,
(forall s h, vc c' P s h) -> {{wp c' P}} (proj_cmd c') {{ P }}.
Lemma wp_vc_util: forall c Q P ,
(forall s h, vc c Q s h) -> (P ==> wp c Q) -> {{P}} proj_cmd c {{Q}}.
Close Local Scope vc_scope.