Library axiomatic
Require Import List.
Require Import ZArith.
Require Import List.
Require Import EqNat.
Require Import Classical.
Require Import util.
Require Import heap.
Require Import heap_tactic.
Require Import bipl.
Inductive cmd : Set :=
skip : cmd
| assign : var.v -> expr -> cmd
| lookup : var.v -> expr -> cmd
| mutation : expr -> expr -> cmd
| malloc : var.v -> expr -> cmd
| free : expr -> cmd
| while : expr_b -> cmd -> cmd
| seq : cmd -> cmd -> cmd
| ifte : expr_b -> cmd -> cmd -> cmd.
Notation "x <- e" := (assign x e) (at level 80) : sep_scope.
Notation "x '<-*' e" := (lookup x e) (at level 80) : sep_scope.
Notation "e '*<-' f" := (mutation e f) (at level 80) : sep_scope.
Notation "x '<-malloc' e" := (malloc x e) (at level 80) : sep_scope.
Notation "c ; d" := (seq c d) (at level 81, right associativity) : sep_scope.
Notation "'ifte' a 'thendo' x 'elsedo' y" := (ifte a x y)(at level 80) : sep_scope.
Open Local Scope sep_scope.
Open Local Scope heap_scope.
Close Local Scope nat_scope.
Definition state := prod store.s heap.h.
Open Local Scope nat_scope.
operational semantics
Inductive exec : option state -> cmd -> option state -> Prop :=
exec_none : forall c, exec None c None
| exec_skip: forall s h, exec (Some (s, h)) skip (Some (s, h))
| exec_assign : forall s h x e,
exec (Some (s, h)) (x <- e) (Some (store.update x (eval e s) s, h))
| exec_lookup : forall s h x e p v,
val2loc (eval e s) = p -> heap.lookup p h = Some v -> exec (Some (s, h)) (x <-* e) (Some (store.update x v s, h))
| exec_lookup_err : forall s h x e p,
val2loc (eval e s) = p ->
heap.lookup p h = None -> exec (Some (s, h)) (x <-* e) None
| exec_mutation : forall s h e e' p v,
val2loc (eval e s) = p -> heap.lookup p h = Some v -> exec (Some (s, h)) (e *<- e') (Some (s, heap.update p (eval e' s) h))
| exec_mutation_err : forall s h e e' p,
val2loc (eval e s) = p -> heap.lookup p h = None -> exec (Some (s, h)) (e *<- e') None
| exec_malloc : forall s h x e n,
h # heap.singleton n (eval e s) ->
exec (Some (s, h)) (x <-malloc e)
(Some (store.update x (Z_of_nat n) s, h +++ (heap.singleton n (eval e s))))
| exec_free : forall s h e v p,
val2loc (eval e s) = p ->
heap.lookup p h = Some v->
exec (Some (s, h)) (free e) (Some (s, h --- p))
| exec_free_err : forall s h e p,
val2loc (eval e s) = p -> heap.lookup p h = None -> exec (Some (s, h)) (free e) None
| exec_seq : forall s s' s'' c d,
exec s c s' -> exec s' d s'' -> exec s (c ; d) s''
| exec_while_true : forall s h s' s'' b c,
eval_b b s = true ->
exec (Some (s, h)) c s' ->
exec s' (while b c) s'' ->
exec (Some (s, h)) (while b c) s''
| exec_while_false : forall s h b c,
eval_b b s = false ->
exec (Some (s, h)) (while b c) (Some (s, h))
| exec_ifte_true : forall b c d s h s',
eval_b b s = true ->
exec (Some (s, h)) c s' ->
exec (Some (s, h)) (ifte b thendo c elsedo d) s'
| exec_ifte_false : forall b c d s h s',
eval_b b s = false ->
exec (Some (s, h)) d s' ->
exec (Some (s, h)) (ifte b thendo c elsedo d) s'.
Notation "s -- c --> t" := (exec s c t) (at level 82) : sep_scope.
Lemma from_none' : forall s0 c s, (s0 -- c --> s) -> s0 = None -> s = None.
Lemma from_none : forall c s, (None -- c --> s) -> s = None.
Definition update_store2 (x:var.v) (e:expr) (P:assert) : assert :=
fun s h => P (store.update x (eval e s) s) h.
Definition lookup2 (x:var.v) (e:expr) (P:assert) : assert :=
fun s h =>
exists p, val2loc (eval e s) = p /\
exists z, heap.lookup p h = Some z /\
P (store.update x z s) h.
Definition update_heap2 (e:expr) (e':expr) (P:assert) : assert :=
fun s h =>
exists p, val2loc (eval e s) = p /\
exists z, heap.lookup p h = Some z /\
P s (heap.update p (eval e' s) h).
Inductive semax : assert -> cmd -> assert -> Prop :=
semax_skip: forall P, semax P skip P
| semax_assign : forall P x e,
semax (update_store2 x e P) (x <- e) P
| semax_lookup : forall P x e,
semax (lookup2 x e P) (x <-* e) P
| semax_mutation : forall P e e',
semax (update_heap2 e e' P) (e *<- e') P
| semax_malloc : forall P x e,
semax (fun s h => forall n, ((nat_e n)|->e -* update_store2 x (nat_e n) P) s h)
(x <-malloc e) P
| semax_free : forall P e,
semax
(fun s h =>
exists l, val2loc (eval e s) = l /\
exists v, heap.lookup l h = Some v /\
P s (h --- l))
(free e)
P
| semax_seq : forall P Q R c d,
semax P c Q -> semax Q d R -> semax P (c ; d) R
| semax_while : forall P b c,
semax (fun s h => (P s h /\ eval_b b s = true)) c P ->
semax P (while b c) (fun s h => P s h /\ eval_b b s = false)
| semax_conseq : forall (P P' Q Q':assert) c,
(Q' ==> Q) -> (P ==> P') ->
semax P' c Q' -> semax P c Q
| semax_ifte : forall P Q b c d,
semax (fun s h => P s h /\ eval_b b s = true) c Q ->
semax (fun s h => P s h /\ eval_b b s = false) d Q ->
semax P (ifte b thendo c elsedo d) Q.
Notation "{{ P }} c {{ Q }}" := (semax P c Q) (at level 82) : sep_scope.
axiomatic semantic lemmas
Lemma semax_weaken_post : forall P (Q Q':assert) c,
(Q' ==> Q) -> {{ P }} c {{ Q' }} -> {{ P }} c {{ Q }}.
Lemma semax_strengthen_pre : forall (P P':assert) Q c,
(P ==> P') -> {{ P' }} c {{ Q }} -> {{ P }} c {{ Q }}.
Definition semax' (P:assert) (c:cmd) (Q:assert) : Prop :=
forall s h, (P s h -> ~(Some (s, h) -- c --> None)) /\
(forall s' h', P s h -> (Some (s, h) -- c --> Some (s', h')) -> Q s' h').
Lemma semax_sound : forall P Q c,
{{ P }} c {{ Q }} -> semax' P c Q.
Definition wp_semantics (c: cmd) (Q: assert): assert :=
fun s h =>
~ ((Some (s, h)) -- c --> None) /\
forall s' h',
(Some (s, h)) -- c --> (Some (s', h')) -> Q s' h'.
Lemma exec_lookup_not_None: forall s h v e,
~ (Some (s, h) -- v <-* e --> None) ->
exists p : loc,
val2loc (eval e s) = p /\
(exists z : heap.v, heap.lookup p h = Some z).
Lemma exec_mutation_not_None: forall s h e e0,
~ (Some (s, h) -- e *<- e0 --> None) ->
exists p : loc,
val2loc (eval e s) = p /\
(exists z : heap.v, heap.lookup p h = Some z).
Lemma exec_free_not_None: forall s h e,
~ (Some (s, h) -- free e --> None) ->
exists p : loc,
val2loc (eval e s) = p /\
(exists z : heap.v, heap.lookup p h = Some z).
Lemma exec_seq1_not_None: forall s h c1 c2,
~ (Some (s, h) -- (c1; c2) --> None) ->
~ (Some (s, h) -- c1 --> None).
Lemma exec_seq2_not_None: forall s h c1 c2 s' h',
~ (Some (s, h) -- (c1; c2) --> None) ->
Some (s, h) -- c1 --> Some (s',h') ->
~ (Some (s', h') -- c2 --> None).
Lemma exec_ifte1_not_None: forall s h c1 c2 e,
~ (Some (s, h) -- (ifte e thendo c1 elsedo c2) --> None) ->
eval_b e s = true ->
(~ (Some (s, h) -- c1 --> None)).
Lemma exec_ifte2_not_None: forall s h c1 c2 e,
~ (Some (s, h) -- (ifte e thendo c1 elsedo c2) --> None) ->
eval_b e s = false ->
~ (Some (s, h) -- c2 --> None).
Lemma exec_while1_not_None: forall s h e c,
~ (Some (s, h) -- while e c --> None) ->
eval_b e s = true ->
~ (Some (s, h) -- c --> None).
Lemma exec_while2_not_None: forall s h e c s' h',
~ (Some (s, h) -- while e c --> None) ->
eval_b e s = true ->
Some (s, h) -- c --> Some (s', h') ->
~ (Some (s', h') -- while e c --> None).
Lemma wp_semantics_sound: forall c Q,
{{wp_semantics c Q}} c {{Q}}.
Lemma semax_complete : forall P Q c,
semax' P c Q -> {{ P }} c {{ Q }}.
Definition semax_alternative (P:assert) (c:cmd) (Q:assert) : Prop :=
forall s h, P s h ->
(forall s' h', ((Some (s, h)) -- c --> (Some (s', h'))) -> (Q s' h')).
Lemma semax_sound_alternative : forall P Q c,
{{ P }} c {{ Q }} -> semax_alternative P c Q.
Lemma semax_lookup_backwards : forall x e P,
{{ fun s h => exists e0, (e|->e0 ** (e|->e0 -* update_store2 x e0 P)) s h }} x <-* e {{ P }}.
Lemma semax_lookup_backwards_alternative : forall x e P e0,
{{ e|->e0 ** (e|->e0 -* update_store2 x e0 P) }} x <-* e {{ P }}.
Lemma semax_mutation_local: forall x v v',
{{ x |-> v }} x *<- v' {{ x |-> v' }}.
Fixpoint modified_cmd_var (c:cmd) {struct c} : list var.v :=
match c with
skip => nil
| x <- e => x::nil
| x <-* e => x::nil
| e *<- f => nil
| x <-malloc e => x::nil
| free e => nil
| c1 ; c2 => modified_cmd_var c1 ++ modified_cmd_var c2
| ifte a thendo c1 elsedo c2 => modified_cmd_var c1 ++ modified_cmd_var c2
| while a c1 => modified_cmd_var c1
end.
Fixpoint modified_loc_expr (c:cmd) {struct c} : list expr :=
match c with
skip => nil
| x <- e => nil
| x <-* e => nil
| e *<- f => e::nil
| x <-malloc e => nil
| free e => e::nil
| c1 ; c2 => modified_loc_expr c1 ++ modified_loc_expr c2
| ifte a thendo c1 elsedo c2 => modified_loc_expr c1 ++ modified_loc_expr c2
| while a c1 => modified_loc_expr c1
end.
Lemma inde_seq : forall R c d,
inde (modified_cmd_var (c; d)) R ->
inde (modified_cmd_var c) R /\ inde (modified_cmd_var d) R.
Lemma inde_ifte : forall R b c d,
inde (modified_cmd_var (ifte b thendo c elsedo d)) R ->
inde (modified_cmd_var c) R /\ inde (modified_cmd_var d) R.
Lemma frame_rule : forall P c Q,
{{P}} c {{Q}} ->
forall R ,
(inde (modified_cmd_var c) R ->
{{ P ** R }} c {{ Q ** R }}).
Lemma semax_free_global_backwards : forall e v P,
{{ (e |-> v) ** P }} free e {{ sep.emp ** P }}.
Lemma semax_mutation_global : forall P e e',
{{(fun s' h' => exists e'', ((e |-> e'') s' h'))**P}} e *<- e' {{(e |-> e')**P}}.
Lemma semax_mutation_global_alternative : forall P e e' e'',
{{ (e |-> e'') ** P }} e *<- e' {{ (e |-> e') ** P }}.
Lemma semax_mutation_backwards : forall P e e',
{{fun s h => exists e'', (e|->e''**(e|->e'-*P)) s h}} e *<- e' {{P}}.
Lemma semax_mutation_backwards_alternative : forall P e e' e'',
{{ (e |-> e'') ** ((e |-> e') -* P) }} e *<- e' {{ P }}.