Library frag
Load seplog_header.
Require Import Omega.
Require Import Bool.
Definition Pi := expr_b.
Inductive Sigma : Set :=
singl : expr -> expr -> Sigma
| cell : expr -> Sigma
| epsi : Sigma
| star : Sigma -> Sigma -> Sigma.
Fixpoint Sigma_interp (a : Sigma) : assert :=
match a with
singl e1 e2 => e1 |-> e2
| epsi => sep.emp
| star s1 s2 => (Sigma_interp s1) ** (Sigma_interp s2)
| cell e => (fun s h => exists v, (e |-> int_e v) s h)
end.
Definition assrt := prod Pi Sigma.
Definition assrt_interp (a: assrt) : assert :=
match a with
(pi, sigm) => (fun s h => eval_b pi s = true /\ Sigma_interp sigm s h)
end.
Inductive ps1 : assrt -> assrt -> Prop :=
ps1_incons: forall pi1 pi2 sig1 sig2,
(forall s , eval_b pi1 s = true -> False) ->
ps1 (pi1,sig1) (pi2,sig2)
| ps1_tauto: forall pi1 pi2,
(forall s , eval_b pi1 s = true -> eval_b pi2 s = true) ->
ps1 (pi1,epsi) (pi2,epsi)
| ps1_coml: forall pi1 sig1 sig2 L,
ps1 (pi1,star sig2 sig1) L ->
ps1 (pi1,star sig1 sig2) L
| ps1_comr: forall pi1 sig1 sig2 L,
ps1 L (pi1,star sig2 sig1) ->
ps1 L (pi1,star sig1 sig2)
| ps1_assocl: forall pi1 sig1 sig2 sig3 L,
ps1 (pi1,star (star sig1 sig2) sig3) L ->
ps1 (pi1,star sig1 (star sig2 sig3)) L
| ps1_assocr: forall pi1 sig1 sig2 sig3 L,
ps1 L (pi1,star (star sig1 sig2) sig3) ->
ps1 L (pi1,star sig1 (star sig2 sig3))
| ps1_epseliml: forall pi1 sig1 L,
ps1 (pi1,sig1) L ->
ps1 (pi1,star sig1 epsi) L
| ps1_epselimr: forall pi1 sig1 L,
ps1 L (pi1,sig1) ->
ps1 L (pi1,star sig1 epsi)
| ps1_epsintrol: forall pi1 sig1 L,
ps1 (pi1,star sig1 epsi) L ->
ps1 (pi1,sig1) L
| ps1_epsintror: forall pi1 sig1 L,
ps1 L (pi1,star sig1 epsi) ->
ps1 L (pi1,sig1)
| ps1_star_elim: forall pi1 pi2 sig1 sig2 e1 e2 e3 e4,
(forall s, eval_b pi1 s = true -> eval_b (e1 == e3) s = true) ->
(forall s , eval_b pi1 s = true -> eval_b (e2 == e4) s = true) ->
ps1 (pi1,sig1) (pi2,sig2) ->
ps1 (pi1,star sig1 (singl e1 e2)) (pi2, star sig2 (singl e3 e4))
| ps1_star_elim': forall pi1 pi2 sig1 sig2 e1 e2 e3,
(forall s, eval_b pi1 s = true -> eval_b (e1 == e3) s = true) ->
ps1 (pi1,sig1) (pi2,sig2) ->
ps1 (pi1,star sig1 (singl e1 e2)) (pi2, star sig2 (cell e3))
| ps1_star_elim'': forall pi1 pi2 sig1 sig2 e1 e3,
(forall s, eval_b pi1 s = true -> eval_b (e1 == e3) s = true) ->
ps1 (pi1,sig1) (pi2,sig2) ->
ps1 (pi1,star sig1 (cell e1)) (pi2, star sig2 (cell e3)).
Lemma ps1_soundness: forall P Q, ps1 P Q ->
(assrt_interp P) ==> (assrt_interp Q).
Ltac ps1_turnl :=
match goal with
| |- ps1 (?Pi, cell ?e) ?L => eapply ps1_epsintrol; eapply ps1_coml; repeat eapply ps1_assocl
| |- ps1 (?Pi, singl ?e1 ?e2) ?L => eapply ps1_epsintrol; eapply ps1_coml; repeat eapply ps1_assocl
| |- ps1 (?Pi, star ?e1 ?e2) ?L => eapply ps1_coml; repeat eapply ps1_assocl
end.
Ltac ps1_turnr :=
match goal with
| |- ps1 ?L (?Pi, cell ?e) => eapply ps1_epsintror; eapply ps1_comr; repeat eapply ps1_assocr
| |- ps1 ?L (?Pi, singl ?e1 ?e2) => eapply ps1_epsintror; eapply ps1_comr; repeat eapply ps1_assocr
| |- ps1 ?L (?Pi, star ?e1 ?e2) => eapply ps1_comr; repeat eapply ps1_assocr
end.
Ltac ps1_resolve := repeat eapply ps1_assocr; repeat eapply ps1_assocl;
match goal with
| |- ps1 (?pi1, star ?sig1 epsi) ?L => ps1_turnl; idtac
| |- ps1 ?L (?Pi, cell ?e) => ps1_turnr; idtac
| |- ps1 ?L (?Pi, singl ?e1 ?e2) => ps1_turnr; idtac
| |- ps1 (?Pi, cell ?e) ?L => ps1_turnl; idtac
| |- ps1 (?Pi, singl ?e1 ?e2) ?L => ps1_turnl; idtac
| |- ps1 (?pi1, epsi) (?pi2, epsi) => eapply ps1_tauto; [intros; Omega_exprb]
| |- ps1 (?pi1, epsi) (?pi2, epsi) => eapply ps1_incons; [intros; Omega_exprb]
(*****)
| |- ps1 (?pi1, star ?e epsi) ?L => eapply ps1_epseliml; idtac
| |- ps1 ?L (?pi2, star ?e epsi) => eapply ps1_epselimr; idtac
(*****)
| |- ps1 (?pi1, star ?sig1 (singl ?e1 ?e2)) (?pi2, star ?sig2 (singl ?e3 ?e4)) =>
(eapply ps1_star_elim; [ Omega_exprb | Omega_exprb | idtac] || ps1_turnl; idtac)
| |- ps1 (?pi1, star ?sig1 (singl ?e1 ?e2)) (?pi2, star ?sig2 (cell ?e3)) =>
(eapply ps1_star_elim'; [ Omega_exprb | idtac] || ps1_turnl; idtac)
| |- ps1 (?pi1, star ?sig1 (cell ?e1)) (?pi2, star ?sig2 (cell ?e3)) =>
(eapply ps1_star_elim''; [ Omega_exprb | idtac] || ps1_turnl; idtac)
| |- ps1 (?pi2, star ?sig2 (cell ?e3)) (?pi1, star ?sig1 (singl ?e1 ?e2)) => ps1_turnl
end.
Ltac ps1_Resolve := repeat ps1_resolve.
Require Import expr_b_dp.
Fixpoint Sigma_clean_epsi (t: Sigma) : Sigma :=
match t with
| star t1 t2 =>
match (Sigma_clean_epsi t1) with
| epsi => (Sigma_clean_epsi t2)
| t1' => match (Sigma_clean_epsi t2) with
| epsi => t1'
| t2' => star t1' t2'
end
end
| _ => t
end.
Definition Sigma_elt_eq (e1 e2: Sigma) (env: Pi) : bool :=
match (e1,e2) with
(epsi,epsi) => true
| ((singl x1 x2),(singl x3 x4)) => (andb (expr_b_dp (env =b> (x1 == x3))) (expr_b_dp (env =b> (x2 == x4))))
| ((singl x1 x2),(cell x3)) => (expr_b_dp (env =b> (x1 == x3)))
| ((cell x1),(cell x3)) => (expr_b_dp (env =b> (x1 == x3)))
| (_,_) => false
end.
Fixpoint Sigma_elt_term_extract (e t: Sigma) (env: Pi) {struct t}: option Sigma :=
match t with
| star t1 t2 =>
match Sigma_elt_term_extract e t1 env with
| None => match (Sigma_elt_term_extract e t2 env) with
| None => None
| Some t2' => Some (Sigma_clean_epsi (star t1 t2'))
end
| Some t1' => Some (Sigma_clean_epsi (star t1' t2))
end
| _ => if (Sigma_elt_eq e t env) then Some epsi else None
end.
Fixpoint Sigma_elt_term_extract' (e t: Sigma) (env: Pi) {struct t}: option Sigma :=
match t with
| star t1 t2 =>
match Sigma_elt_term_extract' e t1 env with
| None => match (Sigma_elt_term_extract' e t2 env) with
| None => None
| Some t2' => Some (Sigma_clean_epsi (star t1 t2'))
end
| Some t1' => Some (Sigma_clean_epsi (star t1' t2))
end
| _ => if Sigma_elt_eq t e env then Some epsi else None
end.
Fixpoint Sigma_term_term_eq (t1 t2: Sigma) (env: Pi) {struct t1}: option Sigma :=
match t1 with
| star t11 t12 =>
match (Sigma_term_term_eq t11 t2 env) with
| None => None
| Some t2' => match (Sigma_term_term_eq t12 t2' env) with
| None => None
| Some t2'' => Some t2''
end
end
| _ => match (Sigma_elt_term_extract t1 t2 env) with
| None => None
| Some t2' => Some t2'
end
end.
Fixpoint Sigma_get_cell_val (e: expr) (sig: Sigma) (env: Pi) {struct sig}: option expr :=
match sig with
| epsi => None
| singl e1 e2 => if (expr_b_dp (env =b> (e1 == e))) then (Some e2) else None
| (cell e1) => None
| star s1 s2 =>
match (Sigma_get_cell_val e s1 env) with
| None => (Sigma_get_cell_val e s2 env)
| Some e2 => Some e2
end
end.
Definition frag_decision (P Q: Pi * Sigma) : bool :=
let (pi1,sig1) := P in
( let (pi2, sig2) := Q in
(
if (expr_b_dp (pi1 =b> pi2)) then
match (Sigma_term_term_eq sig1 sig2 pi1) with
| Some epsi => true
| _ => false
end
else false
)
).
Lemma Sigma_clean_epsi_correct: forall t,
(Sigma_interp (Sigma_clean_epsi t) ==> Sigma_interp t).
Lemma Sigma_clean_epsi_correct': forall t,
(Sigma_interp t ==> Sigma_interp (Sigma_clean_epsi t)).
Lemma Sigma_elt_eq_correct: forall e1 e2 env,
Sigma_elt_eq e1 e2 env = true ->
(forall s h,
(s |b= env) ->
(Sigma_interp e1 s h -> Sigma_interp e2 s h)
).
Opaque Sigma_clean_epsi.
Lemma Sigma_elt_term_extract_correct:
forall e2 e1 env e2',
Sigma_elt_term_extract e1 e2 env = Some e2' ->
(forall s h,
(s |b= env) ->
(Sigma_interp (star e1 e2') s h -> Sigma_interp e2 s h)
).
Lemma Sigma_elt_term_extract'_correct:
forall e2 e1 env e2',
Sigma_elt_term_extract' e1 e2 env = Some e2' ->
(forall s h,
(s |b= env) ->
Sigma_interp e2 s h ->
Sigma_interp (star e1 e2') s h
).
Lemma Sigma_term_term_eq_correct:
forall t1 t2 env t1',
(Sigma_term_term_eq t1 t2 env = Some t1') ->
(forall s h,
(s |b= env) ->
(Sigma_interp (star t1 t1') s h -> Sigma_interp t2 s h)
).
Lemma Sigma_get_cell_val_correct: forall sig e env e',
Sigma_get_cell_val e sig env = Some e' ->
(forall s h,
(s |b= env) ->
Sigma_interp sig s h -> (Sigma_interp (singl e e') ** TT) s h).
Eval_b_hyp H.
Eval_b_hyp H0.
Lemma frag_decision_correct: forall P Q,
frag_decision P Q = true ->
(assrt_interp P ==> assrt_interp Q).
Transparent Sigma_clean_epsi.
Inductive L_assrt : Set :=
L_elt : assrt -> L_assrt
| L_subst : list (var.v * expr) -> L_assrt -> L_assrt
| L_lookup : var.v -> expr -> L_assrt -> L_assrt
| L_mutation : expr -> expr -> L_assrt -> L_assrt
| L_if : expr_b -> L_assrt -> L_assrt -> L_assrt.
Fixpoint subst_lst2update_store (l:list (var.v * expr)) (P:assert) {struct l} : assert :=
match l with
nil => P
| (x,e)::tl => subst_lst2update_store tl (update_store2 x e P)
end.
Lemma subst_lst2update_store_app : forall l2 l1 P s h,
subst_lst2update_store (l2 ++ l1) P s h ->
subst_lst2update_store l1 (subst_lst2update_store l2 P) s h.
Lemma subst_lst2update_store' : forall l x v s h P,
subst_lst2update_store (l ++ (x,v)::nil) P s h ->
update_store2 x v (subst_lst2update_store l P) s h.
Fixpoint L_assrt_interp (a: L_assrt) : assert :=
match a with
L_elt a1 => assrt_interp a1
| L_subst l L => subst_lst2update_store l (L_assrt_interp L)
| L_lookup x e L => (fun s h => exists e0, (e |-> e0 ** (e |-> e0 -* (update_store2 x e0 (L_assrt_interp L)))) s h)
| L_mutation e1 e2 L => (fun s h => exists e0, (e1 |-> e0 ** (e1 |-> e2 -* (L_assrt_interp L))) s h)
| L_if b L1 L2 => (fun s h => (eval_b b s = true -> L_assrt_interp L1 s h) /\
(eval_b b s = false -> L_assrt_interp L2 s h))
end.
Fixpoint subst_e (e patt repl: expr) {struct e} : expr :=
match e with
var_e w => match expr_eq e patt with
true => repl
| false => e
end
| int_e i => match expr_eq e patt with
true => repl
| false => e
end
| add_e e1 e2 => match expr_eq e patt with
true => repl
| false => add_e (subst_e e1 patt repl) (subst_e e2 patt repl)
end
| min_e e1 e2 => match expr_eq e patt with
true => repl
| false => min_e (subst_e e1 patt repl) (subst_e e2 patt repl)
end
| mul_e e1 e2 => match expr_eq e patt with
true => repl
| false => mul_e (subst_e e1 patt repl) (subst_e e2 patt repl)
end
| div_e e1 e2 => match expr_eq e patt with
true => repl
| false => div_e (subst_e e1 patt repl) (subst_e e2 patt repl)
end
| mod_e e1 e2 => match expr_eq e patt with
true => repl
| false => mod_e (subst_e e1 patt repl) (subst_e e2 patt repl)
end
end.
Fixpoint subst_b (e: expr_b) (patt repl: expr) {struct e} : expr_b :=
match e with
true_b => true_b
| f == g => subst_e f patt repl == subst_e g patt repl
| f =/= g => subst_e f patt repl =/= subst_e g patt repl
| f >>= g => subst_e f patt repl >>= subst_e g patt repl
| f >> g => subst_e f patt repl >> subst_e g patt repl
| f &&& g => (subst_b f patt repl) &&& (subst_b g patt repl)
| f ||| g => (subst_b f patt repl) ||| (subst_b g patt repl)
| neg_b e => neg_b (subst_b e patt repl)
end.
Fixpoint subst_Sigma (a: Sigma) (x: var.v) (e: expr) {struct a} : Sigma :=
match a with
singl e1 e2 => singl (subst_e e1 (var_e x) e) (subst_e e2 (var_e x) e)
| epsi => epsi
| star s1 s2 => star (subst_Sigma s1 x e) (subst_Sigma s2 x e)
| cell e1 => cell (subst_e e1 (var_e x) e)
end.
Definition subst_assrt (a: assrt) (x: var.v) (e: expr): assrt :=
match a with
(pi, sigm) => (subst_b pi (var_e x) e, subst_Sigma sigm x e)
end.
Fixpoint subst_assrt_lst (l:list (var.v * expr)) (a:assrt) {struct l} : assrt :=
match l with
nil => a
| (x,e)::tl => subst_assrt_lst tl (subst_assrt a x e)
end.
Fixpoint subst_e_lst (l: list (var.v * expr)) (e: expr) {struct l}: expr :=
match l with
nil => e
| (x,e')::tl => subst_e_lst tl (subst_e e (var_e x) e')
end.
Fixpoint subst_b_lst (l: list (var.v * expr)) (e: expr_b) {struct l}: expr_b :=
match l with
nil => e
| (x,e')::tl => subst_b_lst tl (subst_b e (var_e x) e')
end.
Lemma subst_e2store_update: forall e s x v,
eval (subst_e e (var_e x) v) s = eval e (store.update x (eval v s) s).
Lemma subst_b2store_update: forall b s x v,
eval_b (subst_b b (var_e x) v) s = eval_b b (store.update x (eval v s) s).
Lemma subst_e_lst_int_e: forall l v s,
eval (subst_e_lst l (int_e v)) s = v.
Lemma subst_Sigma2store_update: forall sigm s h x v,
Sigma_interp (subst_Sigma sigm x v) s h ->
Sigma_interp sigm (store.update x (eval v s) s) h.
Lemma subst_Sigma2store_update': forall sigm s h x v,
Sigma_interp sigm (store.update x (eval v s) s) h ->
Sigma_interp (subst_Sigma sigm x v) s h.
Lemma subst_lst2update_store_assrt_interp: forall l s h pi sigm,
assrt_interp (subst_assrt_lst l (pi, sigm)) s h ->
subst_lst2update_store l (assrt_interp (pi, sigm)) s h.
Lemma subst_lst2update_store_subst_b_lst: forall (b':bool) l b s h,
subst_lst2update_store l (fun s h => eval_b b s = b') s h ->
eval_b (subst_b_lst l b) s = b'.
Lemma entails_subst_lst2update_store: forall l P1 P2 s h,
(P1 ==> P2) ->
subst_lst2update_store l P1 s h ->
subst_lst2update_store l P2 s h.
Lemma subst_lst2update_store_and: forall l P1 P2 s h,
subst_lst2update_store l (fun s h => P1 s h /\ P2 s h) s h ->
subst_lst2update_store l P1 s h /\ subst_lst2update_store l P2 s h.
Lemma subst_lst2update_store_and': forall l P1 P2 s h,
subst_lst2update_store l P1 s h /\ subst_lst2update_store l P2 s h ->
subst_lst2update_store l (fun s h => P1 s h /\ P2 s h) s h.
Lemma subst_lst2update_store_imply: forall l P1 P2 s h,
(subst_lst2update_store l P1 s h -> subst_lst2update_store l P2 s h) ->
subst_lst2update_store l (fun s h => P1 s h -> P2 s h) s h.
Lemma subst_lst2update_store_exists: forall l (P: expr -> assert) s h,
(exists x0, (subst_lst2update_store l (P x0)) s h) ->
subst_lst2update_store l (fun s h => exists e0, P e0 s h) s h.
Lemma subst_lst2update_store_sep_con: forall l P1 P2 s h,
((fun s h => subst_lst2update_store l P1 s h) ** (fun s h => subst_lst2update_store l P2 s h)) s h ->
subst_lst2update_store l (P1 ** P2) s h.
Lemma subst_lst2update_store_mapsto: forall l e1 e2 s h,
(subst_e_lst l e1 |-> subst_e_lst l e2) s h ->
subst_lst2update_store l (e1 |-> e2) s h.
Lemma subst_lst2update_store_mapsto': forall l e1 e2 s h,
subst_lst2update_store l (e1 |-> e2) s h ->
(subst_e_lst l e1 |-> subst_e_lst l e2) s h.
Lemma subst_lst2update_store_sepimp: forall l P1 P2 s h,
((fun s h => subst_lst2update_store l P1 s h) -* (fun s h => subst_lst2update_store l P2 s h)) s h ->
subst_lst2update_store l (P1 -* P2) s h.
Lemma subst_lst2update_store_lookup' : forall e x v s,
exists e', eval e s = eval (subst_e e' (var_e x) v) s.
Lemma subst_lst2update_store_lookup: forall l e1 e2 s h P,
(exists e0,
((subst_e_lst l e1 |-> e0) **
(subst_e_lst l e1 |-> subst_e_lst l e2 -* subst_lst2update_store l P)) s h) ->
subst_lst2update_store l (fun s' h' => exists e0, (e1 |-> e0 ** (e1 |-> e2 -* P)) s' h') s h.
Module Type FRESH.
Parameter fresh_e : var.v -> expr -> Prop.
Parameter fresh_b : var.v -> expr_b -> Prop.
Parameter fresh_Sigma : var.v -> Sigma -> Prop.
Parameter fresh_assrt : var.v -> assrt -> Prop.
Parameter fresh_lst : var.v -> (list (var.v * expr)) -> Prop.
Parameter fresh_L_assrt : var.v -> L_assrt -> Prop.
Parameter fresh_lst_decompose : forall x hd0 hd1 tl, fresh_lst x ((hd0,hd1)::tl) ->
fresh_e x hd1 /\ x <> hd0 /\ fresh_lst x tl.
Parameter fresh_e_var_e_neq : forall x y, fresh_e x (var_e y) -> x <> y.
Parameter fresh_e_eval: forall e x v s, fresh_e x e ->
eval e (store.update x v s) = eval e s.
Parameter fresh_L_assrt_inde: forall L x , fresh_L_assrt x L ->
inde (x::nil) (L_assrt_interp L).
End FRESH.
Require Import Max.
Module Fresh <: FRESH.
Fixpoint var_max_expr (e: expr) : var.v :=
match e with
var_e w => w
| int_e i => 0
| add_e e1 e2 => max (var_max_expr e1) (var_max_expr e2)
| min_e e1 e2 => max (var_max_expr e1) (var_max_expr e2)
| mul_e e1 e2 => max (var_max_expr e1) (var_max_expr e2)
| div_e e1 e2 => max (var_max_expr e1) (var_max_expr e2)
| mod_e e1 e2 => max (var_max_expr e1) (var_max_expr e2)
end.
Fixpoint var_max_expr_b (e: expr_b) : var.v :=
match e with
true_b => 0
| eq_b e1 e2 => max (var_max_expr e1) (var_max_expr e2)
| neq_b e1 e2 => max (var_max_expr e1) (var_max_expr e2)
| ge_b e1 e2 => max (var_max_expr e1) (var_max_expr e2)
| gt_b e1 e2 => max (var_max_expr e1) (var_max_expr e2)
| and_b e1 e2 => max (var_max_expr_b e1) (var_max_expr_b e2)
| or_b e1 e2 => max (var_max_expr_b e1) (var_max_expr_b e2)
| neg_b e => (var_max_expr_b e)
end.
Fixpoint var_max_Sigma (s: Sigma) : var.v :=
match s with
singl e1 e2 => max (var_max_expr e1) (var_max_expr e2)
| epsi => 0
| star s1 s2 => max (var_max_Sigma s1) (var_max_Sigma s2)
| cell e1 => var_max_expr e1
end.
Definition var_max_assrt (a: assrt) : var.v :=
match a with
(pi, sigm) => max (var_max_expr_b pi) (var_max_Sigma sigm)
end.
Fixpoint var_max_lst (l: list (var.v * expr)) : var.v :=
match l with
nil => 0
| (v,e)::tl => max (max v (var_max_expr e)) (var_max_lst tl)
end.
Fixpoint var_max_L_assrt (a: L_assrt) : var.v :=
match a with
L_elt a1 => var_max_assrt a1
| L_subst l L => max (var_max_lst l) (var_max_L_assrt L)
| L_lookup x e L=> max (max x (var_max_expr e)) (var_max_L_assrt L)
| L_mutation e1 e2 L => max (max (var_max_expr e1) (var_max_expr e2)) (var_max_L_assrt L)
| L_if b L1 L2 => max (max (var_max_L_assrt L1) (var_max_L_assrt L2)) (var_max_expr_b b)
end.
Definition fresh_e x e := (x > var_max_expr e).
Definition fresh_b x b := (x > var_max_expr_b b).
Definition fresh_Sigma x s := (x > var_max_Sigma s).
Definition fresh_assrt x a := (x > var_max_assrt a).
Definition fresh_lst x l := (x > var_max_lst l).
Definition fresh_L_assrt x L := (x > var_max_L_assrt L).
Lemma fresh_lst_decompose : forall x hd0 hd1 tl, fresh_lst x ((hd0,hd1)::tl) -> fresh_e x hd1 /\ x <> hd0 /\ fresh_lst x tl.
Lemma fresh_e_var_e_neq : forall x y, fresh_e x (var_e y) -> x <> y.
Ltac Max_inf_resolve :=
Max_inf_clean_hyp; Max_inf_resolve_goal
with
Max_inf_clean_hyp :=
match goal with
| id: fresh_e _ _ |- _ => red in id; simpl in id; Max_inf_clean_hyp
| id: fresh_b _ _ |- _ => red in id; simpl in id; Max_inf_clean_hyp
| id: fresh_Sigma _ _ |- _ => red in id; simpl in id; Max_inf_clean_hyp
| id: fresh_assrt _ _ |- _ => red in id; simpl in id; Max_inf_clean_hyp
| id: fresh_lst _ _ |- _ => red in id; simpl in id; Max_inf_clean_hyp
| id: fresh_L_assrt _ _ |- _ => red in id; simpl in id; Max_inf_clean_hyp
| id: ?x > max ?y ?z |- _ =>
generalize (max_lemma2 _ _ _ id); intro X; inversion_clear X; clear id; Max_inf_clean_hyp
| |- _ => idtac
end
with
Max_inf_resolve_goal :=
match goal with
| |- fresh_e _ _ => red; simpl; Max_inf_resolve_goal
| |- fresh_b _ _ => red; simpl; Max_inf_resolve_goal
| |- fresh_Sigma _ _ => red; simpl; Max_inf_resolve_goal
| |- fresh_assrt _ _ => red; simpl; Max_inf_resolve_goal
| |- fresh_lst _ _ => red; simpl; Max_inf_resolve_goal
| |- fresh_L_assrt _ _ => red; simpl; Max_inf_resolve_goal
| |- ?x > max ?y ?z =>
eapply max_lemma3; split; [Max_inf_resolve_goal | Max_inf_resolve_goal]
| |- _ => omega || tauto || idtac
end.
Lemma fresh_e_eval: forall e x v s,
fresh_e x e ->
eval e (store.update x v s) = eval e s.
Lemma fresh_b_inde: forall b x v,
fresh_b x b ->
inde (x::nil) (fun s h => eval_b b s = v).
Lemma var_max_Sigma_inde: forall sigm x ,
fresh_Sigma x sigm ->
inde (x::nil) (Sigma_interp sigm).
Lemma fresh_assrt_inde: forall a x ,
fresh_assrt x a ->
inde (x::nil) (assrt_interp a).
Lemma fresh_lst_inde: forall l P x,
inde (x::nil) P ->
fresh_lst x l ->
inde (x::nil) (subst_lst2update_store l P).
Lemma fresh_L_assrt_inde: forall L x ,
fresh_L_assrt x L ->
inde (x::nil) (L_assrt_interp L).
End Fresh.
Import Fresh.
Inductive LWP : assrt -> L_assrt -> Prop :=
| LWP_entail: forall pi1 pi2 sig1 sig2,
(assrt_interp (pi1,sig1)) ==> (assrt_interp (pi2,sig2)) ->
LWP (pi1,sig1) (L_elt (pi2,sig2))
| LWP_precond_stre: forall L1 L1' L2,
(assrt_interp L1) ==> (assrt_interp L1') ->
LWP L1' L2 ->
LWP L1 L2
| LWP_if: forall pi1 sig1 L1 L2 b,
LWP (pi1 &&& b,sig1) L1 ->
LWP (pi1 &&& (neg_b b),sig1) L2 ->
LWP (pi1,sig1) (L_if b L1 L2)
| LWP_mutation: forall pi1 sig1 e1 e2 e3 e4 L,
(forall s, eval_b pi1 s = true -> eval_b (e1 == e3) s = true) ->
LWP (pi1,star sig1 (singl e1 e4)) L ->
LWP (pi1,star sig1 (singl e1 e2)) (L_mutation e3 e4 L)
| LWP_mutation': forall pi1 sig1 e1 e3 e4 L,
(forall s, eval_b pi1 s = true -> eval_b (e1 == e3) s = true) ->
LWP (pi1,star sig1 (singl e1 e4)) L ->
LWP (pi1,star sig1 (cell e1)) (L_mutation e3 e4 L)
| LWP_lookup: forall pi1 sig1 e1 e2 e x L,
(forall s, eval_b pi1 s = true -> eval_b (e1 == e) s = true) ->
LWP (pi1,star sig1 (singl e1 e2)) (L_subst ((x,e2)::nil) L) ->
LWP (pi1,star sig1 (singl e1 e2)) (L_lookup x e L)
| LWP_subst_elt: forall pi1 pi2 sig1 sig2 l,
LWP (pi1,sig1) (L_elt (subst_assrt_lst l (pi2,sig2))) ->
LWP (pi1,sig1) (L_subst l (L_elt (pi2,sig2)))
| LWP_subst_subst: forall pi1 sig1 l1 l2 L,
LWP (pi1,sig1) (L_subst (l2 ++ l1) L) ->
LWP (pi1,sig1) (L_subst l1 (L_subst l2 L))
| LWP_subst_lookup: forall pi1 sig1 e1 e2 e x x' l L,
(forall s, eval_b pi1 s = true -> eval_b (e1 == (subst_e_lst l e)) s = true) ->
fresh_lst x' l ->
fresh_L_assrt x' L ->
fresh_e x' (var_e x) ->
LWP (pi1,star sig1 (singl e1 e2)) (L_subst ((x,(var_e x'))::l ++ ((x',e2)::nil)) L) ->
LWP (pi1,star sig1 (singl e1 e2)) (L_subst l (L_lookup x e L))
| LWP_subst_mutation: forall pi1 sig1 e1 e2 l L,
LWP (pi1,sig1) (L_mutation (subst_e_lst l e1) (subst_e_lst l e2) (L_subst l L)) ->
LWP (pi1,sig1) (L_subst l (L_mutation e1 e2 L))
| LWP_subst_if: forall pi1 sig1 l b L1 L2,
LWP (pi1,sig1) (L_if (subst_b_lst l b) (L_subst l L1) (L_subst l L2)) ->
LWP (pi1,sig1) (L_subst l (L_if b L1 L2)).
Ltac apply_entails_subst_lst2update_store id :=
match goal with
| id: subst_lst2update_store ?l ?P' ?s ?h |- subst_lst2update_store ?l ?P ?s ?h =>
eapply entails_subst_lst2update_store with P'; [red; simpl; intros; idtac | auto]
end.
Ltac cut_replace_list P :=
match goal with
| |- subst_lst2update_store ?l ?P' ?s ?h =>
cut (subst_lst2update_store l P s h);
[intro cut_replace_listA1; apply_entails_subst_lst2update_store cut_replace_listA1 | idtac]
end.
Lemma subst_lst2update_store_fresh: forall l x' e s h P,
fresh_lst x' l ->
subst_lst2update_store l P (store.update x' (eval e s) s) h ->
subst_lst2update_store l (fun s' h' => P (store.update x' (eval e s) s') h') s h.
Lemma subst_lst2update_store_fresh': forall l x' e s h P,
fresh_lst x' l ->
subst_lst2update_store l (fun s' h' => P (store.update x' (eval e s) s') h') s h ->
subst_lst2update_store l P (store.update x' (eval e s) s) h.
Lemma intro_fresh_var' : forall l x x' e s h P,
fresh_lst x' l ->
fresh_e x' (var_e x) ->
inde (x'::nil) P ->
subst_lst2update_store l (fun s' h' => P (store.update x (eval (var_e x') s') s') h') (store.update x' (eval e s) s) h ->
subst_lst2update_store l (fun s' h' => P (store.update x (eval e s) s') h') s h.
Lemma intro_fresh_var : forall l x x' e s h L,
fresh_lst x' l -> fresh_L_assrt x' L -> fresh_e x' (var_e x) ->
subst_lst2update_store
l (fun s' h' => L_assrt_interp L (store.update x (eval (var_e x') s') s') h') (store.update x' (eval e s) s) h ->
subst_lst2update_store
l (fun s' h' => L_assrt_interp L (store.update x (eval e s) s') h') s h.
Lemma LWP_soundness: forall P Q, LWP P Q ->
(assrt_interp P) ==> (L_assrt_interp Q).
Fixpoint wp_frag (Q: option L_assrt) (c: cmd) {struct c}: option L_assrt :=
match c with
skip => match Q with
None => None
| Some Q' => Some Q'
end
| assign v e => match Q with
None => None
| Some Q' => Some (L_subst ((v,e)::nil) Q')
end
| lookup v e => match Q with
None => None
| Some Q' => Some (L_lookup v e Q')
end
| mutation e1 e2 => match Q with
None => None
| Some Q' => Some (L_mutation e1 e2 Q')
end
| seq c1 c2 => wp_frag (wp_frag Q c2) c1
| ifte b thendo c1 elsedo c2 => match (wp_frag Q c1) with
None => None
| Some Q1 => match (wp_frag Q c2) with
None => None
| Some Q2 => Some (L_if b (Q1) (Q2))
end
end
| while a c => None
| malloc v e => None
| free e => None
end.
Lemma wp_frag_None_is_None: forall c, wp_frag None c = None.
Lemma wp_frag_soudness: forall c Q Q',
wp_frag (Some Q) c = Some Q' -> {{ L_assrt_interp Q' }} c {{ L_assrt_interp Q }}.
Lemma LWP_use: forall c P Q R,
wp_frag (Some (L_elt Q)) c = Some R ->
LWP P R ->
{{ assrt_interp P }} c {{ assrt_interp Q }}.
Lemma LWP_subst_lookup' : forall pi1 sig1 e1 e2 e x x' l L,
(forall s,eval_b pi1 s = true -> (eval_b (e1 == (subst_e_lst l e))) s = true) ->
x' = (max (max (var_max_lst l) (var_max_L_assrt L)) x) + 1 ->
LWP (pi1,star sig1 (singl e1 e2)) (L_subst ((x,(var_e x'))::l ++ ((x',e2)::nil)) L) ->
LWP (pi1,star sig1 (singl e1 e2)) (L_subst l (L_lookup x e L)).
Fixpoint Sigma_assoc_left (t t': Sigma) {struct t}: Sigma :=
match t with
star sig1 sig2 => Sigma_assoc_left (sig2) (star t' sig1)
| _ => match t' with
epsi => t
| _ => star t' t
end
end.
Definition Sigma_com (t: Sigma) : Sigma :=
match t with
star sig1 sig2 => star sig2 sig1
| _ => t
end.
Ltac Rotate_LWP_sig_lhs :=
match goal with
| |- LWP (?pi,?sig) ?L' =>
eapply LWP_precond_stre with (
(pi, Sigma_clean_epsi (Sigma_assoc_left (Sigma_com sig) epsi))
); [apply ps1_soundness; simpl; ps1_Resolve| simpl]
end.
Ltac LWP_resolve :=
match goal with
| |- LWP (?pi1, ?sig1) (L_elt (?pi2, ?sig2)) => eapply LWP_entail; [eapply ps1_soundness; ps1_Resolve]
| |- LWP (?pi1, star ?sig1 (singl ?e1 ?e2)) (L_mutation ?e3 ?e4 ?L') =>
(eapply LWP_mutation; [Omega_exprb | LWP_resolve] || Rotate_LWP_sig_lhs; idtac)
| |- LWP (?pi1, star ?sig1 (cell ?e1)) (L_mutation ?e3 ?e4 ?L') =>
(eapply LWP_mutation'; [Omega_exprb | LWP_resolve] || Rotate_LWP_sig_lhs; idtac)
| |- LWP (?pi1, star ?sig1 (singl ?e1 ?e2)) (L_lookup ?x ?e ?L') =>
(eapply LWP_lookup; [Omega_exprb | LWP_resolve] || Rotate_LWP_sig_lhs; idtac)
| |- LWP ?L (L_subst ?l (L_elt ?L')) => eapply LWP_subst_elt; simpl; idtac
| |- LWP ?L (L_subst ?l (L_subst ?l' ?L')) => eapply LWP_subst_subst; simpl; idtac
| |- LWP ?L (L_subst ?l (L_lookup ?x ?e ?L')) =>
(eapply LWP_subst_lookup'; [Omega_exprb | simpl; intuition | LWP_resolve] || Rotate_LWP_sig_lhs; idtac)
| |- LWP ?L (L_subst ?l (L_mutation ?e1 ?e2 ?L')) => eapply LWP_subst_mutation; simpl; idtac
| |- LWP ?L (L_subst ?l (L_if ?b ?L1 ?L2)) => eapply LWP_subst_if; simpl; idtac
| |- LWP ?L (L_if ?b ?L1 ?L2) => eapply LWP_if; simpl; idtac
end.
Ltac LWP_Resolve := Rotate_LWP_sig_lhs; repeat LWP_resolve.
Definition LWP_step (pi: Pi) (sig: Sigma) (A: L_assrt) : option (list ((Pi * Sigma) * L_assrt)) :=
match A with
| L_elt L =>
if (frag_decision (pi,sig) L) then
Some nil else None
| L_subst l L =>
match L with
| L_elt L' => Some (((pi,sig), L_elt (subst_assrt_lst l L'))::nil)
| L_subst l' L' => Some (((pi,sig), L_subst (l'++ l) L')::nil)
| L_lookup x e L' =>
match (Sigma_get_cell_val (subst_e_lst l e) sig pi) with
| None => None
| Some e' =>
let x' := (max (max (var_max_lst l) (var_max_L_assrt L')) x) + 1 in (
Some (((pi,sig),(L_subst ((x,(var_e x'))::l ++ ((x',e')::nil)) L'))::nil)
)
end
| L_mutation e1 e2 L' =>
Some (((pi,sig),(L_mutation (subst_e_lst l e1) (subst_e_lst l e2) (L_subst l L')))::nil)
| L_if b L1 L2 =>
Some (((pi,sig), L_if (subst_b_lst l b) (L_subst l L1) (L_subst l L2))::nil)
end
| L_lookup x e L =>
match (Sigma_get_cell_val e sig pi) with
| None => None
| Some e' => Some (((pi,sig),(L_subst ((x,e')::nil) L))::nil)
end
| L_mutation e1 e2 L =>
match (Sigma_elt_term_extract' (cell e1) sig pi) with
| None => None
| Some sig' => Some (((pi,star (singl e1 e2) sig'),L)::nil)
end
| L_if b L1 L2 =>
Some (((pi &&& b,sig),L1)::((pi &&& (! b),sig),L2)::nil)
end.
Fixpoint L_assrt_size (A: L_assrt) : nat :=
match A with
L_elt P => 2
| L_subst l P => 2 + L_assrt_size P
| L_lookup x e P => 2 + L_assrt_size P
| L_mutation e1 e2 P => 2 + L_assrt_size P
| L_if b L1 L2 => 2 + L_assrt_size L1 + L_assrt_size L2
end.
Fixpoint LWP_list (l: (list ((Pi * Sigma) * L_assrt))) : option (list ((Pi * Sigma) * L_assrt)) :=
match l with
| nil => Some nil
| ((pi,sig),A)::tl =>
match (LWP_step pi sig A) with
| None => None
| Some l' =>
match (LWP_list tl) with
| None => None
| Some l'' => Some (l' ++ l'')
end
end
end.
Fixpoint LWP_list_rec (l: (list ((Pi * Sigma) * L_assrt))) (size: nat) {struct size}: option (list ((Pi * Sigma) * L_assrt)) :=
match size with
| 0 => Some l
| S size' =>
match (LWP_list l) with
| None => None
| Some l' => LWP_list_rec l' size'
end
end.
Definition frag_hoare (P Q: Pi * Sigma) (c: cmd) : bool :=
match (wp_frag (Some (L_elt Q)) c) with
| None => false
| Some L =>
let (p,s) := P in (
match (LWP_list_rec (((p,s),L)::nil) (L_assrt_size L)) with
| Some nil => true
| _ => false
end
)
end.
Opaque frag_decision.
Lemma LWP_step_correct: forall pi sig A l,
LWP_step pi sig A = Some l ->
(forall pi' sig' A',
In ((pi',sig'),A') l -> ((assrt_interp (pi',sig')) ==> (L_assrt_interp A'))
) ->
((assrt_interp (pi,sig)) ==> (L_assrt_interp A)).
Transparent frag_decision.
Lemma LWP_list_correct: forall l l',
LWP_list l = Some l' ->
(forall pi sig A, In ((pi,sig),A) l' ->
((assrt_interp (pi,sig)) ==> (L_assrt_interp A))) ->
(forall pi sig A, In ((pi,sig),A) l -> ((assrt_interp (pi,sig)) ==> (L_assrt_interp A))).
Lemma LWP_list_rec_correct: forall a l l',
LWP_list_rec l a = Some l' ->
(forall pi sig A, In ((pi,sig),A) l' ->
((assrt_interp (pi,sig)) ==> (L_assrt_interp A))) ->
(forall pi sig A, In ((pi,sig),A) l -> ((assrt_interp (pi,sig)) ==> (L_assrt_interp A))).
Lemma frag_hoare_correct: forall P Q c,
frag_hoare P Q c = true ->
{{assrt_interp P}}
c
{{assrt_interp Q}}.