diff -r 4334c91b7405 -r f682f3f7b726 src/HOL/IMP/Abs_Int2.thy --- a/src/HOL/IMP/Abs_Int2.thy Wed Nov 23 23:31:32 2011 +0100 +++ b/src/HOL/IMP/Abs_Int2.thy Thu Nov 24 19:58:37 2011 +0100 @@ -64,32 +64,32 @@ end -instantiation up :: (WN)WN +instantiation option :: (WN)WN begin -fun widen_up where -"widen_up Bot x = x" | -"widen_up x Bot = x" | -"widen_up (Up x) (Up y) = Up(x \ y)" +fun widen_option where +"None \ x = x" | +"x \ None = x" | +"(Some x) \ (Some y) = Some(x \ y)" -fun narrow_up where -"narrow_up Bot x = Bot" | -"narrow_up x Bot = Bot" | -"narrow_up (Up x) (Up y) = Up(x \ y)" +fun narrow_option where +"None \ x = None" | +"x \ None = None" | +"(Some x) \ (Some y) = Some(x \ y)" instance proof case goal1 show ?case - by(induct x y rule: widen_up.induct) (simp_all add: widen1) + by(induct x y rule: widen_option.induct) (simp_all add: widen1) next case goal2 show ?case - by(induct x y rule: widen_up.induct) (simp_all add: widen2) + by(induct x y rule: widen_option.induct) (simp_all add: widen2) next case goal3 thus ?case - by(induct x y rule: narrow_up.induct) (simp_all add: narrow1) + by(induct x y rule: narrow_option.induct) (simp_all add: narrow1) next case goal4 thus ?case - by(induct x y rule: narrow_up.induct) (simp_all add: narrow2) + by(induct x y rule: narrow_option.induct) (simp_all add: narrow2) qed end @@ -129,7 +129,7 @@ "prefp f = while_option (\x. \ x \ f x) f" definition - pfp_WN :: "(('a::WN)up acom \ 'a up acom) \ com \ 'a up acom option" + pfp_WN :: "(('a::WN)option acom \ 'a option acom) \ com \ 'a option acom option" where "pfp_WN f c = (case lpfp\<^isub>c (\c. c \\<^sub>c f c) c of None \ None | Some c' \ prefp (\c. c \\<^sub>c f c) c')" @@ -183,13 +183,26 @@ locale Abs_Int2 = Abs_Int1_mono rep for rep :: "'a::{WN,L_top_bot} \ val set" begin -definition AI_WN :: "com \ 'a st up acom option" where +definition AI_WN :: "com \ 'a st option acom option" where "AI_WN = pfp_WN (step \)" -lemma AI_sound: "\ AI_WN c = Some c'; (c,s) \ t \ \ t <:up post c'" -unfolding AI_WN_def -by(metis step_sound[of "\" c' s t] strip_pfp_WN strip_step - pfp_WN_pfp mono_def mono_step[OF le_refl] in_rep_Top_up) +lemma AI_WN_sound: "AI_WN c = Some c' \ CS UNIV c \ \\<^isub>c c'" +proof(simp add: CS_def AI_WN_def) + assume 1: "pfp_WN (step \) c = Some c'" + from pfp_WN_pfp[OF allI[OF strip_step] mono_step 1] + have 2: "step \ c' \ c'" . + have 3: "strip (\\<^isub>c (step \ c')) = c" by(simp add: strip_pfp_WN[OF _ 1]) + have "lfp c (step_cs UNIV) \ \\<^isub>c (step \ c')" + proof(rule lfp_lowerbound[OF 3]) + show "step_cs UNIV (\\<^isub>c (step \ c')) \ \\<^isub>c (step \ c')" + proof(rule step_preserves_le[OF _ _ 3]) + show "UNIV \ \\<^isub>u \" by simp + show "\\<^isub>c (step \ c') \ \\<^isub>c c'" by(rule mono_rep_c[OF 2]) + qed + qed + from this 2 show "lfp c (step_cs UNIV) \ \\<^isub>c c'" + by (blast intro: mono_rep_c order_trans) +qed end @@ -198,12 +211,16 @@ defines AI_ivl' is AI_WN proof qed -value [code] "show_acom_opt (AI_ivl test3_ivl)" -value [code] "show_acom_opt (AI_ivl' test3_ivl)" + +subsubsection "Tests" definition "step_up_ivl n = ((\c. c \\<^sub>c step_ivl \ c)^^n)" definition "step_down_ivl n = ((\c. c \\<^sub>c step_ivl \ c)^^n)" +text{* For @{const test3_ivl}, @{const AI_ivl} needed as many iterations as +the loop took to execute. In contrast, @{const AI_ivl} converges in a +constant number of steps: *} + value [code] "show_acom (step_up_ivl 1 (\\<^sub>c test3_ivl))" value [code] "show_acom (step_up_ivl 2 (\\<^sub>c test3_ivl))" value [code] "show_acom (step_up_ivl 3 (\\<^sub>c test3_ivl))" @@ -213,6 +230,8 @@ value [code] "show_acom (step_down_ivl 2 (step_up_ivl 5 (\\<^sub>c test3_ivl)))" value [code] "show_acom (step_down_ivl 3 (step_up_ivl 5 (\\<^sub>c test3_ivl)))" +text{* Now all the analyses terminate: *} + value [code] "show_acom_opt (AI_ivl' test4_ivl)" value [code] "show_acom_opt (AI_ivl' test5_ivl)" value [code] "show_acom_opt (AI_ivl' test6_ivl)"