Theory RefineMonadicVCG
theory RefineMonadicVCG
imports "NREST" "DataRefinement"
"Case_Labeling.Case_Labeling"
begin
method repeat_all_new methods m = (m;repeat_all_new ‹m›)?
lemma R_intro: "A ≤ ⇓Id B ⟹ A ≤ B" by simp
subsection "ASSERT"
lemma le_R_ASSERTI: "(Φ ⟹ M ≤ ⇓ R M') ⟹ M ≤ ⇓ R (ASSERT Φ ⤜ (λ_. M'))"
by (auto simp: pw_le_iff refine_pw_simps)
lemma T_ASSERT[vcg_simp_rules]: "Some t ≤ lst (ASSERT Φ) Q ⟷ Some t ≤ Q () ∧ Φ"
by (cases Φ) vcg'
lemma T_ASSERT_I: "(Φ ⟹ Some t ≤ Q ()) ⟹ Φ ⟹ Some t ≤ lst (ASSERT Φ) Q"
by (simp add: T_ASSERT T_RETURNT)
lemma T_RESTemb_iff: "Some t'
≤ lst (REST (emb' P t)) Q ⟷ (∀x. P x ⟶ Some (t' + t x) ≤ Q x ) "
by (auto simp: emb'_def T_pw mii_alt aux1)
lemma T_RESTemb: "(⋀x. P x ⟹ Some (t' + t x) ≤ Q x)
⟹ Some t' ≤ lst (REST (emb' P t)) Q"
by (auto simp: T_RESTemb_iff)
lemma T_SPEC: "(⋀x. P x ⟹ Some (t' + t x) ≤ Q x)
⟹ Some t' ≤ lst (SPEC P t) Q"
unfolding SPEC_REST_emb'_conv
by (auto simp: T_RESTemb_iff)
lemma T_SPECT_I: "(Some (t' + t ) ≤ Q x)
⟹ Some t' ≤ lst (SPECT [ x ↦ t]) Q"
by(auto simp: T_pw mii_alt aux1)
lemma mm2_map_option:
assumes "Some (t'+t) ≤ mm2 (Q x) (x2 x)"
shows "Some t' ≤ mm2 (Q x) (map_option ((+) t) (x2 x))"
proof(cases "Q x")
case None
from assms None show ?thesis
by (auto simp: mm2_def split: option.splits if_splits)
next
case (Some a)
have arith: "¬ a < b ⟹ t' + t ≤ a - b ⟹ a < t + b ⟹ False" for b
by (cases a; cases b; cases t'; cases t) auto
moreover have arith': "¬ a < b ⟹ t' + t ≤ a - b ⟹ ¬ a < t + b ⟹ t' ≤ a - (t + b)" for b
by (cases a; cases b; cases t'; cases t) auto
ultimately show ?thesis
using Some assms by (auto simp: mm2_def split: option.splits if_splits)
qed
lemma T_consume: "(Some (t' + t) ≤ lst M Q)
⟹ Some t' ≤ lst (consume M t) Q"
by (auto intro!: mm2_map_option simp: consume_def T_pw mii_alt miiFailt
split: nrest.splits option.splits if_splits)
definition "valid t Q M = (Some t ≤ lst M Q)"
subsection ‹VCG splitter›
ML ‹
structure VCG_Case_Splitter = struct
fun dest_case ctxt t =
case strip_comb t of
(Const (case_comb, _), args) =>
(case Ctr_Sugar.ctr_sugar_of_case ctxt case_comb of
NONE => NONE
| SOME {case_thms, ...} =>
let
val lhs = Thm.prop_of (hd (case_thms))
|> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst;
val arity = length (snd (strip_comb lhs));
in
SOME (nth args (arity - 1), case_thms)
end)
| _ => NONE;
fun rewrite_with_asm_tac ctxt k =
Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
Local_Defs.unfold0_tac ctxt' [nth prems k]) ctxt;
fun split_term_tac ctxt case_term = (
case dest_case ctxt case_term of
NONE => no_tac
| SOME (arg,case_thms) => let
val stac = asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps case_thms)
in
(
Induct.cases_tac ctxt false [[SOME arg]] NONE []
THEN_ALL_NEW stac
)
end 1
)
fun split_tac ctxt = Subgoal.FOCUS_PARAMS (fn {context = ctxt, ...} => ALLGOALS (
SUBGOAL (fn (t, _) => case Logic.strip_imp_concl t of
@{mpat "Trueprop (Some _ ≤ lst ?prog _)"} => split_term_tac ctxt prog
| @{mpat "Trueprop (progress ?prog)"} => split_term_tac ctxt prog
| @{mpat "Trueprop (Case_Labeling.CTXT _ _ _ (valid _ _ ?prog))"} => split_term_tac ctxt prog
| _ => no_tac
))
) ctxt
THEN_ALL_NEW TRY o Hypsubst.hyp_subst_tac ctxt
end
›
method_setup vcg_split_case =
‹Scan.succeed (fn ctxt => SIMPLE_METHOD' (CHANGED o (VCG_Case_Splitter.split_tac ctxt)))›
subsection ‹mm3 and emb›
lemma Some_eq_mm3_Some_conv[vcg_simp_rules]: "Some t = mm3 t' (Some t'') ⟷ (t'' ≤ t' ∧ t = enat (t' - t''))"
by (simp add: mm3_def)
lemma Some_eq_mm3_Some_conv': "mm3 t' (Some t'') = Some t ⟷ (t'' ≤ t' ∧ t = enat (t' - t''))"
using Some_eq_mm3_Some_conv by metis
lemma Some_le_emb'_conv[vcg_simp_rules]: "Some t ≤ emb' Q ft x ⟷ Q x ∧ t ≤ ft x"
by (auto simp: emb'_def)
lemma Some_eq_emb'_conv[vcg_simp_rules]: "emb' Q tf s = Some t ⟷ (Q s ∧ t = tf s)"
unfolding emb'_def by(auto split: if_splits)
subsection ‹Setup Labeled VCG›
context
begin
interpretation Labeling_Syntax .
lemma LCondRule:
fixes IC CT defines "CT' ≡ (''cond'', IC, []) # CT "
assumes "b ⟹ C⟨Suc IC,(''the'', IC, []) # (''cond'', IC, []) # CT,OC1: valid t Q c1 ⟩"
and "~b ⟹ C⟨Suc OC1,(''els'', Suc OC1, []) # (''cond'', IC, []) # CT,OC: valid t Q c2 ⟩"
shows "C⟨IC,CT,OC: valid t Q (if b then c1 else c2)⟩"
using assms(2-) unfolding LABEL_simps by (simp add: valid_def)
lemma LouterCondRule:
fixes IC CT defines "CT' ≡ (''cond2'', IC, []) # CT "
assumes "b ⟹ C⟨Suc IC,(''the'', IC, []) # (''cond2'', IC, []) # CT,OC1: t ≤ A ⟩"
and "~b ⟹ C⟨Suc OC1,(''els'', Suc OC1, []) # (''cond2'', IC, []) # CT,OC: t ≤ B ⟩"
shows "C⟨IC,CT,OC: t ≤ (if b then A else B)⟩"
using assms(2-) unfolding LABEL_simps by (simp add: valid_def)
lemma While:
assumes "I s0" "(⋀s. I s ⟹ b s ⟹ Some 0 ≤ lst (C s) (λs'. mm3 (E s) (if I s' then Some (E s') else None)))"
"(⋀s. progress (C s))"
"(⋀x. ¬ b x ⟹ I x ⟹ (E x) ≤ (E s0) ⟹ Some (t + enat ((E s0) - E x)) ≤ Q x)"
shows "Some t ≤ lst (whileIET I E b C s0) Q"
apply(rule whileIET_rule'[THEN T_conseq4])
subgoal using assms(2) by simp
subgoal using assms(3) by simp
subgoal using assms(1) by simp
subgoal for x using assms(4) by (cases "I x") (auto simp: Some_eq_mm3_Some_conv' split: if_splits)
done
definition "monadic_WHILE b f s ≡ do {
RECT (λD s. do {
bv ← b s;
if bv then do {
s ← f s;
D s
} else do {RETURNT s}
}) s
}"
lemma monadic_WHILE_mono:
assumes "⋀x. bm x ≤ bm' x"
and "⋀x t. nofailT (bm' x) ⟹ inresT (bm x) True t ⟹ c x ≤ c' x"
shows "(monadic_WHILE bm c x) ≤ (monadic_WHILE bm' c' x)"
unfolding monadic_WHILE_def apply(rule RECT_mono)
subgoal by (refine_mono)
using assms by (auto intro!: bindT_mono)
lemma z: "inresT A x t ⟹ A ≤ B ⟹ inresT B x t"
by (meson fail_inresT pw_le_iff)
lemma monadic_WHILE_mono':
assumes
"⋀x. bm x ≤ bm' x"
and "⋀x t. nofailT (bm' x) ⟹ inresT (bm' x) True t ⟹ c x ≤ c' x"
shows " (monadic_WHILE bm c x) ≤ (monadic_WHILE bm' c' x)"
unfolding monadic_WHILE_def apply(rule RECT_mono)
subgoal by(refine_mono)
apply(rule bindT_mono)
apply fact
using assms by (auto intro!: bindT_mono dest: z[OF _ assms(1)])
lemma monadic_WHILE_refine:
assumes
"(x, x') ∈ R"
"⋀x x'. (x, x') ∈ R ⟹ bm x ≤ ⇓Id (bm' x')"
and "⋀x x' t. (x, x') ∈ R ⟹ nofailT (bm' x') ⟹ inresT (bm' x') True t ⟹ c x ≤ ⇓R (c' x')"
shows "(monadic_WHILE bm c x) ≤ ⇓R (monadic_WHILE bm' c' x')"
unfolding monadic_WHILE_def apply(rule RECT_refine[OF _ assms(1)])
subgoal by(refine_mono)
apply(rule bindT_refine'[OF assms(2)])
subgoal by auto
by (auto intro!: assms(3) bindT_refine RETURNT_refine)
lemma monadic_WHILE_aux: "monadic_WHILE b f s = monadic_WHILEIT (λ_. True) b f s"
unfolding monadic_WHILEIT_def monadic_WHILE_def
by simp
lemma "lst (c x) Q = Some t ⟹ Some t ≤ lst (c x) Q'"
apply(rule T_conseq6) oops
lemma TbindT_I2: "tt ≤ lst M (λy. lst (f y) Q) ⟹ tt ≤ lst (M ⤜ f) Q"
by (simp add: T_bindT)
lemma T_conseq7:
assumes
"lst f Q' ≥ tt"
"⋀x t'' M. f = SPECT M ⟹ M x ≠ None ⟹ Q' x = Some t'' ⟹ (Q x) ≥ Some ( t'')"
shows "lst f Q ≥ tt"
using assms by (cases tt) (auto intro: T_conseq6)
lemma monadic_WHILE_ruleaaa'':
assumes "monadic_WHILE bm c s = r"
assumes IS[vcg_rules]: "⋀s.
lst (bm s) (λb. if b then lst (c s) (λs'. if (s',s)∈R then I s' else None) else Q s) ≥ I s"
assumes wf: "wf R"
shows "lst r Q ≥ I s"
using assms(1)
unfolding monadic_WHILE_def
proof (induction rule: RECT_wf_induct[where R="R"])
case 1
show ?case by fact
next
case 2
then show ?case by refine_mono
next
case step: (3 x D r )
note IH[vcg_rules] = step.IH[OF _ refl]
note step.hyps[symmetric, simp]
from step.prems
show ?case
apply simp
apply (rule TbindT_I2)
apply(rule T_conseq7)
apply (rule IS)
apply simp
apply safe
subgoal
apply (rule TbindT_I)
apply(rule T_conseq6[where Q'="(λs'. if (s', x) ∈ R then I s' else None)"])
subgoal by simp
by (auto split: if_splits dest: IH)
subgoal by(simp add: T_RETURNT)
done
qed
lemma monadic_WHILE_rule'':
assumes "monadic_WHILE bm c s = r"
assumes IS[vcg_rules]: "⋀s t'. I s = Some t'
⟹ lst (bm s) (λb. if b then lst (c s) (λs'. if (s',s)∈R then I s' else None)else Q s) ≥ Some t'"
assumes "I s = Some t"
assumes wf: "wf R"
shows "lst r Q ≥ Some t"
using assms(1,3)
unfolding monadic_WHILE_def
proof (induction arbitrary: t rule: RECT_wf_induct[where R="R"])
case 1
show ?case by fact
next
case 2
then show ?case by refine_mono
next
case step: (3 x D r t')
note IH[vcg_rules] = step.IH[OF _ refl]
note step.hyps[symmetric, simp]
from step.prems
show ?case
apply simp
apply (rule TbindT_I)
apply(rule T_conseq6)
apply (rule IS) apply simp
apply simp
apply safe
subgoal
apply (rule TbindT_I)
apply(rule T_conseq6[where Q'="(λs'. if (s', x) ∈ R then I s' else None)"])
subgoal by simp
by (auto split: if_splits intro: IH)
subgoal by vcg'
done
qed
lemma whileT_rule''':
fixes I :: "'a ⇒ nat option"
assumes "whileT b c s0 = r"
assumes progress: "⋀s. progress (c s)"
assumes IS[vcg_rules]: "⋀s t t'. I s = Some t ⟹ b s ⟹
lst (c s) (λs'. mm3 t (I s') ) ≥ Some 0"
assumes [simp]: "I s0 = Some t0"
shows "lst r (λx. if b x then None else mm3 t0 (I x)) ≥ Some 0"
apply(rule T_conseq4)
apply(rule whileT_rule''[where I="λs. mm3 t0 (I s)"
and R="measure (the_enat o the o I)", OF assms(1)])
subgoal for s t'
apply(cases "I s"; simp)
subgoal for ti
using IS[of s ti]
apply (cases "c s") apply(simp)
subgoal for M
using progress[of s, THEN progressD, of M]
apply(auto simp: T_pw mm3_Some_conv mii_alt mm2_def mm3_def split: option.splits if_splits)
apply fastforce
subgoal
by (metis enat_ord_simps(1) le_diff_iff le_less_trans option.distinct(1))
subgoal
by (metis diff_is_0_eq' leI less_option_Some option.simps(3) zero_enat_def)
subgoal
by (smt Nat.add_diff_assoc enat_ile enat_ord_code(1) idiff_enat_enat leI le_add_diff_inverse2 nat_le_iff_add option.simps(3))
subgoal
using dual_order.trans by blast
done
done
done
by auto
fun Someplus where
"Someplus None _ = None"
| "Someplus _ None = None"
| "Someplus (Some a) (Some b) = Some (a+b)"
lemma l: "~ (a::enat) < b ⟷ a ≥ b" by auto
lemma kk: "a≥b ⟹ (b::enat) + (a -b) = a"
by (cases a; cases b) auto
lemma Tea: "Someplus A B = Some t ⟷ (∃a b. A = Some a ∧ B = Some b ∧ t = a + b)"
by (cases A; cases B) auto
lemma TTT_Some_nofailT: "lst c Q = Some l ⟹ c ≠ FAILT"
unfolding lst_def mii_alt by auto
lemma GRR: assumes "lst (SPECT Mf) Q = Some l"
shows "Mf x = None ∨ (Q x≠ None ∧ (Q x) ≥ Mf x) "
proof -
from assms have "None ∉ {mii Q (SPECT Mf) x |x. True}"
unfolding lst_def
unfolding Inf_option_def by (auto split: if_splits)
then have "None ≠ (case Mf x of None ⇒ Some ∞
| Some mt ⇒ case Q x of None ⇒ None
| Some rt ⇒ if rt < mt then None else Some (rt - mt))"
unfolding mii_alt mm2_def
by (auto)
then show ?thesis by (auto split: option.splits if_splits)
qed
lemma Someplus_None: "Someplus A B = None ⟷ (A = None ∨ B = None)"
by (cases A; cases B) auto
lemma Somemm3: "A ≥ B ⟹ mm3 A (Some B) = Some (A - B)"
unfolding mm3_def by auto
lemma neueWhile_rule: assumes "monadic_WHILE bm c s0 = r"
and step: "⋀s. I s ⟹
Some 0 ≤ lst (bm s) (λb. if b
then lst (c s) (λs'. (if I s' ∧ (E s' ≤ E s) then Some (enat (E s - E s')) else None))
else mm2 (Q s) (Someplus (Some t) (mm3 (E s0) (Some (E s)))) )
"
and progress: "⋀s. progress (c s)"
and I0: "I s0"
shows "Some t ≤ lst r Q"
proof -
show "Some t ≤ lst r Q"
apply (rule monadic_WHILE_rule''[where I="λs. Someplus (Some t) (mm3 (E s0) ((λe. if I e
then Some (E e) else None) s))" and R="measure (the_enat o the o (λe. if I e
then Some (E e) else None))", simplified])
apply fact
subgoal for s t'
apply(auto split: if_splits)
apply(rule T_conseq4)
apply(rule step)
apply simp
apply auto
proof (goal_cases)
case (1 b t'')
from 1(3) TTT_Some_nofailT obtain M where cs: "c s = SPECT M" by force
{ assume A: "⋀x. M x = None"
with A have "?case" unfolding cs lst_def mii_alt by simp
}
moreover
{ assume "∃x. M x ≠ None"
then obtain x where i: "M x ≠ None" by blast
let ?T = "lst (c s) (λs'. if I s' ∧ E s' ≤ E s then Some (enat (E s - E s')) else None)"
from GRR[OF 1(3)[unfolded cs], where x=x]
i have "(if I x ∧ E x ≤ E s then Some (enat (E s - E x)) else None) ≠ None ∧ M x ≤ (if I x ∧ E x ≤ E s then Some (enat (E s - E x)) else None)"
by simp
then have pf: " I x" "E x ≤ E s" "M x ≤ Some (enat (E s - E x)) " by (auto split: if_splits)
then have "M x < Some ∞"
using enat_ord_code(4) le_less_trans less_option_Some by blast
have "Some t'' = ?T" using 1(3) by simp
also have oo: "?T ≤ mm2 (Some (enat (E s - E x))) (M x)"
unfolding lst_def apply(rule Inf_lower) apply (simp add: mii_alt cs) apply(rule exI[where x=x])
using pf by simp
also from i have o: "… < Some ∞" unfolding mm2_def
apply auto
using fl by blast
finally have tni: "t'' < ∞" by auto
then have tt: "t' + t'' - t'' = t'" apply(cases t''; cases t') by auto
have ka: "⋀x. mii (λs'. if I s' ∧ E s' ≤ E s then Some (enat (E s - E s')) else None) (c s) x ≥ Some t''" unfolding lst_def
using "1"(3) T_pw by fastforce
{ fix x assume nN: "M x ≠ None"
with progress[of s, unfolded cs progress_def] have strict: "Some 0 < M x" by auto
from ka[of x] nN have "E x < E s" unfolding mii_alt cs progress_def mm2_def
using strict less_le zero_enat_def by (fastforce simp: l split: if_splits)
} note strict = this
have ?case
apply(rule T_conseq5[where Q'="(λs'. if I s' ∧ E s' ≤ E s then Some (enat (E s - E s')) else None)"])
using 1(3) apply(auto) [] using 1(2)
apply (auto simp add: tt Tea split: if_splits)
subgoal apply(auto simp add: Some_eq_mm3_Some_conv')
apply(rule strict) using cs by simp
subgoal by(simp add: Some_eq_mm3_Some_conv' Somemm3)
done
}
ultimately show ?case by blast
next
case (2 x t'')
then show ?case unfolding mm3_def mm2_def by (auto simp add: l kk split: if_splits option.splits)
qed
subgoal
using I0 by simp
done
qed
definition monadic_WHILEIE where
"monadic_WHILEIE I E bm c s = monadic_WHILE bm c s"
definition "G b d = (if b then Some d else None)"
definition "H Qs t Es0 Es = mm2 Qs (Someplus (Some t) (mm3 (Es0) (Some (Es))))"
lemma neueWhile_rule':
fixes s0 :: 'a and I :: "'a ⇒ bool" and E :: "'a ⇒ nat"
assumes
step: "(⋀s. I s ⟹ Some 0 ≤ lst (bm s) (λb. if b then lst (c s) (λs'. if I s' ∧ E s' ≤ E s then Some (enat (E s - E s')) else None) else mm2 (Q s) (Someplus (Some t) (mm3 (E s0) (Some (E s))))))"
and progress: "⋀s. progress (c s)"
and i: "I s0"
shows "Some t ≤ lst (monadic_WHILEIE I E bm c s0) Q"
unfolding monadic_WHILEIE_def
apply(rule neueWhile_rule[OF refl]) by fact+
lemma neueWhile_rule'':
fixes s0 :: 'a and I :: "'a ⇒ bool" and E :: "'a ⇒ nat"
assumes
step: "(⋀s. I s ⟹ Some 0 ≤ lst (bm s) (λb. if b then lst (c s) (λs'. G (I s' ∧ E s' ≤ E s) (enat (E s - E s'))) else H (Q s) t (E s0) (E s)))"
and progress: "⋀s. progress (c s)"
and i: "I s0"
shows "Some t ≤ lst (monadic_WHILEIE I E bm c s0) Q"
unfolding monadic_WHILEIE_def apply(rule neueWhile_rule[OF refl, where I=I and E=E ])
using assms unfolding G_def H_def by auto
lemma LmonWhileRule:
fixes IC CT
assumes "V⟨(''precondition'', IC, []),(''monwhile'', IC, []) # CT: I s0⟩"
and "⋀s. I s ⟹ C⟨Suc IC,(''invariant'', Suc IC, []) # (''monwhile'', IC, []) # CT,OC: valid 0 (λb. if b then lst (C s) (λs'. if I s' ∧ E s' ≤ E s then Some (enat (E s - E s')) else None) else mm2 (Q s) (Someplus (Some t) (mm3 (E s0) (Some (E s))))) (bm s)⟩"
and "⋀s. V⟨(''progress'', IC, []),(''monwhile'', IC, []) # CT: progress (C s)⟩"
shows "C⟨IC,CT,OC: valid t Q (monadic_WHILEIE I E bm C s0)⟩"
using assms(2,3,1) unfolding valid_def LABEL_simps
by (rule neueWhile_rule')
lemma LWhileRule:
fixes IC CT
assumes "V⟨(''precondition'', IC, []),(''while'', IC, []) # CT: I s0⟩"
and "⋀s. I s ⟹ b s ⟹ C⟨Suc IC,(''invariant'', Suc IC, []) # (''while'', IC, []) # CT,OC1: valid 0 (λs'. mm3 (E s) (if I s' then Some (E s') else None)) (C s)⟩"
and "⋀s. V⟨(''progress'', IC, []),(''while'', IC, []) # CT: progress (C s)⟩"
and "⋀x. ¬ b x ⟹ I x ⟹ (E x) ≤ (E s0) ⟹ C⟨Suc OC1,(''postcondition'', IC, [])#(''while'', IC, []) # CT,OC: Some (t + enat ((E s0) - E x)) ≤ Q x⟩"
shows "C⟨IC,CT,OC: valid t Q (whileIET I E b C s0)⟩"
using assms unfolding valid_def LABEL_simps
by (rule While)
lemma validD: "valid t Q M ⟹ Some t ≤ lst M Q" by(simp add: valid_def)
lemma LABELs_to_concl:
"C⟨IC, CT, OC: True⟩ ⟹ C⟨IC, CT, OC: P⟩ ⟹ P"
"V⟨x, ct: True⟩ ⟹ V⟨x, ct: P⟩ ⟹ P"
unfolding LABEL_simps .
lemma LASSERTRule:
assumes "V⟨(''ASSERT'', IC, []),CT: Φ⟩"
"C⟨Suc IC, CT,OC: valid t Q (RETURNT ())⟩"
shows "C⟨IC,CT,OC: valid t Q (ASSERT Φ)⟩"
using assms unfolding LABEL_simps by (simp add: valid_def )
lemma LbindTRule:
assumes "C⟨IC,CT,OC: valid t (λy. lst (f y) Q) m⟩"
shows "C⟨IC,CT,OC: valid t Q (bindT m f)⟩"
using assms unfolding LABEL_simps by(simp add: T_bindT valid_def )
lemma LRETURNTRule:
assumes "C⟨IC,CT,OC: Some t ≤ Q x⟩"
shows "C⟨IC,CT,OC: valid t Q (RETURNT x)⟩"
using assms unfolding LABEL_simps by (simp add: valid_def T_RETURNT)
lemma LSELECTRule:
fixes IC CT defines "CT' ≡ (''cond'', IC, []) # CT "
assumes "⋀x. P x ⟹ C⟨Suc IC,(''Some'', IC, []) # (''SELECT'', IC, []) # CT,OC1: valid (t+T) Q (RETURNT (Some x)) ⟩"
and "∀x. ¬ P x ⟹ C⟨Suc OC1,(''None'', Suc OC1, []) # (''SELECT'', IC, []) # CT,OC: valid (t+T) Q (RETURNT None) ⟩"
shows "C⟨IC,CT,OC: valid t Q (SELECT P T)⟩"
using assms(2-) unfolding LABEL_simps by(auto intro: T_SELECT T_SPECT_I simp add: valid_def T_RETURNT)
lemma LRESTembRule:
assumes "⋀x. P x ⟹ C⟨IC,CT,OC: Some (t + T x) ≤ Q x⟩"
shows "C⟨IC,CT,OC: valid t Q (REST (emb' P T))⟩"
using assms unfolding LABEL_simps by (simp add: valid_def T_RESTemb)
lemma LRESTsingleRule:
assumes "C⟨IC,CT,OC: Some (t + t') ≤ Q x⟩"
shows "C⟨IC,CT,OC: valid t Q (REST [x↦t'])⟩"
using assms unfolding LABEL_simps by (simp add: valid_def T_pw mii_alt aux1)
lemma LTTTinRule:
assumes "C⟨IC,CT,OC: valid t Q M⟩"
shows "C⟨IC,CT,OC: Some t ≤ lst M Q⟩"
using assms unfolding LABEL_simps by(simp add: valid_def)
lemma LfinaltimeRule:
assumes "V⟨(''time'', IC, []), CT: t ≤ t' ⟩"
shows "C⟨IC,CT,IC: Some t ≤ Some t'⟩"
using assms unfolding LABEL_simps by simp
lemma LfinalfuncRule:
assumes "V⟨(''func'', IC, []), CT: False ⟩"
shows "C⟨IC,CT,IC: Some t ≤ None⟩"
using assms unfolding LABEL_simps by simp
lemma LembRule:
assumes "V⟨(''time'', IC, []), CT: t ≤ T x ⟩"
and "V⟨(''func'', IC, []), CT: P x ⟩"
shows "C⟨IC,CT,IC: Some t ≤ emb' P T x⟩"
using assms unfolding LABEL_simps by (simp add: emb'_def)
lemma Lmm3Rule:
assumes "V⟨(''time'', IC, []), CT: Va' ≤ Va ∧ t ≤ enat (Va - Va') ⟩"
and "V⟨(''func'', IC, []), CT: b ⟩"
shows "C⟨IC,CT,IC: Some t ≤ mm3 Va (if b then Some Va' else None) ⟩"
using assms unfolding LABEL_simps by (simp add: mm3_def)
lemma LinjectRule:
assumes "Some t ≤ lst A Q ⟹ Some t ≤ lst B Q"
"C⟨IC,CT,OC: valid t Q A⟩"
shows "C⟨IC,CT,OC: valid t Q B⟩"
using assms unfolding LABEL_simps by (simp add: valid_def)
lemma Linject2Rule:
assumes "A = B"
"C⟨IC,CT,OC: valid t Q A⟩"
shows "C⟨IC,CT,OC: valid t Q B⟩"
using assms unfolding LABEL_simps by (simp add: valid_def)
method labeled_VCG_init = ((rule T_specifies_I validD)+), rule Initial_Label
method labeled_VCG_step uses rules = (rule rules[symmetric, THEN Linject2Rule]
LTTTinRule LbindTRule
LembRule Lmm3Rule
LRETURNTRule LASSERTRule LCondRule LSELECTRule
LRESTsingleRule LRESTembRule
LouterCondRule
LfinaltimeRule LfinalfuncRule
LmonWhileRule LWhileRule) | vcg_split_case
method labeled_VCG uses rules = labeled_VCG_init, repeat_all_new ‹labeled_VCG_step rules: rules›
method casified_VCG uses rules = labeled_VCG rules: rules, casify
subsection ‹Examples, labeled vcg›
lemma "do { x ← SELECT P T;
(case x of None ⇒ RETURNT (1::nat) | Some t ⇒ RETURNT (2::nat))
} ≤ SPECT (emb Q T')"
apply labeled_VCG oops
lemma
assumes "b" "c"
shows "do { ASSERT b;
ASSERT c;
RETURNT 1 } ≤ SPECT (emb (λx. x>(0::nat)) 1)"
proof (labeled_VCG, casify)
case ASSERT then show ?case by fact
case ASSERTa then show ?case by fact
case func then show ?case by simp
case time then show ?case by simp
qed
lemma "do {
(if b then RETURNT (1::nat) else RETURNT 2)
} ≤ SPECT (emb (λx. x>0) 1)"
proof (labeled_VCG, casify)
case cond {
case the {
case time
then show ?case by simp
next
case func
then show ?case by simp
}
next
case els {
case func
then show ?case by simp
}
}
qed simp
lemma
assumes "b"
shows "do {
ASSERT b;
(if b then RETURNT (1::nat) else RETURNT 2)
} ≤ SPECT (emb (λx. x>0) 1)"
proof (labeled_VCG, casify)
case ASSERT then show ?case by fact
case cond {
case the {
case time
then show ?case by simp
next
case func
then show ?case by simp
}
next
case els {
case func
then show ?case by simp
}
}
qed simp
end
lemma SPECT_ub': "T≤T' ⟹ SPECT (emb' M' T) ≤ ⇓Id (SPECT (emb' M' T'))"
unfolding emb'_def by (auto simp: pw_le_iff le_funD order_trans refine_pw_simps)
lemma REST_single_rule[vcg_simp_rules]: "Some t ≤ lst (REST [x↦t']) Q ⟷ Some (t+t') ≤ (Q x)"
by (simp add: T_REST aux1)
subsection "progress solver"
method progress methods solver =
(rule asm_rl[of "progress _"],
(simp split: prod.splits | intro allI impI conjI | determ ‹rule progress_rules›
| rule disjI1; progress ‹solver›; fail | rule disjI2; progress ‹solver›; fail | solver)+) []
method progress' methods solver =
(rule asm_rl[of "progress _"], (vcg_split_case | intro allI impI conjI | determ ‹rule progress_rules›
| rule disjI1 disjI2; progress'‹solver› | (solver;fail))+) []
lemma
assumes "(⋀s t. P s = Some t ⟹ ∃s'. Some t ≤ Q s' ∧ (s, s') ∈ R)"
shows SPECT_refine: "SPECT P ≤ ⇓ R (SPECT Q)"
proof-
have "P x ≤ ⨆ {Q a |a. (x, a) ∈ R}" for x
proof(cases "P x")
case [simp]: (Some y)
from assms[of x, OF Some] obtain s' where s': "Some y ≤ Q s'" "(x, s') ∈ R"
by blast+
show ?thesis
by (rule order.trans[where b="Q s'"]) (auto intro: s' Sup_upper)
qed auto
then show ?thesis
by (auto simp add: conc_fun_def le_fun_def)
qed
subsection ‹more stuff involving mm3 and emb›
lemma Some_le_mm3_Some_conv[vcg_simp_rules]: "Some t ≤ mm3 t' (Some t'') ⟷ (t'' ≤ t' ∧ t ≤ enat (t' - t''))"
by (simp add: mm3_def)
lemma embtimeI: "T ≤ T' ⟹ emb P T ≤ emb P T'" unfolding emb'_def by (auto simp: le_fun_def split: if_splits)
lemma not_cons_is_Nil_conv[simp]: "(∀y ys. l ≠ y # ys) ⟷ l=[]" by (cases l) auto
lemma mm3_Some0_eq[simp]: "mm3 t (Some 0) = Some t"
by (auto simp: mm3_def)
lemma ran_emb': "c ∈ ran (emb' Q t) ⟷ (∃s'. Q s' ∧ t s' = c)"
by(auto simp: emb'_def ran_def)
lemma ran_emb_conv: "Ex Q ⟹ ran (emb Q t) = {t}"
by (auto simp: ran_emb')
lemma in_ran_emb_special_case: "c∈ran (emb Q t) ⟹ c≤t"
apply (cases "Ex Q")
subgoal by (auto simp: ran_emb_conv)
subgoal by (auto simp: emb'_def)
done
lemma dom_emb'_eq[simp]: "dom (emb' Q f) = Collect Q"
by(auto simp: emb'_def split: if_splits)
lemma emb_le_emb: "emb' P T ≤ emb' P' T' ⟷ (∀x. P x ⟶ P' x ∧ T x ≤ T' x)"
unfolding emb'_def by (auto simp: le_fun_def split: if_splits)
subsection ‹VCG for monadic programs›
subsubsection ‹old›
declare mm3_Some_conv [vcg_simp_rules]
lemma SS[vcg_simp_rules]: "Some t = Some t' ⟷ t = t'" by simp
lemma SS': "(if b then Some t else None) = Some t' ⟷ (b ∧ t = t')" by simp
term "(case s of (a,b) ⇒ M a b)"
lemma case_T[vcg_rules]: "(⋀a b. s = (a, b) ⟹ t ≤ lst Q (M a b)) ⟹ t ≤ lst Q (case s of (a,b) ⇒ M a b)"
by (simp add: split_def)
subsubsection ‹new setup›
named_theorems vcg_rules'
lemma if_T[vcg_rules']: "(b ⟹ t ≤ lst Ma Q) ⟹ (¬b ⟹ t ≤ lst Mb Q) ⟹ t ≤ lst (if b then Ma else Mb) Q"
by (simp add: split_def)
lemma RETURNT_T_I[vcg_rules']: "t ≤ Q x ⟹ t ≤ lst (RETURNT x) Q"
by (simp add: T_RETURNT)
declare T_SPECT_I [vcg_rules']
declare TbindT_I [vcg_rules']
declare T_RESTemb [vcg_rules']
declare T_ASSERT_I [vcg_rules']
declare While[ vcg_rules']
named_theorems vcg_simps'
declare option.case [vcg_simps']
declare neueWhile_rule'' [vcg_rules']
method vcg'_step methods solver uses rules =
(intro rules vcg_rules' | vcg_split_case | (progress simp;fail) | (solver; fail))
method vcg' methods solver uses rules = repeat_all_new ‹vcg'_step solver rules: rules›
declare T_SELECT [vcg_rules']
lemma "⋀c. do { c ← RETURNT None;
(case_option (RETURNT (1::nat)) (λp. RETURNT (2::nat))) c
} ≤ SPECT (emb (λx. x>(0::nat)) 1)"
apply(rule T_specifies_I)
apply(vcg'‹-›) unfolding option.case
oops
thm option.case
subsection ‹setup for ‹refine_vcg››
lemma If_refine[refine]: "b = b' ⟹
(b ⟹ b' ⟹ S1 ≤ ⇓ R S1') ⟹
(¬ b ⟹ ¬ b' ⟹ S2 ≤ ⇓ R S2') ⟹ (if b then S1 else S2) ≤ ⇓ R (if b' then S1' else S2')"
by auto
lemma Case_option_refine[refine]: "(x,x')∈ ⟨S⟩option_rel ⟹
(⋀y y'. (y,y')∈S ⟹ S2 y ≤ ⇓ R (S2' y')) ⟹ S1 ≤ ⇓ R S1'
⟹ (case x of None ⇒ S1 | Some y ⇒ S2 y) ≤ ⇓ R (case x' of None ⇒ S1' | Some y' ⇒ S2' y')"
by(auto split: option.split)
lemma conc_fun_Id_refined[refine0]: "⋀S. S ≤ ⇓ Id S" by simp
lemma conc_fun_ASSERT_refine[refine0]: "Φ ⟹ (Φ ⟹ S ≤ ⇓ R S') ⟹ ASSERT Φ ⤜ (λ_. S) ≤ ⇓ R S'"
by auto
declare le_R_ASSERTI [refine0]
declare bindT_refine [refine]
declare WHILET_refine [refine]
lemma SPECT_refine_vcg_cons[refine_vcg_cons]: "m ≤ SPECT Φ ⟹ (⋀x. Φ x ≤ Ψ x) ⟹ m ≤ SPECT Ψ"
by (metis dual_order.trans le_funI nres_order_simps(2))
end