diff -r b716b16ab2ac -r e72e44cee6f2 src/HOL/IMP/Abs_Int1.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/Abs_Int1.thy Thu Apr 19 20:19:13 2012 +0200 @@ -0,0 +1,313 @@ +(* Author: Tobias Nipkow *) + +theory Abs_Int1 +imports Abs_State +begin + +lemma le_iff_le_annos_zip: "C1 \ C2 \ + (\ (a1,a2) \ set(zip (annos C1) (annos C2)). a1 \ a2) \ strip C1 = strip C2" +by(induct C1 C2 rule: le_acom.induct) (auto simp: size_annos_same2) + +lemma le_iff_le_annos: "C1 \ C2 \ + strip C1 = strip C2 \ (\ i annos C2 ! i)" +by(auto simp add: le_iff_le_annos_zip set_zip) (metis size_annos_same2) + + +lemma mono_fun_wt[simp]: "wt F X \ F \ G \ x : X \ fun F x \ fun G x" +by(simp add: mono_fun wt_st_def) + +lemma wt_bot[simp]: "wt (bot c) (vars c)" +by(simp add: wt_acom_def bot_def) + +lemma wt_acom_simps[simp]: "wt (SKIP {P}) X \ wt P X" + "wt (x ::= e {P}) X \ x : X \ vars e \ X \ wt P X" + "wt (C1;C2) X \ wt C1 X \ wt C2 X" + "wt (IF b THEN C1 ELSE C2 {P}) X \ + vars b \ X \ wt C1 X \ wt C2 X \ wt P X" + "wt ({I} WHILE b DO C {P}) X \ + wt I X \ vars b \ X \ wt C X \ wt P X" +by(auto simp add: wt_acom_def) + +lemma wt_post[simp]: "wt c X \ wt (post c) X" +by(induction c)(auto simp: wt_acom_def) + +lemma lpfp_inv: +assumes "lpfp f x0 = Some x" and "\x. P x \ P(f x)" and "P(bot x0)" +shows "P x" +using assms unfolding lpfp_def pfp_def +by (metis (lifting) while_option_rule) + + +subsection "Computable Abstract Interpretation" + +text{* Abstract interpretation over type @{text st} instead of +functions. *} + +context Gamma +begin + +fun aval' :: "aexp \ 'av st \ 'av" where +"aval' (N n) S = num' n" | +"aval' (V x) S = fun S x" | +"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" + +lemma aval'_sound: "s : \\<^isub>f S \ vars a \ dom S \ aval a s : \(aval' a S)" +by (induction a) (auto simp: gamma_num' gamma_plus' \_st_def) + +end + +text{* The for-clause (here and elsewhere) only serves the purpose of fixing +the name of the type parameter @{typ 'av} which would otherwise be renamed to +@{typ 'a}. *} + +locale Abs_Int = Gamma where \=\ for \ :: "'av::SL_top \ val set" +begin + +fun step' :: "'av st option \ 'av st option acom \ 'av st option acom" where +"step' S (SKIP {P}) = (SKIP {S})" | +"step' S (x ::= e {P}) = + x ::= e {case S of None \ None | Some S \ Some(update S x (aval' e S))}" | +"step' S (C1; C2) = step' S C1; step' (post C1) C2" | +"step' S (IF b THEN C1 ELSE C2 {P}) = + (IF b THEN step' S C1 ELSE step' S C2 {post C1 \ post C2})" | +"step' S ({Inv} WHILE b DO C {P}) = + {S \ post C} WHILE b DO step' Inv C {Inv}" + +definition AI :: "com \ 'av st option acom option" where +"AI c = lpfp (step' (top c)) c" + + +lemma strip_step'[simp]: "strip(step' S C) = strip C" +by(induct C arbitrary: S) (simp_all add: Let_def) + + +text{* Soundness: *} + +lemma in_gamma_update: + "\ s : \\<^isub>f S; i : \ a \ \ s(x := i) : \\<^isub>f(update S x a)" +by(simp add: \_st_def) + +theorem step_preserves_le: + "\ S \ \\<^isub>o S'; C \ \\<^isub>c C'; wt C' X; wt S' X \ \ step S C \ \\<^isub>c (step' S' C')" +proof(induction C arbitrary: C' S S') + case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) +next + case Assign thus ?case + by(fastforce simp: Assign_le map_acom_Assign wt_st_def + intro: aval'_sound in_gamma_update split: option.splits) +next + case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi) + by (metis le_post post_map_acom wt_post) +next + case (If b C1 C2 P) + then obtain C1' C2' P' where + "C' = IF b THEN C1' ELSE C2' {P'}" + "P \ \\<^isub>o P'" "C1 \ \\<^isub>c C1'" "C2 \ \\<^isub>c C2'" + by (fastforce simp: If_le map_acom_If) + moreover from this(1) `wt C' X` have wt: "wt C1' X" "wt C2' X" + by simp_all + moreover have "post C1 \ \\<^isub>o(post C1' \ post C2')" + by (metis (no_types) `C1 \ \\<^isub>c C1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom wt wt_post) + moreover have "post C2 \ \\<^isub>o(post C1' \ post C2')" + by (metis (no_types) `C2 \ \\<^isub>c C2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom wt wt_post) + ultimately show ?case using `S \ \\<^isub>o S'` `wt S' X` + by (simp add: If.IH subset_iff) +next + case (While I b C1 P) + then obtain C1' I' P' where + "C' = {I'} WHILE b DO C1' {P'}" + "I \ \\<^isub>o I'" "P \ \\<^isub>o P'" "C1 \ \\<^isub>c C1'" + by (fastforce simp: map_acom_While While_le) + moreover from this(1) `wt C' X` + have wt: "wt C1' X" "wt I' X" by simp_all + moreover note compat = `wt S' X` wt_post[OF wt(1)] + moreover have "S \ post C1 \ \\<^isub>o (S' \ post C1')" + using `S \ \\<^isub>o S'` le_post[OF `C1 \ \\<^isub>c C1'`, simplified] + by (metis (no_types) join_ge1[OF compat] join_ge2[OF compat] le_sup_iff mono_gamma_o order_trans) + ultimately show ?case by (simp add: While.IH subset_iff) +qed + +lemma wt_step'[simp]: + "\ wt C X; wt S X \ \ wt (step' S C) X" +proof(induction C arbitrary: S) + case Assign thus ?case + by(auto simp: wt_st_def update_def split: option.splits) +qed auto + +theorem AI_sound: "AI c = Some C \ CS c \ \\<^isub>c C" +proof(simp add: CS_def AI_def) + assume 1: "lpfp (step' (top c)) c = Some C" + have "wt C (vars c)" + by(rule lpfp_inv[where P = "%C. wt C (vars c)", OF 1 _ wt_bot]) + (erule wt_step'[OF _ wt_top]) + have 2: "step' (top c) C \ C" by(rule lpfpc_pfp[OF 1]) + have 3: "strip (\\<^isub>c (step' (top c) C)) = c" + by(simp add: strip_lpfp[OF _ 1]) + have "lfp (step UNIV) c \ \\<^isub>c (step' (top c) C)" + proof(rule lfp_lowerbound[simplified,OF 3]) + show "step UNIV (\\<^isub>c (step' (top c) C)) \ \\<^isub>c (step' (top c) C)" + proof(rule step_preserves_le[OF _ _ `wt C (vars c)` wt_top]) + show "UNIV \ \\<^isub>o (top c)" by simp + show "\\<^isub>c (step' (top c) C) \ \\<^isub>c C" by(rule mono_gamma_c[OF 2]) + qed + qed + from this 2 show "lfp (step UNIV) c \ \\<^isub>c C" + by (blast intro: mono_gamma_c order_trans) +qed + +end + + +subsubsection "Monotonicity" + +lemma le_join_disj: "wt y X \ wt (z::_::SL_top_wt) X \ x \ y \ x \ z \ x \ y \ z" +by (metis join_ge1 join_ge2 preord_class.le_trans) + +locale Abs_Int_mono = Abs_Int + +assumes mono_plus': "a1 \ b1 \ a2 \ b2 \ plus' a1 a2 \ plus' b1 b2" +begin + +lemma mono_aval': "S1 \ S2 \ wt S1 X \ vars e \ X \ aval' e S1 \ aval' e S2" +by(induction e) (auto simp: le_st_def mono_plus' wt_st_def) + +theorem mono_step': "wt S1 X \ wt S2 X \ wt C1 X \ wt C2 X \ + S1 \ S2 \ C1 \ C2 \ step' S1 C1 \ step' S2 C2" +apply(induction C1 C2 arbitrary: S1 S2 rule: le_acom.induct) +apply (auto simp: Let_def mono_aval' mono_post + le_join_disj le_join_disj[OF wt_post wt_post] + split: option.split) +done + +lemma mono_step'_top: "wt c (vars c0) \ wt c' (vars c0) \ c \ c' \ step' (top c0) c \ step' (top c0) c'" +by (metis wt_top mono_step' preord_class.le_refl) + +end + + +subsubsection "Termination" + +abbreviation sqless (infix "\" 50) where +"x \ y == x \ y \ \ y \ x" + +lemma pfp_termination: +fixes x0 :: "'a::preord" and m :: "'a \ nat" +assumes mono: "\x y. I x \ I y \ x \ y \ f x \ f y" +and m: "\x y. I x \ I y \ x \ y \ m x > m y" +and I: "\x y. I x \ I(f x)" and "I x0" and "x0 \ f x0" +shows "EX x. pfp f x0 = Some x" +proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. I x & x \ f x"]) + show "wf {(y,x). ((I x \ x \ f x) \ \ f x \ x) \ y = f x}" + by(rule wf_subset[OF wf_measure[of m]]) (auto simp: m I) +next + show "I x0 \ x0 \ f x0" using `I x0` `x0 \ f x0` by blast +next + fix x assume "I x \ x \ f x" thus "I(f x) \ f x \ f(f x)" + by (blast intro: I mono) +qed + +lemma lpfp_termination: +fixes f :: "'a::preord option acom \ 'a option acom" +and m :: "'a option acom \ nat" and I :: "'a option acom \ bool" +assumes "\x y. I x \ I y \ x \ y \ m x > m y" +and "\x y. I x \ I y \ x \ y \ f x \ f y" +and "\x y. I x \ I(f x)" and "I(bot c)" +and "\C. strip (f C) = strip C" +shows "\c'. lpfp f c = Some c'" +unfolding lpfp_def +by(fastforce intro: pfp_termination[where I=I and m=m] assms bot_least + simp: assms(5)) + + +locale Abs_Int_measure = + Abs_Int_mono where \=\ for \ :: "'av::SL_top \ val set" + +fixes m :: "'av \ nat" +fixes h :: "nat" +assumes m1: "x \ y \ m x \ m y" +assumes m2: "x \ y \ m x > m y" +assumes h: "m x \ h" +begin + +definition "m_st S = (\ x \ dom S. m(fun S x))" + +lemma m_st1: "S1 \ S2 \ m_st S1 \ m_st S2" +proof(auto simp add: le_st_def m_st_def) + assume "\x\dom S2. fun S1 x \ fun S2 x" + hence "\x\dom S2. m(fun S1 x) \ m(fun S2 x)" by (metis m1) + thus "(\x\dom S2. m (fun S2 x)) \ (\x\dom S2. m (fun S1 x))" + by (metis setsum_mono) +qed + +lemma m_st2: "finite(dom S1) \ S1 \ S2 \ m_st S1 > m_st S2" +proof(auto simp add: le_st_def m_st_def) + assume "finite(dom S2)" and 0: "\x\dom S2. fun S1 x \ fun S2 x" + hence 1: "\x\dom S2. m(fun S1 x) \ m(fun S2 x)" by (metis m1) + fix x assume "x \ dom S2" "\ fun S2 x \ fun S1 x" + hence 2: "\x\dom S2. m(fun S1 x) > m(fun S2 x)" using 0 m2 by blast + from setsum_strict_mono_ex1[OF `finite(dom S2)` 1 2] + show "(\x\dom S2. m (fun S2 x)) < (\x\dom S2. m (fun S1 x))" . +qed + + +definition m_o :: "nat \ 'av st option \ nat" where +"m_o d opt = (case opt of None \ h*d+1 | Some S \ m_st S)" + +definition m_c :: "'av st option acom \ nat" where +"m_c c = (\i finite X \ m_st x \ h * card X" +by(simp add: wt_st_def m_st_def) + (metis nat_mult_commute of_nat_id setsum_bounded[OF h]) + +lemma m_o1: "finite X \ wt o1 X \ wt o2 X \ + o1 \ o2 \ m_o (card X) o1 \ m_o (card X) o2" +proof(induction o1 o2 rule: le_option.induct) + case 1 thus ?case by (simp add: m_o_def)(metis m_st1) +next + case 2 thus ?case + by(simp add: wt_option_def m_o_def le_SucI m_st_h split: option.splits) +next + case 3 thus ?case by simp +qed + +lemma m_o2: "finite X \ wt o1 X \ wt o2 X \ + o1 \ o2 \ m_o (card X) o1 > m_o (card X) o2" +proof(induction o1 o2 rule: le_option.induct) + case 1 thus ?case by (simp add: m_o_def wt_st_def m_st2) +next + case 2 thus ?case + by(auto simp add: m_o_def le_imp_less_Suc m_st_h) +next + case 3 thus ?case by simp +qed + +lemma m_c2: "wt c1 (vars(strip c1)) \ wt c2 (vars(strip c2)) \ + c1 \ c2 \ m_c c1 > m_c c2" +proof(auto simp add: le_iff_le_annos m_c_def size_annos_same[of c1 c2] wt_acom_def) + let ?X = "vars(strip c2)" + let ?n = "card ?X" + assume V1: "\a\set(annos c1). wt a ?X" + and V2: "\a\set(annos c2). wt a ?X" + and strip_eq: "strip c1 = strip c2" + and 0: "\i annos c2 ! i" + hence 1: "\i m_o ?n (annos c2 ! i)" + by (auto simp: all_set_conv_all_nth) + (metis finite_cvars m_o1 size_annos_same2) + fix i assume "i < size(annos c2)" "\ annos c2 ! i \ annos c1 ! i" + hence "m_o ?n (annos c1 ! i) > m_o ?n (annos c2 ! i)" (is "?P i") + by(metis m_o2[OF finite_cvars] V1 V2 strip_eq nth_mem size_annos_same 0) + hence 2: "\i < size(annos c2). ?P i" using `i < size(annos c2)` by blast + show "(\iiC. AI c = Some C" +unfolding AI_def +apply(rule lpfp_termination[where I = "%C. strip C = c \ wt C (vars c)" + and m="m_c"]) +apply(simp_all add: m_c2 mono_step'_top) +done + +end + +end