diff -r 323414474c1f -r 73fb17ed2e08 src/HOL/IMP/Abs_Int1.thy --- a/src/HOL/IMP/Abs_Int1.thy Sun Sep 16 10:33:25 2012 +0200 +++ b/src/HOL/IMP/Abs_Int1.thy Sun Sep 16 11:50:03 2012 +0200 @@ -13,26 +13,26 @@ 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 mono_fun_L[simp]: "F \ L X \ F \ G \ x : X \ fun F x \ fun G x" +by(simp add: mono_fun L_st_def) -lemma wt_bot[simp]: "wt (bot c) (vars c)" -by(simp add: wt_acom_def bot_def) +lemma bot_in_L[simp]: "bot c \ L(vars c)" +by(simp add: L_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 {P1} C1 ELSE {P2} C2 {Q}) X \ - vars b \ X \ wt C1 X \ wt C2 X \ wt P1 X \ wt P2 X \ wt Q X" - "wt ({I} WHILE b DO {P} C {Q}) X \ - wt I X \ vars b \ X \ wt P X \ wt C X \ wt Q X" -by(auto simp add: wt_acom_def) +lemma L_acom_simps[simp]: "SKIP {P} \ L X \ P \ L X" + "(x ::= e {P}) \ L X \ x : X \ vars e \ X \ P \ L X" + "(C1;C2) \ L X \ C1 \ L X \ C2 \ L X" + "(IF b THEN {P1} C1 ELSE {P2} C2 {Q}) \ L X \ + vars b \ X \ C1 \ L X \ C2 \ L X \ P1 \ L X \ P2 \ L X \ Q \ L X" + "({I} WHILE b DO {P} C {Q}) \ L X \ + I \ L X \ vars b \ X \ P \ L X \ C \ L X \ Q \ L X" +by(auto simp add: L_acom_def) lemma post_in_annos: "post C : set(annos C)" by(induction C) auto -lemma wt_post[simp]: "wt C X \ wt (post C) X" -by(simp add: wt_acom_def post_in_annos) +lemma post_in_L[simp]: "C \ L X \ post C \ L X" +by(simp add: L_acom_def post_in_annos) lemma lpfp_inv: assumes "lpfp f x0 = Some x" and "\x. P x \ P(f x)" and "P(bot x0)" @@ -63,7 +63,7 @@ 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" +locale Abs_Int = Gamma where \=\ for \ :: "'av::semilattice \ val set" begin fun step' :: "'av st option \ 'av st option acom \ 'av st option acom" where @@ -91,29 +91,29 @@ 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')" + "\ S \ \\<^isub>o S'; C \ \\<^isub>c C'; C' \ L X; S' \ L 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 + by(fastforce simp: Assign_le map_acom_Assign L_st_def intro: aval'_sound in_gamma_update split: option.splits) next case Seq thus ?case apply (auto simp: Seq_le map_acom_Seq) - by (metis le_post post_map_acom wt_post) + by (metis le_post post_map_acom post_in_L) next case (If b P1 C1 P2 C2 Q) then obtain P1' P2' C1' C2' Q' where "C' = IF b THEN {P1'} C1' ELSE {P2'} C2' {Q'}" "P1 \ \\<^isub>o P1'" "P2 \ \\<^isub>o P2'" "Q \ \\<^isub>o Q'" "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" "wt P1' X" "wt P2' X" - by simp_all + moreover from this(1) `C' \ L X` + have L: "C1' \ L X" "C2' \ L X" "P1' \ L X" "P2' \ L 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) + by (metis (no_types) `C1 \ \\<^isub>c C1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom L post_in_L) 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 (metis (no_types) `C2 \ \\<^isub>c C2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom L post_in_L) + ultimately show ?case using `S \ \\<^isub>o S'` `S' \ L X` by (simp add: If.IH subset_iff) next case (While I b P1 C1 Q) @@ -121,35 +121,35 @@ "C' = {I'} WHILE b DO {P1'} C1' {Q'}" "I \ \\<^isub>o I'" "P1 \ \\<^isub>o P1'" "C1 \ \\<^isub>c C1'" "Q \ \\<^isub>o Q'" by (fastforce simp: map_acom_While While_le) - moreover from this(1) `wt C' X` - have wt: "wt C1' X" "wt I' X" "wt P1' X" by simp_all - moreover note compat = `wt S' X` wt_post[OF wt(1)] + moreover from this(1) `C' \ L X` + have L: "C1' \ L X" "I' \ L X" "P1' \ L X" by simp_all + moreover note compat = `S' \ L X` post_in_L[OF L(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" +lemma step'_in_L[simp]: + "\ C \ L X; S \ L X \ \ (step' S C) \ L X" proof(induction C arbitrary: S) case Assign thus ?case - by(auto simp: wt_st_def update_def split: option.splits) + by(auto simp: L_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 "C \ L(vars c)" + by(rule lpfp_inv[where P = "%C. C \ L(vars c)", OF 1 _ bot_in_L]) + (erule step'_in_L[OF _ top_in_L]) 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 c (step UNIV) \ \\<^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]) + proof(rule step_preserves_le[OF _ _ `C \ L(vars c)` top_in_L]) 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 @@ -163,26 +163,29 @@ subsubsection "Monotonicity" -lemma le_join_disj: "wt y X \ wt (z::_::SL_top_wt) X \ x \ y \ x \ z \ x \ y \ z" +lemma le_join_disj: "y \ L X \ (z::_::semilatticeL) \ L 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) +lemma mono_aval': + "S1 \ S2 \ S1 \ L X \ vars e \ X \ aval' e S1 \ aval' e S2" +by(induction e) (auto simp: le_st_def mono_plus' L_st_def) -theorem mono_step': "wt S1 X \ wt S2 X \ wt C1 X \ wt C2 X \ +theorem mono_step': "S1 \ L X \ S2 \ L X \ C1 \ L X \ C2 \ L 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] + le_join_disj le_join_disj[OF post_in_L post_in_L] 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) +lemma mono_step'_top: "C \ L(vars c) \ C' \ L(vars c) \ + C \ C' \ step' (top c) C \ step' (top c) C'" +by (metis top_in_L mono_step' preord_class.le_refl) end @@ -222,7 +225,7 @@ locale Abs_Int_measure = - Abs_Int_mono where \=\ for \ :: "'av::SL_top \ val set" + + Abs_Int_mono where \=\ for \ :: "'av::semilattice \ val set" + fixes m :: "'av \ nat" fixes h :: "nat" assumes m1: "x \ y \ m x \ m y" @@ -255,27 +258,27 @@ "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 = (\ii finite X \ m_st x \ h * card X" -by(simp add: wt_st_def m_st_def) +lemma m_st_h: "x \ L X \ finite X \ m_st x \ h * card X" +by(simp add: L_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 \ +lemma m_o1: "finite X \ o1 \ L X \ o2 \ L 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) + by(simp add: L_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 \ +lemma m_o2: "finite X \ o1 \ L X \ o2 \ L 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) + case 1 thus ?case by (simp add: m_o_def L_st_def m_st2) next case 2 thus ?case by(auto simp add: m_o_def le_imp_less_Suc m_st_h) @@ -283,30 +286,30 @@ 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)" +lemma m_c2: "C1 \ L(vars(strip C1)) \ C2 \ L(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] L_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)" + assume V1: "\a\set(annos C1). a \ L ?X" + and V2: "\a\set(annos C2). a \ L ?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 "(\ii 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 nth_mem size_annos_same[OF strip_eq] 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)" +apply(rule lpfp_termination[where I = "%C. strip C = c \ C \ L(vars c)" and m="m_c"]) apply(simp_all add: m_c2 mono_step'_top) done