# HG changeset patch # User nipkow # Date 1347789003 -7200 # Node ID 73fb17ed2e08c3a63ae814611431c402d84209c7 # Parent 323414474c1f68a87981fd53588747c7f4a105ad converted wt into a set, tuned names diff -r 323414474c1f -r 73fb17ed2e08 src/HOL/IMP/Abs_Int0.thy --- a/src/HOL/IMP/Abs_Int0.thy Sun Sep 16 10:33:25 2012 +0200 +++ b/src/HOL/IMP/Abs_Int0.thy Sun Sep 16 11:50:03 2012 +0200 @@ -27,7 +27,7 @@ class join = preord + fixes join :: "'a \ 'a \ 'a" (infixl "\" 65) -class SL_top = join + +class semilattice = join + fixes Top :: "'a" ("\") assumes join_ge1 [simp]: "x \ x \ y" and join_ge2 [simp]: "y \ x \ y" @@ -56,7 +56,7 @@ end -instantiation "fun" :: (type, SL_top) SL_top +instantiation "fun" :: (type, semilattice) semilattice begin definition "f \ g = (\x. f x \ g x)" @@ -154,7 +154,7 @@ end -instantiation option :: (SL_top)SL_top +instantiation option :: (semilattice)semilattice begin definition "\ = Some \" @@ -171,19 +171,19 @@ end -class Bot = preord + -fixes Bot :: "'a" ("\") -assumes Bot[simp]: "\ \ x" +class bot = preord + +fixes bot :: "'a" ("\") +assumes bot[simp]: "\ \ x" -instantiation option :: (preord)Bot +instantiation option :: (preord)bot begin -definition Bot_option :: "'a option" where +definition bot_option :: "'a option" where "\ = None" instance proof - case goal1 thus ?case by(auto simp: Bot_option_def) + case goal1 thus ?case by(auto simp: bot_option_def) qed end @@ -250,7 +250,7 @@ text{* The interface for abstract values: *} locale Val_abs = -fixes \ :: "'av::SL_top \ val set" +fixes \ :: "'av::semilattice \ val set" assumes mono_gamma: "a \ b \ \ a \ \ b" and gamma_Top[simp]: "\ \ = UNIV" fixes num' :: "val \ 'av" @@ -261,7 +261,7 @@ type_synonym 'av st = "(vname \ 'av)" -locale Abs_Int_Fun = Val_abs \ for \ :: "'av::SL_top \ val set" +locale Abs_Int_Fun = Val_abs \ for \ :: "'av::semilattice \ val set" begin fun aval' :: "aexp \ 'av st \ 'av" where 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 diff -r 323414474c1f -r 73fb17ed2e08 src/HOL/IMP/Abs_Int1_const.thy --- a/src/HOL/IMP/Abs_Int1_const.thy Sun Sep 16 10:33:25 2012 +0200 +++ b/src/HOL/IMP/Abs_Int1_const.thy Sun Sep 16 11:50:03 2012 +0200 @@ -20,7 +20,7 @@ (case (a1,a2) of (Const m, Const n) \ Const(m+n) | _ \ Any)" by(auto split: prod.split const.split) -instantiation const :: SL_top +instantiation const :: semilattice begin fun le_const where @@ -74,6 +74,10 @@ subsubsection "Tests" +(* FIXME dirty trick to get around some problem with the code generator *) +lemma [code]: "L X = (L X :: 'av::semilattice st set)" +by(rule refl) + definition "steps c i = (step_const(top c) ^^ i) (bot c)" value "show_acom (steps test1_const 0)" diff -r 323414474c1f -r 73fb17ed2e08 src/HOL/IMP/Abs_Int1_parity.thy --- a/src/HOL/IMP/Abs_Int1_parity.thy Sun Sep 16 10:33:25 2012 +0200 +++ b/src/HOL/IMP/Abs_Int1_parity.thy Sun Sep 16 11:50:03 2012 +0200 @@ -33,9 +33,9 @@ end -text{* Instantiation of class @{class SL_top} with type @{typ parity}: *} +text{* Instantiation of class @{class semilattice} with type @{typ parity}: *} -instantiation parity :: SL_top +instantiation parity :: semilattice begin definition join_parity where @@ -118,11 +118,14 @@ subsubsection "Tests" +(* FIXME dirty trick to get around some problem with the code generator *) +lemma [code]: "L X = (L X :: 'av::semilattice st set)" +by(rule refl) + definition "test1_parity = ''x'' ::= N 1; WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 2)" - -value "show_acom_opt (AI_parity test1_parity)" +value [code] "show_acom_opt (AI_parity test1_parity)" definition "test2_parity = ''x'' ::= N 1; diff -r 323414474c1f -r 73fb17ed2e08 src/HOL/IMP/Abs_Int2.thy --- a/src/HOL/IMP/Abs_Int2.thy Sun Sep 16 10:33:25 2012 +0200 +++ b/src/HOL/IMP/Abs_Int2.thy Sun Sep 16 11:50:03 2012 +0200 @@ -21,7 +21,7 @@ subsection "Backward Analysis of Expressions" -class L_top_bot = SL_top + Bot + +class lattice = semilattice + bot + fixes meet :: "'a \ 'a \ 'a" (infixl "\" 65) assumes meet_le1 [simp]: "x \ y \ x" and meet_le2 [simp]: "x \ y \ y" @@ -34,10 +34,10 @@ end locale Val_abs1_gamma = - Gamma where \ = \ for \ :: "'av::L_top_bot \ val set" + + Gamma where \ = \ for \ :: "'av::lattice \ val set" + assumes inter_gamma_subset_gamma_meet: "\ a1 \ \ a2 \ \(a1 \ a2)" -and gamma_Bot[simp]: "\ \ = {}" +and gamma_bot[simp]: "\ \ = {}" begin lemma in_gamma_meet: "x : \ a1 \ x : \ a2 \ x : \(a1 \ a2)" @@ -51,7 +51,7 @@ locale Val_abs1 = Val_abs1_gamma where \ = \ - for \ :: "'av::L_top_bot \ val set" + + for \ :: "'av::lattice \ val set" + fixes test_num' :: "val \ 'av \ bool" and filter_plus' :: "'av \ 'av \ 'av \ 'av * 'av" and filter_less' :: "bool \ 'av \ 'av \ 'av * 'av" @@ -63,19 +63,19 @@ locale Abs_Int1 = - Val_abs1 where \ = \ for \ :: "'av::L_top_bot \ val set" + Val_abs1 where \ = \ for \ :: "'av::lattice \ val set" begin lemma in_gamma_join_UpI: - "wt S1 X \ wt S2 X \ s : \\<^isub>o S1 \ s : \\<^isub>o S2 \ s : \\<^isub>o(S1 \ S2)" -by (metis (hide_lams, no_types) SL_top_wt_class.join_ge1 SL_top_wt_class.join_ge2 mono_gamma_o subsetD) + "S1 \ L X \ S2 \ L X \ s : \\<^isub>o S1 \ s : \\<^isub>o S2 \ s : \\<^isub>o(S1 \ S2)" +by (metis (hide_lams, no_types) semilatticeL_class.join_ge1 semilatticeL_class.join_ge2 mono_gamma_o subsetD) fun aval'' :: "aexp \ 'av st option \ 'av" where "aval'' e None = \" | "aval'' e (Some sa) = aval' e sa" -lemma aval''_sound: "s : \\<^isub>o S \ wt S X \ vars a \ X \ aval a s : \(aval'' a S)" -by(simp add: wt_option_def wt_st_def aval'_sound split: option.splits) +lemma aval''_sound: "s : \\<^isub>o S \ S \ L X \ vars a \ X \ aval a s : \(aval'' a S)" +by(simp add: L_option_def L_st_def aval'_sound split: option.splits) subsubsection "Backward analysis" @@ -108,12 +108,12 @@ (let (res1,res2) = filter_less' res (aval'' e1 S) (aval'' e2 S) in afilter e1 res1 (afilter e2 res2 S))" -lemma wt_afilter: "wt S X \ vars e \ X ==> wt (afilter e a S) X" +lemma afilter_in_L: "S \ L X \ vars e \ X \ afilter e a S \ L X" by(induction e arbitrary: a S) - (auto simp: Let_def update_def wt_st_def + (auto simp: Let_def update_def L_st_def split: option.splits prod.split) -lemma afilter_sound: "wt S X \ vars e \ X \ +lemma afilter_sound: "S \ L X \ vars e \ X \ s : \\<^isub>o S \ aval e s : \ a \ s : \\<^isub>o (afilter e a S)" proof(induction e arbitrary: a S) case N thus ?case by simp (metis test_num') @@ -122,22 +122,21 @@ obtain S' where "S = Some S'" and "s : \\<^isub>f S'" using `s : \\<^isub>o S` by(auto simp: in_gamma_option_iff) moreover hence "s x : \ (fun S' x)" - using V(1,2) by(simp add: \_st_def wt_st_def) + using V(1,2) by(simp add: \_st_def L_st_def) moreover have "s x : \ a" using V by simp ultimately show ?case using V(3) by(simp add: Let_def \_st_def) - (metis mono_gamma emptyE in_gamma_meet gamma_Bot subset_empty) + (metis mono_gamma emptyE in_gamma_meet gamma_bot subset_empty) next case (Plus e1 e2) thus ?case using filter_plus'[OF _ aval''_sound[OF Plus.prems(3)] aval''_sound[OF Plus.prems(3)]] - by (auto simp: wt_afilter split: prod.split) + by (auto simp: afilter_in_L split: prod.split) qed -lemma wt_bfilter: "wt S X \ vars b \ X \ wt (bfilter b bv S) X" -by(induction b arbitrary: bv S) - (auto simp: wt_afilter split: prod.split) +lemma bfilter_in_L: "S \ L X \ vars b \ X \ bfilter b bv S \ L X" +by(induction b arbitrary: bv S)(auto simp: afilter_in_L split: prod.split) -lemma bfilter_sound: "wt S X \ vars b \ X \ +lemma bfilter_sound: "S \ L X \ vars b \ X \ s : \\<^isub>o S \ bv = bval b s \ s : \\<^isub>o(bfilter b bv S)" proof(induction b arbitrary: S bv) case Bc thus ?case by simp @@ -145,11 +144,11 @@ case (Not b) thus ?case by simp next case (And b1 b2) thus ?case - by simp (metis And(1) And(2) wt_bfilter in_gamma_join_UpI) + by simp (metis And(1) And(2) bfilter_in_L in_gamma_join_UpI) next case (Less e1 e2) thus ?case by(auto split: prod.split) - (metis (lifting) wt_afilter afilter_sound aval''_sound filter_less') + (metis (lifting) afilter_in_L afilter_sound aval''_sound filter_less') qed @@ -181,67 +180,68 @@ 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_option_def wt_st_def + by (fastforce simp: Assign_le map_acom_Assign L_option_def L_st_def intro: aval'_sound in_gamma_update split: option.splits del:subsetD) 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' C1' P2' C2' Q' where "C' = IF b THEN {P1'} C1' ELSE {P2'} C2' {Q'}" "P1 \ \\<^isub>o P1'" "P2 \ \\<^isub>o P2'" "C1 \ \\<^isub>c C1'" "C2 \ \\<^isub>c C2'" "Q \ \\<^isub>o Q'" 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" - and "vars b \ 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" and "vars b \ 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_post wt) + by (metis (no_types) `C1 \ \\<^isub>c C1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom post_in_L 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_post wt) - moreover note vars = `wt S' X` `vars b \ X` + by (metis (no_types) `C2 \ \\<^isub>c C2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom post_in_L L) + moreover note vars = `S' \ L X` `vars b \ X` ultimately show ?case using `S \ \\<^isub>o S'` - by (simp add: If.IH subset_iff bfilter_sound[OF vars] wt_bfilter[OF vars]) + by (simp add: If.IH subset_iff bfilter_sound[OF vars] bfilter_in_L[OF vars]) next case (While I b P C1 Q) then obtain C1' I' P' Q' where "C' = {I'} WHILE b DO {P'} C1' {Q'}" "I \ \\<^isub>o I'" "P \ \\<^isub>o P'" "Q \ \\<^isub>o Q'" "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" "wt P' X" and "vars b \ X" by simp_all - moreover note compat = `wt S' X` wt_post[OF wt(1)] - moreover note vars = `wt I' X` `vars b \ X` + moreover from this(1) `C' \ L X` + have L: "C1' \ L X" "I' \ L X" "P' \ L X" and "vars b \ X" by simp_all + moreover note compat = `S' \ L X` post_in_L[OF L(1)] + moreover note vars = `I' \ L X` `vars b \ X` 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 bfilter_sound[OF vars] wt_bfilter[OF vars]) + by (simp add: While.IH subset_iff bfilter_sound[OF vars] bfilter_in_L[OF vars]) 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(simp add: wt_option_def wt_st_def update_def split: option.splits) -qed (auto simp add: wt_bfilter) + case Assign thus ?case by(simp add: L_option_def L_st_def update_def split: option.splits) +qed (auto simp add: bfilter_in_L) 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 @@ -264,48 +264,49 @@ 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) + "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) lemma mono_aval'': - "S1 \ S2 \ wt S1 X \ vars e \ X \ aval'' e S1 \ aval'' e S2" + "S1 \ S2 \ S1 \ L X \ vars e \ X \ aval'' e S1 \ aval'' e S2" apply(cases S1) apply simp apply(cases S2) apply simp by (simp add: mono_aval') -lemma mono_afilter: "wt S1 X \ wt S2 X \ vars e \ X \ +lemma mono_afilter: "S1 \ L X \ S2 \ L X \ vars e \ X \ r1 \ r2 \ S1 \ S2 \ afilter e r1 S1 \ afilter e r2 S2" apply(induction e arbitrary: r1 r2 S1 S2) apply(auto simp: test_num' Let_def mono_meet split: option.splits prod.splits) apply (metis mono_gamma subsetD) -apply(drule (2) mono_fun_wt) +apply(drule (2) mono_fun_L) apply (metis mono_meet le_trans) -apply(metis mono_aval'' mono_filter_plus'[simplified le_prod_def] fst_conv snd_conv wt_afilter) +apply(metis mono_aval'' mono_filter_plus'[simplified le_prod_def] fst_conv snd_conv afilter_in_L) done -lemma mono_bfilter: "wt S1 X \ wt S2 X \ vars b \ X \ +lemma mono_bfilter: "S1 \ L X \ S2 \ L X \ vars b \ X \ S1 \ S2 \ bfilter b bv S1 \ bfilter b bv S2" apply(induction b arbitrary: bv S1 S2) apply(simp) apply(simp) apply simp -apply(metis join_least le_trans[OF _ join_ge1] le_trans[OF _ join_ge2] wt_bfilter) +apply(metis join_least le_trans[OF _ join_ge1] le_trans[OF _ join_ge2] bfilter_in_L) apply (simp split: prod.splits) -apply(metis mono_aval'' mono_afilter mono_filter_less'[simplified le_prod_def] fst_conv snd_conv wt_afilter) +apply(metis mono_aval'' mono_afilter mono_filter_less'[simplified le_prod_def] fst_conv snd_conv afilter_in_L) done -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_bfilter mono_aval' mono_post - le_join_disj le_join_disj[OF wt_post wt_post] wt_bfilter + le_join_disj le_join_disj[OF post_in_L post_in_L] bfilter_in_L split: option.split) done -lemma mono_step'_top: "wt C1 (vars c) \ wt C2 (vars c) \ C1 \ C2 \ step' (top c) C1 \ step' (top c) C2" -by (metis wt_top mono_step' preord_class.le_refl) +lemma mono_step'_top: "C1 \ L(vars c) \ C2 \ L(vars c) \ + C1 \ C2 \ step' (top c) C1 \ step' (top c) C2" +by (metis top_in_L mono_step' preord_class.le_refl) end diff -r 323414474c1f -r 73fb17ed2e08 src/HOL/IMP/Abs_Int2_ivl.thy --- a/src/HOL/IMP/Abs_Int2_ivl.thy Sun Sep 16 10:33:25 2012 +0200 +++ b/src/HOL/IMP/Abs_Int2_ivl.thy Sun Sep 16 11:50:03 2012 +0200 @@ -58,7 +58,7 @@ definition "plus_ivl i1 i2 = (if is_empty i1 | is_empty i2 then empty else case (i1,i2) of (I l1 h1, I l2 h2) \ I (l1+l2) (h1+h2))" -instantiation ivl :: SL_top +instantiation ivl :: semilattice begin definition le_option :: "bool \ int option \ int option \ bool" where @@ -111,7 +111,7 @@ end -instantiation ivl :: L_top_bot +instantiation ivl :: lattice begin definition "i1 \ i2 = (if is_empty i1 \ is_empty i2 then empty else @@ -131,7 +131,7 @@ case goal4 thus ?case by (cases x, cases y, cases z, auto simp add: le_ivl_def meet_ivl_def empty_def le_option_def max_option_def min_option_def split: option.splits if_splits) next - case goal1 show ?case by(cases x, simp add: Bot_ivl_def empty_def le_ivl_def) + case goal1 show ?case by(cases x, simp add: bot_ivl_def empty_def le_ivl_def) qed end @@ -186,7 +186,7 @@ case goal1 thus ?case by(auto simp add: \_ivl_def meet_ivl_def empty_def min_option_def max_option_def split: ivl.split option.split) next - case goal2 show ?case by(auto simp add: Bot_ivl_def \_ivl_def empty_def) + case goal2 show ?case by(auto simp add: bot_ivl_def \_ivl_def empty_def) qed lemma mono_minus_ivl: @@ -248,6 +248,10 @@ subsubsection "Tests" +(* FIXME dirty trick to get around some problem with the code generator *) +lemma [code]: "L X = (L X :: 'av::semilattice st set)" +by(rule refl) + value "show_acom_opt (AI_ivl test1_ivl)" text{* Better than @{text AI_const}: *} diff -r 323414474c1f -r 73fb17ed2e08 src/HOL/IMP/Abs_Int3.thy --- a/src/HOL/IMP/Abs_Int3.thy Sun Sep 16 10:33:25 2012 +0200 +++ b/src/HOL/IMP/Abs_Int3.thy Sun Sep 16 11:50:03 2012 +0200 @@ -6,51 +6,53 @@ subsubsection "Welltypedness" -class Wt = -fixes Wt :: "'a \ com \ bool" +class Lc = +fixes Lc :: "com \ 'a set" -instantiation st :: (type)Wt +instantiation st :: (type)Lc begin -definition Wt_st :: "'a st \ com \ bool" where -"Wt_st S c = wt S (vars c)" +definition Lc_st :: "com \ 'a st set" where +"Lc_st c = L (vars c)" instance .. end -instantiation acom :: (Wt)Wt +instantiation acom :: (Lc)Lc begin -definition Wt_acom :: "'a acom \ com \ bool" where -"Wt C c = (strip C = c \ (\a\set(annos C). Wt a c))" +definition Lc_acom :: "com \ 'a acom set" where +"Lc c = {C. strip C = c \ (\a\set(annos C). a \ Lc c)}" instance .. end -instantiation option :: (Wt)Wt +instantiation option :: (Lc)Lc begin -fun Wt_option :: "'a option \ com \ bool" where -"Wt None c = True" | -"Wt (Some x) c = Wt x c" +definition Lc_option :: "com \ 'a option set" where +"Lc c = {None} \ Some ` Lc c" + +lemma Lc_option_simps[simp]: "None \ Lc c" "(Some x \ Lc c) = (x \ Lc c)" +by(auto simp: Lc_option_def) instance .. end -lemma Wt_option_iff_wt[simp]: fixes a :: "_ st option" -shows "Wt a c = wt a (vars c)" -by(auto simp add: wt_option_def Wt_st_def split: option.splits) +lemma Lc_option_iff_wt[simp]: fixes a :: "_ st option" +shows "(a \ Lc c) = (a \ L (vars c))" +by(auto simp add: L_option_def Lc_st_def split: option.splits) context Abs_Int1 begin -lemma Wt_step': "Wt C c \ Wt S c \ Wt (step' S C) c" -apply(auto simp add: Wt_acom_def) -by (metis wt_acom_def wt_step' order_refl) +lemma step'_in_Lc: "C \ Lc c \ S \ Lc c \ step' S C \ Lc c" +apply(auto simp add: Lc_acom_def) +by(metis step'_in_L[simplified L_acom_def mem_Collect_eq] order_refl) end @@ -69,13 +71,13 @@ assumes narrow1: "y \ x \ y \ x \ y" assumes narrow2: "y \ x \ x \ y \ x" -class WN_Wt = widen + narrow + preord + Wt + -assumes widen1: "Wt x c \ Wt y c \ x \ x \ y" -assumes widen2: "Wt x c \ Wt y c \ y \ x \ y" +class WN_Lc = widen + narrow + preord + Lc + +assumes widen1: "x \ Lc c \ y \ Lc c \ x \ x \ y" +assumes widen2: "x \ Lc c \ y \ Lc c \ y \ x \ y" assumes narrow1: "y \ x \ y \ x \ y" assumes narrow2: "y \ x \ x \ y \ x" -assumes Wt_widen[simp]: "Wt x c \ Wt y c \ Wt (x \ y) c" -assumes Wt_narrow[simp]: "Wt x c \ Wt y c \ Wt (x \ y) c" +assumes Lc_widen[simp]: "x \ Lc c \ y \ Lc c \ x \ y \ Lc c" +assumes Lc_narrow[simp]: "x \ Lc c \ y \ Lc c \ x \ y \ Lc c" instantiation ivl :: WN @@ -101,7 +103,7 @@ end -instantiation st :: (WN)WN_Wt +instantiation st :: (WN)WN_Lc begin definition "widen_st F1 F2 = FunDom (\x. fun F1 x \ fun F2 x) (dom F1)" @@ -114,7 +116,7 @@ by(simp add: widen_st_def le_st_def WN_class.widen1) next case goal2 thus ?case - by(simp add: widen_st_def le_st_def WN_class.widen2 Wt_st_def wt_st_def) + by(simp add: widen_st_def le_st_def WN_class.widen2 Lc_st_def L_st_def) next case goal3 thus ?case by(auto simp: narrow_st_def le_st_def WN_class.narrow1) @@ -122,15 +124,15 @@ case goal4 thus ?case by(auto simp: narrow_st_def le_st_def WN_class.narrow2) next - case goal5 thus ?case by(auto simp: widen_st_def Wt_st_def wt_st_def) + case goal5 thus ?case by(auto simp: widen_st_def Lc_st_def L_st_def) next - case goal6 thus ?case by(auto simp: narrow_st_def Wt_st_def wt_st_def) + case goal6 thus ?case by(auto simp: narrow_st_def Lc_st_def L_st_def) qed end -instantiation option :: (WN_Wt)WN_Wt +instantiation option :: (WN_Lc)WN_Lc begin fun widen_option where @@ -158,10 +160,10 @@ by(induct x y rule: narrow_option.induct) (simp_all add: narrow2) next case goal5 thus ?case - by(induction x y rule: widen_option.induct)(auto simp: Wt_st_def) + by(induction x y rule: widen_option.induct)(auto simp: Lc_st_def) next case goal6 thus ?case - by(induction x y rule: narrow_option.induct)(auto simp: Wt_st_def) + by(induction x y rule: narrow_option.induct)(auto simp: Lc_st_def) qed end @@ -176,7 +178,7 @@ "map2_acom f ({a1} WHILE b DO {p} C {a2}) ({a3} WHILE b' DO {p'} C' {a4}) = ({f a1 a3} WHILE b DO {f p p'} map2_acom f C C' {f a2 a4})" -instantiation acom :: (WN_Wt)WN_Wt +instantiation acom :: (WN_Lc)WN_Lc begin definition "widen_acom = map2_acom (op \)" @@ -184,16 +186,16 @@ definition "narrow_acom = map2_acom (op \)" lemma widen_acom1: - "\\a\set(annos x). Wt a c; \a\set (annos y). Wt a c; strip x = strip y\ + "\\a\set(annos x). a \ Lc c; \a\set (annos y). a \ Lc c; strip x = strip y\ \ x \ x \ y" by(induct x y rule: le_acom.induct) - (auto simp: widen_acom_def widen1 Wt_acom_def) + (auto simp: widen_acom_def widen1 Lc_acom_def) lemma widen_acom2: - "\\a\set(annos x). Wt a c; \a\set (annos y). Wt a c; strip x = strip y\ + "\\a\set(annos x). a \ Lc c; \a\set (annos y). a \ Lc c; strip x = strip y\ \ y \ x \ y" by(induct x y rule: le_acom.induct) - (auto simp: widen_acom_def widen2 Wt_acom_def) + (auto simp: widen_acom_def widen2 Lc_acom_def) lemma strip_map2_acom[simp]: "strip C1 = strip C2 \ strip(map2_acom f C1 C2) = strip C1" @@ -213,9 +215,9 @@ instance proof - case goal1 thus ?case by(auto simp: Wt_acom_def widen_acom1) + case goal1 thus ?case by(auto simp: Lc_acom_def widen_acom1) next - case goal2 thus ?case by(auto simp: Wt_acom_def widen_acom2) + case goal2 thus ?case by(auto simp: Lc_acom_def widen_acom2) next case goal3 thus ?case by(induct x y rule: le_acom.induct)(simp_all add: narrow_acom_def narrow1) @@ -224,36 +226,36 @@ by(induct x y rule: le_acom.induct)(simp_all add: narrow_acom_def narrow2) next case goal5 thus ?case - by(auto simp: Wt_acom_def widen_acom_def split_conv elim!: in_set_zipE) + by(auto simp: Lc_acom_def widen_acom_def split_conv elim!: in_set_zipE) next case goal6 thus ?case - by(auto simp: Wt_acom_def narrow_acom_def split_conv elim!: in_set_zipE) + by(auto simp: Lc_acom_def narrow_acom_def split_conv elim!: in_set_zipE) qed end -lemma wt_widen_o[simp]: fixes x1 x2 :: "_ st option" -shows "wt x1 X \ wt x2 X \ wt (x1 \ x2) X" +lemma widen_o_in_L[simp]: fixes x1 x2 :: "_ st option" +shows "x1 \ L X \ x2 \ L X \ x1 \ x2 \ L X" by(induction x1 x2 rule: widen_option.induct) - (simp_all add: widen_st_def wt_st_def) + (simp_all add: widen_st_def L_st_def) -lemma wt_narrow_o[simp]: fixes x1 x2 :: "_ st option" -shows "wt x1 X \ wt x2 X \ wt (x1 \ x2) X" +lemma narrow_o_in_L[simp]: fixes x1 x2 :: "_ st option" +shows "x1 \ L X \ x2 \ L X \ x1 \ x2 \ L X" by(induction x1 x2 rule: narrow_option.induct) - (simp_all add: narrow_st_def wt_st_def) + (simp_all add: narrow_st_def L_st_def) -lemma wt_widen_c: fixes C1 C2 :: "_ st option acom" -shows "strip C1 = strip C2 \ wt C1 X \ wt C2 X \ wt (C1 \ C2) X" +lemma widen_c_in_L: fixes C1 C2 :: "_ st option acom" +shows "strip C1 = strip C2 \ C1 \ L X \ C2 \ L X \ C1 \ C2 \ L X" by(induction C1 C2 rule: le_acom.induct) (auto simp: widen_acom_def) -lemma wt_narrow_c: fixes C1 C2 :: "_ st option acom" -shows "strip C1 = strip C2 \ wt C1 X \ wt C2 X \ wt (C1 \ C2) X" +lemma narrow_c_in_L: fixes C1 C2 :: "_ st option acom" +shows "strip C1 = strip C2 \ C1 \ L X \ C2 \ L X \ C1 \ C2 \ L X" by(induction C1 C2 rule: le_acom.induct) (auto simp: narrow_acom_def) -lemma Wt_bot[simp]: "Wt (bot c) c" -by(simp add: Wt_acom_def bot_def) +lemma bot_in_Lc[simp]: "bot c \ Lc c" +by(simp add: Lc_acom_def bot_def) subsubsection "Post-fixed point computation" @@ -264,7 +266,7 @@ definition iter_narrow :: "('a \ 'a) \ 'a \ ('a::{preord,narrow})option" where "iter_narrow f = while_option (\c. \ c \ c \ f c) (\c. c \ f c)" -definition pfp_wn :: "(('a::WN_Wt option acom) \ 'a option acom) \ com \ 'a option acom option" +definition pfp_wn :: "(('a::WN_Lc option acom) \ 'a option acom) \ com \ 'a option acom option" where "pfp_wn f c = (case iter_widen f (bot c) of None \ None | Some c' \ iter_narrow f c')" @@ -284,7 +286,7 @@ using while_option_rule[where P = "\C'. strip C' = strip C", OF _ assms(2)] by (metis assms(1)) -lemma strip_iter_widen: fixes f :: "'a::WN_Wt acom \ 'a acom" +lemma strip_iter_widen: fixes f :: "'a::WN_Lc acom \ 'a acom" assumes "\C. strip (f C) = strip C" and "iter_widen f C = Some C'" shows "strip C' = strip C" proof- @@ -294,7 +296,7 @@ qed lemma iter_narrow_pfp: -assumes mono: "!!c1 c2::_::WN_Wt. P c1 \ P c2 \ c1 \ c2 \ f c1 \ f c2" +assumes mono: "!!c1 c2::_::WN_Lc. P c1 \ P c2 \ c1 \ c2 \ f c1 \ f c2" and Pinv: "!!c. P c \ P(f c)" "!!c1 c2. P c1 \ P c2 \ P(c1 \ c2)" and "P c0" and "f c0 \ c0" and "iter_narrow f c0 = Some c" shows "P c \ f c \ c" @@ -321,7 +323,7 @@ qed lemma pfp_wn_pfp: -assumes mono: "!!c1 c2::_::WN_Wt option acom. P c1 \ P c2 \ c1 \ c2 \ f c1 \ f c2" +assumes mono: "!!c1 c2::_::WN_Lc option acom. P c1 \ P c2 \ c1 \ c2 \ f c1 \ f c2" and Pinv: "P (bot c)" "!!c. P c \ P(f c)" "!!c1 c2. P c1 \ P c2 \ P(c1 \ c2)" "!!c1 c2. P c1 \ P c2 \ P(c1 \ c2)" @@ -343,7 +345,7 @@ locale Abs_Int2 = Abs_Int1_mono -where \=\ for \ :: "'av::{WN,L_top_bot} \ val set" +where \=\ for \ :: "'av::{WN,lattice} \ val set" begin definition AI_wn :: "com \ 'av st option acom option" where @@ -352,17 +354,17 @@ lemma AI_wn_sound: "AI_wn c = Some C \ CS c \ \\<^isub>c C" proof(simp add: CS_def AI_wn_def) assume 1: "pfp_wn (step' (top c)) c = Some C" - have 2: "(strip C = c & wt C (vars c)) \ step' \\<^bsub>c\<^esub> C \ C" + have 2: "(strip C = c & C \ L(vars c)) \ step' \\<^bsub>c\<^esub> C \ C" by(rule pfp_wn_pfp[where c=c]) - (simp_all add: 1 mono_step'_top wt_widen_c wt_narrow_c) + (simp_all add: 1 mono_step'_top widen_c_in_L narrow_c_in_L) have 3: "strip (\\<^isub>c (step' \\<^bsub>c\<^esub> C)) = c" by(simp add: strip_pfp_wn[OF _ 1]) have "lfp c (step UNIV) \ \\<^isub>c (step' \\<^bsub>c\<^esub> C)" proof(rule lfp_lowerbound[simplified,OF 3]) show "step UNIV (\\<^isub>c (step' \\<^bsub>c\<^esub> C)) \ \\<^isub>c (step' \\<^bsub>c\<^esub> C)" - proof(rule step_preserves_le[OF _ _ _ wt_top]) + proof(rule step_preserves_le[OF _ _ _ top_in_L]) show "UNIV \ \\<^isub>o \\<^bsub>c\<^esub>" by simp show "\\<^isub>c (step' \\<^bsub>c\<^esub> C) \ \\<^isub>c C" by(rule mono_gamma_c[OF conjunct2[OF 2]]) - show "wt C (vars c)" using 2 by blast + show "C \ L(vars c)" using 2 by blast qed qed from this 2 show "lfp c (step UNIV) \ \\<^isub>c C" @@ -381,6 +383,10 @@ subsubsection "Tests" +(* FIXME dirty trick to get around some problem with the code generator *) +lemma [code]: "equal_class.equal (x::'a st) y = equal_class.equal x y" +by(rule refl) + definition "step_up_ivl n = ((\C. C \ step_ivl (top(strip C)) C)^^n)" definition "step_down_ivl n = @@ -415,7 +421,7 @@ subsubsection "Generic Termination Proof" locale Abs_Int2_measure = Abs_Int2 - where \=\ for \ :: "'av::{WN,L_top_bot} \ val set" + + where \=\ for \ :: "'av::{WN,lattice} \ val set" + fixes m :: "'av \ nat" fixes n :: "'av \ nat" fixes h :: "nat" @@ -450,9 +456,9 @@ by (metis setsum_mono) qed -lemma m_st_widen: "wt S1 X \ wt S2 X \ finite X \ +lemma m_st_widen: "S1 \ L X \ S2 \ L X \ finite X \ ~ S2 \ S1 \ m_st(S1 \ S2) < m_st S1" -proof(auto simp add: le_st_def m_st_def wt_st_def widen_st_def) +proof(auto simp add: le_st_def m_st_def L_st_def widen_st_def) assume "finite(dom S1)" have 1: "\x\dom S1. m(fun S1 x) \ m(fun S1 x \ fun S2 x)" by (metis m_anti_mono WN_class.widen1) @@ -493,16 +499,16 @@ definition "m_o k opt = (case opt of None \ k+1 | Some x \ m_st x)" -lemma m_o_anti_mono: "wt S1 X \ wt S2 X \ finite X \ +lemma m_o_anti_mono: "S1 \ L X \ S2 \ L X \ finite X \ S1 \ S2 \ m_o (h * card X) S2 \ m_o (h * card X) S1" apply(induction S1 S2 rule: le_option.induct) -apply(auto simp: m_o_def m_st_anti_mono le_SucI h_st wt_st_def +apply(auto simp: m_o_def m_st_anti_mono le_SucI h_st L_st_def split: option.splits) done -lemma m_o_widen: "\ wt S1 X; wt S2 X; finite X; \ S2 \ S1 \ \ +lemma m_o_widen: "\ S1 \ L X; S2 \ L X; finite X; \ S2 \ S1 \ \ m_o (h * card X) (S1 \ S2) < m_o (h * card X) S1" -by(auto simp: m_o_def wt_st_def h_st less_Suc_eq_le m_st_widen +by(auto simp: m_o_def L_st_def h_st less_Suc_eq_le m_st_widen split: option.split) definition "n_o opt = (case opt of None \ 0 | Some x \ n_st x + 1)" @@ -511,14 +517,14 @@ by(induction S1 S2 rule: le_option.induct)(auto simp: n_o_def n_st_mono) lemma n_o_narrow: - "wt S1 X \ wt S2 X \ finite X + "S1 \ L X \ S2 \ L X \ finite X \ S2 \ S1 \ \ S1 \ S1 \ S2 \ n_o (S1 \ S2) < n_o S1" apply(induction S1 S2 rule: narrow_option.induct) -apply(auto simp: n_o_def wt_st_def n_st_narrow) +apply(auto simp: n_o_def L_st_def n_st_narrow) done -lemma annos_narrow_acom[simp]: "strip C2 = strip (C1::'a::WN_Wt acom) \ +lemma annos_narrow_acom[simp]: "strip C2 = strip (C1::'a::WN_Lc acom) \ annos(C1 \ C2) = map (\(x,y).x\y) (zip (annos C1) (annos C2))" by(induction "narrow::'a\'a\'a" C1 C2 rule: map2_acom.induct) (simp_all add: narrow_acom_def size_annos_same2) @@ -528,8 +534,8 @@ \i=0.. Wt C2 c \ \ C2 \ C1 \ m_c (C1 \ C2) < m_c C1" -apply(auto simp: Wt_acom_def m_c_def Let_def widen_acom_def) + "C1 \ Lc c \ C2 \ Lc c \ \ C2 \ C1 \ m_c (C1 \ C2) < m_c C1" +apply(auto simp: Lc_acom_def m_c_def Let_def widen_acom_def) apply(subgoal_tac "length(annos C2) = length(annos C1)") prefer 2 apply (simp add: size_annos_same2) apply (auto) @@ -541,14 +547,14 @@ apply(rule_tac x=i in bexI) prefer 2 apply simp apply(rule m_o_widen) -apply (simp add: finite_cvars)+(*FIXME [simp]*) +apply (simp add: finite_cvars)+ done definition "n_c C = (let as = annos C in \i=0.. Wt C2 c \ C2 \ C1 \ \ C1 \ C1 \ C2 \ n_c (C1 \ C2) < n_c C1" -apply(auto simp: n_c_def Let_def Wt_acom_def narrow_acom_def) +lemma n_c_narrow: "C1 \ Lc c \ C2 \ Lc c \ + C2 \ C1 \ \ C1 \ C1 \ C2 \ n_c (C1 \ C2) < n_c C1" +apply(auto simp: n_c_def Let_def Lc_acom_def narrow_acom_def) apply(subgoal_tac "length(annos C2) = length(annos C1)") prefer 2 apply (simp add: size_annos_same2) apply (auto) @@ -569,7 +575,7 @@ lemma iter_widen_termination: -fixes m :: "'a::WN_Wt \ nat" +fixes m :: "'a::WN_Lc \ nat" assumes P_f: "\C. P C \ P(f C)" and P_widen: "\C1 C2. P C1 \ P C2 \ P(C1 \ C2)" and m_widen: "\C1 C2. P C1 \ P C2 \ ~ C2 \ C1 \ m(C1 \ C2) < m C1" @@ -584,7 +590,7 @@ qed thm mono_step'(*FIXME does not need wt assms*) lemma iter_narrow_termination: -fixes n :: "'a::WN_Wt \ nat" +fixes n :: "'a::WN_Lc \ nat" assumes P_f: "\C. P C \ P(f C)" and P_narrow: "\C1 C2. P C1 \ P C2 \ P(C1 \ C2)" and mono: "\C1 C2. P C1 \ P C2 \ C1 \ C2 \ f C1 \ f C2" @@ -654,19 +660,19 @@ lemma iter_winden_step_ivl_termination: "\C. iter_widen (step_ivl (top c)) (bot c) = Some C" -apply(rule iter_widen_termination[where m = "m_c" and P = "%C. Wt C c"]) -apply (simp_all add: Wt_step' m_c_widen) +apply(rule iter_widen_termination[where m = "m_c" and P = "%C. C \ Lc c"]) +apply (simp_all add: step'_in_Lc m_c_widen) done lemma iter_narrow_step_ivl_termination: - "Wt C0 c \ step_ivl (top c) C0 \ C0 \ + "C0 \ Lc c \ step_ivl (top c) C0 \ C0 \ \C. iter_narrow (step_ivl (top c)) C0 = Some C" -apply(rule iter_narrow_termination[where n = "n_c" and P = "%C. Wt C c"]) -apply (simp add: Wt_step') +apply(rule iter_narrow_termination[where n = "n_c" and P = "%C. C \ Lc c"]) +apply (simp add: step'_in_Lc) apply (simp) apply(rule mono_step'_top) -apply(simp add: Wt_acom_def wt_acom_def) -apply(simp add: Wt_acom_def wt_acom_def) +apply(simp add: Lc_acom_def L_acom_def) +apply(simp add: Lc_acom_def L_acom_def) apply assumption apply(erule (3) n_c_narrow) apply assumption @@ -678,7 +684,7 @@ apply(auto simp: AI_wn_def pfp_wn_def iter_winden_step_ivl_termination split: option.split) apply(rule iter_narrow_step_ivl_termination) -apply(blast intro: iter_widen_inv[where f = "step' \\<^bsub>c\<^esub>" and P = "%C. Wt C c"] Wt_bot Wt_widen Wt_step'[where S = "\\<^bsub>c\<^esub>" and c=c, simplified]) +apply(blast intro: iter_widen_inv[where f = "step' \\<^bsub>c\<^esub>" and P = "%C. C \ Lc c"] bot_in_Lc Lc_widen step'_in_Lc[where S = "\\<^bsub>c\<^esub>" and c=c, simplified]) apply(erule iter_widen_pfp) done diff -r 323414474c1f -r 73fb17ed2e08 src/HOL/IMP/Abs_State.thy --- a/src/HOL/IMP/Abs_State.thy Sun Sep 16 10:33:25 2012 +0200 +++ b/src/HOL/IMP/Abs_State.thy Sun Sep 16 11:50:03 2012 +0200 @@ -4,7 +4,7 @@ imports Abs_Int0 begin -subsubsection "Welltypedness: wt" +subsubsection "Set-based lattices" instantiation com :: vars begin @@ -31,45 +31,45 @@ by(induction c) (simp_all add: finite_avars finite_bvars) -class wt = -fixes wt :: "'a \ vname set \ bool" +class L = +fixes L :: "vname set \ 'a set" -instantiation acom :: (wt)wt +instantiation acom :: (L)L begin -definition wt_acom where -"wt C X = (vars(strip C) \ X \ (\a \ set(annos C). wt a X))" +definition L_acom where +"L X = {C. vars(strip C) \ X \ (\a \ set(annos C). a \ L X)}" instance .. end -instantiation option :: (wt)wt +instantiation option :: (L)L begin -definition wt_option where -"wt opt X = (case opt of None \ True | Some x \ wt x X)" +definition L_option where +"L X = {opt. case opt of None \ True | Some x \ x \ L X}" -lemma wt_option_simps[simp]: "wt None X" "wt (Some x) X = wt x X" -by(simp_all add: wt_option_def) +lemma L_option_simps[simp]: "None \ L X" "(Some x \ L X) = (x \ L X)" +by(simp_all add: L_option_def) instance .. end -class SL_top_wt = join + wt + +class semilatticeL = join + L + fixes top :: "com \ 'a" ("\\<^bsub>_\<^esub>") -assumes join_ge1 [simp]: "wt x X \ wt y X \ x \ x \ y" -and join_ge2 [simp]: "wt x X \ wt y X \ y \ x \ y" +assumes join_ge1 [simp]: "x \ L X \ y \ L X \ x \ x \ y" +and join_ge2 [simp]: "x \ L X \ y \ L X \ y \ x \ y" and join_least[simp]: "x \ z \ y \ z \ x \ y \ z" -and top[simp]: "wt x (vars c) \ x \ top c" -and wt_top[simp]: "wt (top c) (vars c)" -and wt_join[simp]: "wt x X \ wt y X \ wt (x \ y) X" +and top[simp]: "x \ L(vars c) \ x \ top c" +and top_in_L[simp]: "top c \ L(vars c)" +and join_in_L[simp]: "x \ L X \ y \ L X \ x \ y \ L X" -instantiation option :: (SL_top_wt)SL_top_wt +instantiation option :: (semilatticeL)semilatticeL begin definition top_option where "top c = Some(top c)" @@ -85,7 +85,7 @@ next case goal5 thus ?case by(simp add: top_option_def) next - case goal6 thus ?case by(simp add: wt_option_def split: option.splits) + case goal6 thus ?case by(simp add: L_option_def split: option.splits) qed end @@ -143,24 +143,24 @@ end -instantiation st :: (type) wt +instantiation st :: (type) L begin -definition wt_st :: "'a st \ vname set \ bool" where -"wt F X = (dom F = X)" +definition L_st :: "vname set \ 'a st set" where +"L X = {F. dom F = X}" instance .. end -instantiation st :: (SL_top) SL_top_wt +instantiation st :: (semilattice) semilatticeL begin definition top_st where "top c = FunDom (\x. \) (vars c)" instance proof -qed (auto simp: le_st_def join_st_def top_st_def wt_st_def) +qed (auto simp: le_st_def join_st_def top_st_def L_st_def) end @@ -173,7 +173,7 @@ by(auto simp add: le_st_def update_def) -locale Gamma = Val_abs where \=\ for \ :: "'av::SL_top \ val set" +locale Gamma = Val_abs where \=\ for \ :: "'av::semilattice \ val set" begin abbreviation \\<^isub>f :: "'av st \ state set"