# HG changeset patch # User nipkow # Date 1390230859 -3600 # Node ID f69530f22f5a27a6a054c6875ca3598024728765 # Parent c602013f127e233f2054268232d5aad85ffdfdae tuned names diff -r c602013f127e -r f69530f22f5a src/HOL/IMP/Abs_Int2.thy --- a/src/HOL/IMP/Abs_Int2.thy Mon Jan 20 12:20:23 2014 +0100 +++ b/src/HOL/IMP/Abs_Int2.thy Mon Jan 20 16:14:19 2014 +0100 @@ -74,14 +74,14 @@ subsubsection "Backward analysis" -fun inv_aval'' :: "aexp \ 'av \ 'av st option \ 'av st option" where -"inv_aval'' (N n) a S = (if test_num' n a then S else None)" | -"inv_aval'' (V x) a S = (case S of None \ None | Some S \ +fun inv_aval' :: "aexp \ 'av \ 'av st option \ 'av st option" where +"inv_aval' (N n) a S = (if test_num' n a then S else None)" | +"inv_aval' (V x) a S = (case S of None \ None | Some S \ let a' = fun S x \ a in if a' = \ then None else Some(update S x a'))" | -"inv_aval'' (Plus e1 e2) a S = +"inv_aval' (Plus e1 e2) a S = (let (a1,a2) = inv_plus' a (aval'' e1 S) (aval'' e2 S) - in inv_aval'' e1 a1 (inv_aval'' e2 a2 S))" + in inv_aval' e1 a1 (inv_aval' e2 a2 S))" text{* The test for @{const bot} in the @{const V}-case is important: @{const bot} indicates that a variable has no possible values, i.e.\ that the current @@ -93,17 +93,17 @@ making the analysis less precise. *} -fun inv_bval'' :: "bexp \ bool \ 'av st option \ 'av st option" where -"inv_bval'' (Bc v) res S = (if v=res then S else None)" | -"inv_bval'' (Not b) res S = inv_bval'' b (\ res) S" | -"inv_bval'' (And b1 b2) res S = - (if res then inv_bval'' b1 True (inv_bval'' b2 True S) - else inv_bval'' b1 False S \ inv_bval'' b2 False S)" | -"inv_bval'' (Less e1 e2) res S = +fun inv_bval' :: "bexp \ bool \ 'av st option \ 'av st option" where +"inv_bval' (Bc v) res S = (if v=res then S else None)" | +"inv_bval' (Not b) res S = inv_bval' b (\ res) S" | +"inv_bval' (And b1 b2) res S = + (if res then inv_bval' b1 True (inv_bval' b2 True S) + else inv_bval' b1 False S \ inv_bval' b2 False S)" | +"inv_bval' (Less e1 e2) res S = (let (a1,a2) = inv_less' res (aval'' e1 S) (aval'' e2 S) - in inv_aval'' e1 a1 (inv_aval'' e2 a2 S))" + in inv_aval' e1 a1 (inv_aval' e2 a2 S))" -lemma inv_aval''_correct: "s : \\<^sub>o S \ aval e s : \ a \ s : \\<^sub>o (inv_aval'' e a S)" +lemma inv_aval'_correct: "s : \\<^sub>o S \ aval e s : \ a \ s : \\<^sub>o (inv_aval' e a S)" proof(induction e arbitrary: a S) case N thus ?case by simp (metis test_num') next @@ -122,7 +122,7 @@ by (auto split: prod.split) qed -lemma inv_bval''_correct: "s : \\<^sub>o S \ bv = bval b s \ s : \\<^sub>o(inv_bval'' b bv S)" +lemma inv_bval'_correct: "s : \\<^sub>o S \ bv = bval b s \ s : \\<^sub>o(inv_bval' b bv S)" proof(induction b arbitrary: S bv) case Bc thus ?case by simp next @@ -133,12 +133,12 @@ next case (Less e1 e2) thus ?case by(auto split: prod.split) - (metis (lifting) inv_aval''_correct aval''_correct inv_less') + (metis (lifting) inv_aval'_correct aval''_correct inv_less') qed definition "step' = Step (\x e S. case S of None \ None | Some S \ Some(update S x (aval' e S))) - (\b S. inv_bval'' b True S)" + (\b S. inv_bval' b True S)" definition AI :: "com \ 'av st option acom option" where "AI c = pfp (step' \) (bot c)" @@ -146,23 +146,23 @@ lemma strip_step'[simp]: "strip(step' S c) = strip c" by(simp add: step'_def) -lemma top_on_inv_aval'': "\ top_on_opt S X; vars e \ -X \ \ top_on_opt (inv_aval'' e a S) X" +lemma top_on_inv_aval': "\ top_on_opt S X; vars e \ -X \ \ top_on_opt (inv_aval' e a S) X" by(induction e arbitrary: a S) (auto simp: Let_def split: option.splits prod.split) -lemma top_on_inv_bval'': "\top_on_opt S X; vars b \ -X\ \ top_on_opt (inv_bval'' b r S) X" -by(induction b arbitrary: r S) (auto simp: top_on_inv_aval'' top_on_sup split: prod.split) +lemma top_on_inv_bval': "\top_on_opt S X; vars b \ -X\ \ top_on_opt (inv_bval' b r S) X" +by(induction b arbitrary: r S) (auto simp: top_on_inv_aval' top_on_sup split: prod.split) lemma top_on_step': "top_on_acom C (- vars C) \ top_on_acom (step' \ C) (- vars C)" unfolding step'_def by(rule top_on_Step) - (auto simp add: top_on_top top_on_inv_bval'' split: option.split) + (auto simp add: top_on_top top_on_inv_bval' split: option.split) subsubsection "Correctness" lemma step_step': "step (\\<^sub>o S) (\\<^sub>c C) \ \\<^sub>c (step' S C)" unfolding step_def step'_def by(rule gamma_Step_subcomm) - (auto simp: intro!: aval'_correct inv_bval''_correct in_gamma_update split: option.splits) + (auto simp: intro!: aval'_correct inv_bval'_correct in_gamma_update split: option.splits) lemma AI_correct: "AI c = Some C \ CS c \ \\<^sub>c C" proof(simp add: CS_def AI_def) @@ -204,7 +204,7 @@ apply simp by (simp add: mono_aval') -lemma mono_inv_aval'': "r1 \ r2 \ S1 \ S2 \ inv_aval'' e r1 S1 \ inv_aval'' e r2 S2" +lemma mono_inv_aval': "r1 \ r2 \ S1 \ S2 \ inv_aval' e r1 S1 \ inv_aval' e r2 S2" apply(induction e arbitrary: r1 r2 S1 S2) apply(auto simp: test_num' Let_def inf_mono split: option.splits prod.splits) apply (metis mono_gamma subsetD) @@ -213,19 +213,19 @@ apply(metis mono_aval'' mono_inv_plus'[simplified less_eq_prod_def] fst_conv snd_conv) done -lemma mono_inv_bval'': "S1 \ S2 \ inv_bval'' b bv S1 \ inv_bval'' b bv S2" +lemma mono_inv_bval': "S1 \ S2 \ inv_bval' b bv S1 \ inv_bval' b bv S2" apply(induction b arbitrary: bv S1 S2) apply(simp) apply(simp) apply simp apply(metis order_trans[OF _ sup_ge1] order_trans[OF _ sup_ge2]) apply (simp split: prod.splits) -apply(metis mono_aval'' mono_inv_aval'' mono_inv_less'[simplified less_eq_prod_def] fst_conv snd_conv) +apply(metis mono_aval'' mono_inv_aval' mono_inv_less'[simplified less_eq_prod_def] fst_conv snd_conv) done theorem mono_step': "S1 \ S2 \ C1 \ C2 \ step' S1 C1 \ step' S2 C2" unfolding step'_def -by(rule mono2_Step) (auto simp: mono_aval' mono_inv_bval'' split: option.split) +by(rule mono2_Step) (auto simp: mono_aval' mono_inv_bval' split: option.split) lemma mono_step'_top: "C1 \ C2 \ step' \ C1 \ step' \ C2" by (metis mono_step' order_refl) diff -r c602013f127e -r f69530f22f5a src/HOL/IMP/Abs_Int2_ivl.thy --- a/src/HOL/IMP/Abs_Int2_ivl.thy Mon Jan 20 12:20:23 2014 +0100 +++ b/src/HOL/IMP/Abs_Int2_ivl.thy Mon Jan 20 16:14:19 2014 +0100 @@ -361,8 +361,8 @@ where \ = \_ivl and num' = num_ivl and plus' = "op +" and test_num' = in_ivl and inv_plus' = inv_plus_ivl and inv_less' = inv_less_ivl -defines inv_aval_ivl is inv_aval'' -and inv_bval_ivl is inv_bval'' +defines inv_aval_ivl is inv_aval' +and inv_bval_ivl is inv_bval' and step_ivl is step' and AI_ivl is AI and aval_ivl' is aval''