diff -r f692657e0f71 -r df3426139651 src/HOL/IMP/Abs_Int1.thy --- a/src/HOL/IMP/Abs_Int1.thy Wed Apr 17 20:53:26 2013 +0200 +++ b/src/HOL/IMP/Abs_Int1.thy Wed Apr 17 21:11:01 2013 +0200 @@ -4,6 +4,20 @@ imports Abs_State begin +(* FIXME mv *) +instantiation acom :: (type) vars +begin + +definition "vars_acom = vars o strip" + +instance .. + +end + +lemma finite_Cvars: "finite(vars(C::'a acom))" +by(simp add: vars_acom_def) + + 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: less_eq_acom.induct) (auto simp: size_annos_same2) @@ -12,28 +26,10 @@ 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_L[simp]: "F \ L X \ F \ G \ x : X \ fun F x \ fun G x" -by(simp add: mono_fun L_st_def) - -lemma bot_in_L[simp]: "bot c \ L(vars c)" -by(simp add: L_acom_def bot_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)" +(* FIXME mv *) +lemma post_in_annos: "post C \ set(annos C)" by(induction C) auto -lemma post_in_L[simp]: "C \ L X \ post C \ L X" -by(simp add: L_acom_def post_in_annos) - subsection "Computable Abstract Interpretation" @@ -47,15 +43,15 @@ "aval' (V x) S = fun S x" | "aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" -lemma aval'_sound: "s : \\<^isub>s S \ vars a \ dom S \ aval a s : \(aval' a S)" +lemma aval'_sound: "s : \\<^isub>s S \ aval a s : \(aval' a S)" by (induction a) (auto simp: gamma_num' gamma_plus' \_st_def) lemma gamma_Step_subcomm: fixes C1 C2 :: "'a::semilattice_sup acom" - assumes "!!x e S. x : X \ vars e \ X \ S \ L X \ f1 x e (\\<^isub>o S) \ \\<^isub>o (f2 x e S)" - "!!b S. S \ L X \ vars b \ X \ g1 b (\\<^isub>o S) \ \\<^isub>o (g2 b S)" - shows "C \ L X \ S \ L X \ Step f1 g1 (\\<^isub>o S) (\\<^isub>c C) \ \\<^isub>c (Step f2 g2 S C)" + assumes "!!x e S. f1 x e (\\<^isub>o S) \ \\<^isub>o (f2 x e S)" + "!!b S. g1 b (\\<^isub>o S) \ \\<^isub>o (g2 b S)" + shows "Step f1 g1 (\\<^isub>o S) (\\<^isub>c C) \ \\<^isub>c (Step f2 g2 S C)" proof(induction C arbitrary: S) -qed (auto simp: assms intro!: mono_gamma_o post_in_L sup_ge1 sup_ge2) +qed (auto simp: assms intro!: mono_gamma_o sup_ge1 sup_ge2) lemma in_gamma_update: "\ s : \\<^isub>s S; i : \ a \ \ s(x := i) : \\<^isub>s(update S x a)" by(simp add: \_st_def) @@ -63,16 +59,6 @@ end -lemma Step_in_L: fixes C :: "'a::semilatticeL acom" -assumes "!!x e S. \S \ L X; x \ X; vars e \ X\ \ f x e S \ L X" - "!!b S. S \ L X \ vars b \ X \ g b S \ L X" -shows"\ C \ L X; S \ L X \ \ Step f g S C \ L X" -proof(induction C arbitrary: S) - case Assign thus ?case - by(auto simp: L_st_def assms split: option.splits) -qed (auto simp: assms) - - 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}. *} @@ -85,7 +71,7 @@ (\b S. S)" definition AI :: "com \ 'av st option acom option" where -"AI c = pfp (step' (Top(vars c))) (bot c)" +"AI c = pfp (step' \) (bot c)" lemma strip_step'[simp]: "strip(step' S C) = strip C" @@ -94,32 +80,23 @@ text{* Soundness: *} -lemma step_step': "C \ L X \ S \ L X \ step (\\<^isub>o S) (\\<^isub>c C) \ \\<^isub>c (step' S C)" +lemma step_step': "step (\\<^isub>o S) (\\<^isub>c C) \ \\<^isub>c (step' S C)" unfolding step_def step'_def by(rule gamma_Step_subcomm) - (auto simp: L_st_def intro!: aval'_sound in_gamma_update split: option.splits) - -lemma step'_in_L[simp]: "\ C \ L X; S \ L X \ \ step' S C \ L X" -unfolding step'_def -by(rule Step_in_L)(auto simp: L_st_def step'_def split: option.splits) + (auto simp: intro!: aval'_sound in_gamma_update split: option.splits) lemma AI_sound: "AI c = Some C \ CS c \ \\<^isub>c C" proof(simp add: CS_def AI_def) - assume 1: "pfp (step' (Top(vars c))) (bot c) = Some C" - have "C \ L(vars c)" - by(rule pfp_inv[where P = "%C. C \ L(vars c)", OF 1 _ bot_in_L]) - (erule step'_in_L[OF _ Top_in_L]) - have pfp': "step' (Top(vars c)) C \ C" by(rule pfp_pfp[OF 1]) - have 2: "step (\\<^isub>o(Top(vars c))) (\\<^isub>c C) \ \\<^isub>c C" + assume 1: "pfp (step' \) (bot c) = Some C" + have pfp': "step' \ C \ C" by(rule pfp_pfp[OF 1]) + have 2: "step (\\<^isub>o \) (\\<^isub>c C) \ \\<^isub>c C" --"transfer the pfp'" proof(rule order_trans) - show "step (\\<^isub>o (Top(vars c))) (\\<^isub>c C) \ \\<^isub>c (step' (Top(vars c)) C)" - by(rule step_step'[OF `C \ L(vars c)` Top_in_L]) - show "\\<^isub>c (step' (Top(vars c)) C) \ \\<^isub>c C" - by(rule mono_gamma_c[OF pfp']) + show "step (\\<^isub>o \) (\\<^isub>c C) \ \\<^isub>c (step' \ C)" by(rule step_step') + show "... \ \\<^isub>c C" by (metis mono_gamma_c[OF pfp']) qed have 3: "strip (\\<^isub>c C) = c" by(simp add: strip_pfp[OF _ 1] step'_def) - have "lfp c (step (\\<^isub>o(Top(vars c)))) \ \\<^isub>c C" - by(rule lfp_lowerbound[simplified,where f="step (\\<^isub>o(Top(vars c)))", OF 3 2]) + have "lfp c (step (\\<^isub>o \)) \ \\<^isub>c C" + by(rule lfp_lowerbound[simplified,where f="step (\\<^isub>o \)", OF 3 2]) thus "lfp c (step UNIV) \ \\<^isub>c C" by simp qed @@ -128,49 +105,34 @@ subsubsection "Monotonicity" -lemma le_sup_disj: "y \ L X \ (z::_::semilatticeL) \ L X \ - x \ y \ x \ z \ x \ y \ z" +lemma le_sup_disj: "x \ y \ x \ z \ x \ y \ (z::_::semilattice_sup)" by (metis sup_ge1 sup_ge2 order_trans) -theorem mono2_Step: fixes C1 :: "'a::semilatticeL acom" - assumes "!!x e S1 S2. \S1 \ L X; S2 \ L X; x \ X; vars e \ X; S1 \ S2\ \ f x e S1 \ f x e S2" - "!!b S1 S2. S1 \ L X \ S2 \ L X \ vars b \ X \ S1 \ S2 \ g b S1 \ g b S2" - shows "S1 \ L X \ S2 \ L X \ C1 \ L X \ C2 \ L X \ - S1 \ S2 \ C1 \ C2 \ Step f g S1 C1 \ Step f g S2 C2" +theorem mono2_Step: fixes C1 :: "'a::semilattice acom" + assumes "!!x e S1 S2. S1 \ S2 \ f x e S1 \ f x e S2" + "!!b S1 S2. S1 \ S2 \ g b S1 \ g b S2" + shows "S1 \ S2 \ C1 \ C2 \ Step f g S1 C1 \ Step f g S2 C2" by(induction C1 C2 arbitrary: S1 S2 rule: less_eq_acom.induct) - (auto simp: mono_post assms le_sup_disj le_sup_disj[OF post_in_L post_in_L]) + (auto simp: mono_post assms le_sup_disj) 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 \ S1 \ L X \ S2 \ L X \ vars e \ X \ aval' e S1 \ aval' e S2" -by(induction e) (auto simp: less_eq_st_def mono_plus' L_st_def) +lemma mono_aval': "S1 \ S2 \ aval' e S1 \ aval' e S2" +by(induction e) (auto simp: mono_plus' mono_fun) -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" +theorem mono_step': "S1 \ S2 \ C1 \ C2 \ step' S1 C1 \ step' S2 C2" unfolding step'_def by(rule mono2_Step) (auto simp: mono_aval' split: option.split) -lemma mono_step'_top: "C \ L X \ C' \ L X \ - C \ C' \ step' (Top X) C \ step' (Top X) C'" -by (metis Top_in_L mono_step' order_refl) +lemma mono_step'_top: "C \ C' \ step' \ C \ step' \ C'" +by (metis mono_step' order_refl) -lemma pfp_bot_least: -assumes "\x\L(vars c)\{C. strip C = c}.\y\L(vars c)\{C. strip C = c}. - x \ y \ f x \ f y" -and "\C. C \ L(vars c)\{C. strip C = c} \ f C \ L(vars c)\{C. strip C = c}" -and "f C' \ C'" "strip C' = c" "C' \ L(vars c)" "pfp f (bot c) = Some C" +lemma AI_least_pfp: assumes "AI c = Some C" "step' \ C' \ C'" "strip C' = c" shows "C \ C'" -by(rule while_least[OF assms(1,2) _ _ assms(3) _ assms(6)[unfolded pfp_def]]) - (simp_all add: assms(4,5) bot_least) - -lemma AI_least_pfp: assumes "AI c = Some C" -and "step' (Top(vars c)) C' \ C'" "strip C' = c" "C' \ L(vars c)" -shows "C \ C'" -by(rule pfp_bot_least[OF _ _ assms(2-4) assms(1)[unfolded AI_def]]) +by(rule pfp_bot_least[OF _ _ assms(2,3) assms(1)[unfolded AI_def]]) (simp_all add: mono_step'_top) end @@ -196,44 +158,112 @@ locale Measure1 = -fixes m :: "'av::order \ nat" +fixes m :: "'av::{order,top} \ nat" fixes h :: "nat" assumes h: "m x \ h" begin -definition m_s :: "'av st \ nat" ("m\<^isub>s") where -"m_s S = (\ x \ dom S. m(fun S x))" +definition m_s :: "vname set \ 'av st \ nat" ("m\<^isub>s") where +"m_s X S = (\ x \ X. m(fun S x))" + +lemma m_s_h: "finite X \ m_s X S \ h * card X" +by(simp add: m_s_def) (metis nat_mult_commute of_nat_id setsum_bounded[OF h]) -lemma m_s_h: "x \ L X \ finite X \ m_s x \ h * card X" -by(simp add: L_st_def m_s_def) - (metis nat_mult_commute of_nat_id setsum_bounded[OF h]) +definition m_o :: "vname set \ 'av st option \ nat" ("m\<^isub>o") where +"m_o X opt = (case opt of None \ h * card X + 1 | Some S \ m_s X S)" -definition m_o :: "nat \ 'av st option \ nat" ("m\<^isub>o") where -"m_o d opt = (case opt of None \ h*d+1 | Some S \ m_s S)" - -lemma m_o_h: "ost \ L X \ finite X \ m_o (card X) ost \ (h*card X + 1)" -by(auto simp add: m_o_def m_s_h split: option.split dest!:m_s_h) +lemma m_o_h: "finite X \ m_o X opt \ (h*card X + 1)" +by(auto simp add: m_o_def m_s_h le_SucI split: option.split dest:m_s_h) definition m_c :: "'av st option acom \ nat" ("m\<^isub>c") where -"m_c C = (\ii L(vars(strip C))" -shows "m_c C \ size(annos C) * (h * card(vars(strip C)) + 1)" +text{* Upper complexity bound: *} +lemma m_c_h: "m_c C \ size(annos C) * (h * card(vars C) + 1)" proof- - let ?X = "vars(strip C)" let ?n = "card ?X" let ?a = "size(annos C)" - { fix i assume "i < ?a" - hence "annos C ! i \ L ?X" using assms by(simp add: L_acom_def) - note m_o_h[OF this finite_cvars] - } note 1 = this - have "m_c C = (\ii \ (\i = ?a * (h * ?n + 1)" by simp finally show ?thesis . qed end +text{* The predicates @{text "top_on X a"} that follow describe that @{text a} is some object +that maps all variables in @{text X} to @{term \}. +This is an important invariant for the termination proof where we argue that only +the finitely many variables in the program change. That the others do not change +follows because they remain @{term \}. *} + +class top_on = +fixes top_on :: "vname set \ 'a \ bool" + +instantiation st :: (top)top_on +begin + +fun top_on_st :: "vname set \ 'a st \ bool" where +"top_on_st X F = (\x\X. fun F x = \)" + +instance .. + +end + +instantiation option :: (top_on)top_on +begin + +fun top_on_option :: "vname set \ 'a option \ bool" where +"top_on_option X (Some F) = top_on X F" | +"top_on_option X None = True" + +instance .. + +end + +instantiation acom :: (top_on)top_on +begin + +definition top_on_acom :: "vname set \ 'a acom \ bool" where +"top_on_acom X C = (\a \ set(annos C). top_on X a)" + +instance .. + +end + +lemma top_on_top: "top_on X (\::_ st option)" +by(auto simp: top_option_def fun_top) + +lemma top_on_bot: "top_on X (bot c)" +by(auto simp add: top_on_acom_def bot_def) + +lemma top_on_post: "top_on X C \ top_on X (post C)" +by(simp add: top_on_acom_def post_in_annos) + +lemma top_on_acom_simps: + "top_on X (SKIP {Q}) = top_on X Q" + "top_on X (x ::= e {Q}) = top_on X Q" + "top_on X (C1;C2) = (top_on X C1 \ top_on X C2)" + "top_on X (IF b THEN {P1} C1 ELSE {P2} C2 {Q}) = + (top_on X P1 \ top_on X C1 \ top_on X P2 \ top_on X C2 \ top_on X Q)" + "top_on X ({I} WHILE b DO {P} C {Q}) = + (top_on X I \ top_on X C \ top_on X P \ top_on X Q)" +by(auto simp add: top_on_acom_def) + +lemma top_on_sup: + "top_on X o1 \ top_on X o2 \ top_on X (o1 \ o2 :: _ st option)" +apply(induction o1 o2 rule: sup_option.induct) +apply(auto) +by transfer simp + +lemma top_on_Step: fixes C :: "('a::semilattice)st option acom" +assumes "!!x e S. \top_on X S; x \ X; vars e \ -X\ \ top_on X (f x e S)" + "!!b S. top_on X S \ vars b \ -X \ top_on X (g b S)" +shows "\ vars C \ -X; top_on X S; top_on X C \ \ top_on X (Step f g S C)" +proof(induction C arbitrary: S) +qed (auto simp: top_on_acom_simps vars_acom_def top_on_post top_on_sup assms) + + locale Measure = Measure1 + assumes m2: "x < y \ m x > m y" begin @@ -241,62 +271,79 @@ lemma m1: "x \ y \ m x \ m y" by(auto simp: le_less m2) -lemma m_s2: "finite(dom S1) \ S1 < S2 \ m_s S1 > m_s S2" -proof(auto simp add: less_st_def less_eq_st_def m_s_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)" by (metis 0 m2 less_le_not_le) - 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))" . +lemma m_s2_rep: assumes "finite(X)" and "S1 = S2 on -X" and "\x. S1 x \ S2 x" and "S1 \ S2" +shows "(\x\X. m (S2 x)) < (\x\X. m (S1 x))" +proof- + from assms(3) have 1: "\x\X. m(S1 x) \ m(S2 x)" by (simp add: m1) + from assms(2,3,4) have "EX x:X. S1 x < S2 x" + by(simp add: fun_eq_iff) (metis Compl_iff le_neq_trans) + hence 2: "\x\X. m(S1 x) > m(S2 x)" by (metis m2) + from setsum_strict_mono_ex1[OF `finite X` 1 2] + show "(\x\X. m (S2 x)) < (\x\X. m (S1 x))" . qed -lemma m_o2: "finite X \ o1 \ L X \ o2 \ L X \ - o1 < o2 \ m_o (card X) o1 > m_o (card X) o2" +lemma m_s2: "finite(X) \ fun S1 = fun S2 on -X \ S1 < S2 \ m_s X S1 > m_s X S2" +apply(auto simp add: less_st_def m_s_def) +apply (transfer fixing: m) +apply(simp add: less_eq_st_rep_iff eq_st_def m_s2_rep) +done + +lemma m_o2: "finite X \ top_on (-X) o1 \ top_on (-X) o2 \ + o1 < o2 \ m_o X o1 > m_o X o2" proof(induction o1 o2 rule: less_eq_option.induct) - case 1 thus ?case by (auto simp: m_o_def L_st_def m_s2 less_option_def) + case 1 thus ?case by (auto simp: m_o_def m_s2 less_option_def) next case 2 thus ?case by(auto simp: m_o_def less_option_def le_imp_less_Suc m_s_h) next case 3 thus ?case by (auto simp: less_option_def) qed -lemma m_o1: "finite X \ o1 \ L X \ o2 \ L X \ - o1 \ o2 \ m_o (card X) o1 \ m_o (card X) o2" +lemma m_o1: "finite X \ top_on (-X) o1 \ top_on (-X) o2 \ + o1 \ o2 \ m_o X o1 \ m_o X o2" by(auto simp: le_less m_o2) -lemma m_c2: "C1 \ L(vars(strip C1)) \ C2 \ L(vars(strip C2)) \ + +lemma m_c2: "top_on (-vars C1) C1 \ top_on (-vars C2) 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] less_acom_def L_acom_def) +proof(auto simp add: le_iff_le_annos m_c_def size_annos_same[of C1 C2] vars_acom_def less_acom_def) let ?X = "vars(strip C2)" - let ?n = "card ?X" - 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 nth_mem size_annos_same[OF strip_eq] 0 less_option_def) + assume top: "top_on (- vars(strip C2)) C1" "top_on (- vars(strip C2)) C2" + and strip_eq: "strip C1 = strip C2" + and 0: "\i annos C2 ! i" + hence 1: "\i m_o ?X (annos C2 ! i)" + apply (auto simp: all_set_conv_all_nth vars_acom_def top_on_acom_def) + by (metis finite_cvars m_o1 size_annos_same2) + fix i assume i: "i < size(annos C2)" "\ annos C2 ! i \ annos C1 ! i" + have topo1: "top_on (- ?X) (annos C1 ! i)" + using i(1) top(1) by(simp add: top_on_acom_def size_annos_same[OF strip_eq]) + have topo2: "top_on (- ?X) (annos C2 ! i)" + using i(1) top(2) by(simp add: top_on_acom_def size_annos_same[OF strip_eq]) + from i have "m_o ?X (annos C1 ! i) > m_o ?X (annos C2 ! i)" (is "?P i") + by (metis 0 less_option_def m_o2[OF finite_cvars topo1] topo2) hence 2: "\i < size(annos C2). ?P i" using `i < size(annos C2)` by blast - show "(\iiii=\ + Measure where m=m for \ :: "'av::semilattice \ val set" and m :: "'av \ nat" begin +lemma top_on_step': "\ vars C \ -X; top_on X C \ \ top_on X (step' \ C)" +unfolding step'_def +by(rule top_on_Step) + (auto simp add: top_option_def fun_top split: option.splits) + lemma AI_Some_measure: "\C. AI c = Some C" unfolding AI_def -apply(rule pfp_termination[where I = "\C. strip C = c \ C \ L(vars c)" and m="m_c"]) -apply(simp_all add: m_c2 mono_step'_top bot_least) +apply(rule pfp_termination[where I = "\C. strip C = c \ top_on (- vars C) C" and m="m_c"]) +apply(simp_all add: m_c2 mono_step'_top bot_least top_on_bot) +apply(auto simp add: top_on_step' vars_acom_def) done end