# HG changeset patch # User nipkow # Date 1326113294 -3600 # Node ID 8b5f1f91ef383c625e94baff804c047a57f81cc2 # Parent 3d518b508bbb10552140d3ca85aa1f03cf6f9e3d Added termination to IMP Abs_Int diff -r 3d518b508bbb -r 8b5f1f91ef38 src/HOL/IMP/Abs_Int0.thy --- a/src/HOL/IMP/Abs_Int0.thy Mon Jan 09 11:41:38 2012 +0100 +++ b/src/HOL/IMP/Abs_Int0.thy Mon Jan 09 13:48:14 2012 +0100 @@ -132,4 +132,260 @@ end + +subsubsection "Ascending Chain Condition" + +hide_const acc + +abbreviation "strict r == r \ -(r^-1)" +abbreviation "acc r == wf((strict r)^-1)" + +lemma strict_inv_image: "strict(inv_image r f) = inv_image (strict r) f" +by(auto simp: inv_image_def) + +lemma acc_inv_image: + "acc r \ acc (inv_image r f)" +by (metis converse_inv_image strict_inv_image wf_inv_image) + +text{* ACC for option type: *} + +lemma acc_option: assumes "acc {(x,y::'a::preord). x \ y}" +shows "acc {(x,y::'a option). x \ y}" +proof(auto simp: wf_eq_minimal) + fix xo :: "'a option" and Qo assume "xo : Qo" + let ?Q = "{x. Some x \ Qo}" + show "\yo\Qo. \zo. yo \ zo \ ~ zo \ yo \ zo \ Qo" (is "\zo\Qo. ?P zo") + proof cases + assume "?Q = {}" + hence "?P None" by auto + moreover have "None \ Qo" using `?Q = {}` `xo : Qo` + by auto (metis not_Some_eq) + ultimately show ?thesis by blast + next + assume "?Q \ {}" + with assms show ?thesis + apply(auto simp: wf_eq_minimal) + apply(erule_tac x="?Q" in allE) + apply auto + apply(rule_tac x = "Some z" in bexI) + by auto + qed +qed + +text{* ACC for abstract states, via measure functions. *} + +(*FIXME mv*) +lemma setsum_strict_mono1: +fixes f :: "'a \ 'b::{comm_monoid_add, ordered_cancel_ab_semigroup_add}" +assumes "finite A" and "ALL x:A. f x \ g x" and "EX a:A. f a < g a" +shows "setsum f A < setsum g A" +proof- + from assms(3) obtain a where a: "a:A" "f a < g a" by blast + have "setsum f A = setsum f ((A-{a}) \ {a})" + by(simp add:insert_absorb[OF `a:A`]) + also have "\ = setsum f (A-{a}) + setsum f {a}" + using `finite A` by(subst setsum_Un_disjoint) auto + also have "setsum f (A-{a}) \ setsum g (A-{a})" + by(rule setsum_mono)(simp add: assms(2)) + also have "setsum f {a} < setsum g {a}" using a by simp + also have "setsum g (A - {a}) + setsum g {a} = setsum g((A-{a}) \ {a})" + using `finite A` by(subst setsum_Un_disjoint[symmetric]) auto + also have "\ = setsum g A" by(simp add:insert_absorb[OF `a:A`]) + finally show ?thesis by (metis add_right_mono add_strict_left_mono) +qed + +lemma measure_st: assumes "(strict{(x,y::'a::SL_top). x \ y})^-1 <= measure m" +and "\x y::'a. x \ y \ y \ x \ m x = m y" +shows "(strict{(S,S'::'a st). S \ S'})^-1 \ + measure(%fd. \x| x\set(dom fd) \ ~ \ \ fun fd x. m(fun fd x)+1)" +proof- + { fix S S' :: "'a st" assume "S \ S'" "~ S' \ S" + let ?X = "set(dom S)" let ?Y = "set(dom S')" + let ?f = "fun S" let ?g = "fun S'" + let ?X' = "{x:?X. ~ \ \ ?f x}" let ?Y' = "{y:?Y. ~ \ \ ?g y}" + from `S \ S'` have "ALL y:?Y'\?X. ?f y \ ?g y" + by(auto simp: le_st_def lookup_def) + hence 1: "ALL y:?Y'\?X. m(?g y)+1 \ m(?f y)+1" + using assms(1,2) by(fastforce) + from `~ S' \ S` obtain u where u: "u : ?X" "~ lookup S' u \ ?f u" + by(auto simp: le_st_def) + hence "u : ?X'" by simp (metis preord_class.le_trans top) + have "?Y'-?X = {}" using `S \ S'` by(fastforce simp: le_st_def lookup_def) + have "?Y'\?X <= ?X'" apply auto + apply (metis `S \ S'` le_st_def lookup_def preord_class.le_trans) + done + have "(\y\?Y'. m(?g y)+1) = (\y\(?Y'-?X) \ (?Y'\?X). m(?g y)+1)" + by (metis Un_Diff_Int) + also have "\ = (\y\?Y'\?X. m(?g y)+1)" + using `?Y'-?X = {}` by (metis Un_empty_left) + also have "\ < (\x\?X'. m(?f x)+1)" + proof cases + assume "u \ ?Y'" + hence "m(?g u) < m(?f u)" using assms(1) `S \ S'` u + by (fastforce simp: le_st_def lookup_def) + have "(\y\?Y'\?X. m(?g y)+1) < (\y\?Y'\?X. m(?f y)+1)" + using `u:?X` `u:?Y'` `m(?g u) < m(?f u)` + by(fastforce intro!: setsum_strict_mono1[OF _ 1]) + also have "\ \ (\y\?X'. m(?f y)+1)" + by(simp add: setsum_mono3[OF _ `?Y'\?X <= ?X'`]) + finally show ?thesis . + next + assume "u \ ?Y'" + with `?Y'\?X <= ?X'` have "?Y'\?X - {u} <= ?X' - {u}" by blast + have "(\y\?Y'\?X. m(?g y)+1) = (\y\?Y'\?X - {u}. m(?g y)+1)" + proof- + have "?Y'\?X = ?Y'\?X - {u}" using `u \ ?Y'` by auto + thus ?thesis by metis + qed + also have "\ < (\y\?Y'\?X-{u}. m(?g y)+1) + (\y\{u}. m(?f y)+1)" by simp + also have "(\y\?Y'\?X-{u}. m(?g y)+1) \ (\y\?Y'\?X-{u}. m(?f y)+1)" + using 1 by(blast intro: setsum_mono) + also have "\ \ (\y\?X'-{u}. m(?f y)+1)" + by(simp add: setsum_mono3[OF _ `?Y'\?X-{u} <= ?X'-{u}`]) + also have "\ + (\y\{u}. m(?f y)+1)= (\y\(?X'-{u}) \ {u}. m(?f y)+1)" + using `u:?X'` by(subst setsum_Un_disjoint[symmetric]) auto + also have "\ = (\x\?X'. m(?f x)+1)" + using `u : ?X'` by(simp add:insert_absorb) + finally show ?thesis by (blast intro: add_right_mono) + qed + finally have "(\y\?Y'. m(?g y)+1) < (\x\?X'. m(?f x)+1)" . + } thus ?thesis by(auto simp add: measure_def inv_image_def) +qed + +text{* ACC for acom. First the ordering on acom is related to an ordering on +lists of annotations. *} + +(* FIXME mv and add [simp] *) +lemma listrel_Cons_iff: + "(x#xs, y#ys) : listrel r \ (x,y) \ r \ (xs,ys) \ listrel r" +by (blast intro:listrel.Cons) + +lemma listrel_app: "(xs1,ys1) : listrel r \ (xs2,ys2) : listrel r + \ (xs1@xs2, ys1@ys2) : listrel r" +by(auto simp add: listrel_iff_zip) + +lemma listrel_app_same_size: "size xs1 = size ys1 \ size xs2 = size ys2 \ + (xs1@xs2, ys1@ys2) : listrel r \ + (xs1,ys1) : listrel r \ (xs2,ys2) : listrel r" +by(auto simp add: listrel_iff_zip) + +lemma listrel_converse: "listrel(r^-1) = (listrel r)^-1" +proof- + { fix xs ys + have "(xs,ys) : listrel(r^-1) \ (ys,xs) : listrel r" + apply(induct xs arbitrary: ys) + apply (fastforce simp: listrel.Nil) + apply (fastforce simp: listrel_Cons_iff) + done + } thus ?thesis by auto +qed + +(* It would be nice to get rid of refl & trans and build them into the proof *) +lemma acc_listrel: fixes r :: "('a*'a)set" assumes "refl r" and "trans r" +and "acc r" shows "acc (listrel r - {([],[])})" +proof- + have refl: "!!x. (x,x) : r" using `refl r` unfolding refl_on_def by blast + have trans: "!!x y z. (x,y) : r \ (y,z) : r \ (x,z) : r" + using `trans r` unfolding trans_def by blast + from assms(3) obtain mx :: "'a set \ 'a" where + mx: "!!S x. x:S \ mx S : S \ (\y. (mx S,y) : strict r \ y \ S)" + by(simp add: wf_eq_minimal) metis + let ?R = "listrel r - {([], [])}" + { fix Q and xs :: "'a list" + have "xs \ Q \ \ys. ys\Q \ (\zs. (ys, zs) \ strict ?R \ zs \ Q)" + (is "_ \ \ys. ?P Q ys") + proof(induction xs arbitrary: Q rule: length_induct) + case (1 xs) + { have "!!ys Q. size ys < size xs \ ys : Q \ EX ms. ?P Q ms" + using "1.IH" by blast + } note IH = this + show ?case + proof(cases xs) + case Nil with `xs : Q` have "?P Q []" by auto + thus ?thesis by blast + next + case (Cons x ys) + let ?Q1 = "{a. \bs. size bs = size ys \ a#bs : Q}" + have "x : ?Q1" using `xs : Q` Cons by auto + from mx[OF this] obtain m1 where + 1: "m1 \ ?Q1 \ (\y. (m1,y) \ strict r \ y \ ?Q1)" by blast + then obtain ms1 where "size ms1 = size ys" "m1#ms1 : Q" by blast+ + hence "size ms1 < size xs" using Cons by auto + let ?Q2 = "{bs. \m1'. (m1',m1):r \ (m1,m1'):r \ m1'#bs : Q \ size bs = size ms1}" + have "ms1 : ?Q2" using `m1#ms1 : Q` by(blast intro: refl) + from IH[OF `size ms1 < size xs` this] + obtain ms where 2: "?P ?Q2 ms" by auto + then obtain m1' where m1': "(m1',m1) : r \ (m1,m1') : r \ m1'#ms : Q" + by blast + hence "\ab. (m1'#ms,ab) : strict ?R \ ab \ Q" using 1 2 + apply (auto simp: listrel_Cons_iff) + apply (metis `length ms1 = length ys` listrel_eq_len trans) + by (metis `length ms1 = length ys` listrel_eq_len trans) + with m1' show ?thesis by blast + qed + qed + } + thus ?thesis unfolding wf_eq_minimal by (metis converse_iff) +qed + + +fun annos :: "'a acom \ 'a list" where +"annos (SKIP {a}) = [a]" | +"annos (x::=e {a}) = [a]" | +"annos (c1;c2) = annos c1 @ annos c2" | +"annos (IF b THEN c1 ELSE c2 {a}) = a # annos c1 @ annos c2" | +"annos ({i} WHILE b DO c {a}) = i # a # annos c" + +lemma size_annos_same: "strip c1 = strip c2 \ size(annos c1) = size(annos c2)" +apply(induct c2 arbitrary: c1) +apply (auto simp: strip_eq_SKIP strip_eq_Assign strip_eq_Semi strip_eq_If strip_eq_While) +done + +lemma le_iff_le_annos: "c1 \ c2 \ + (annos c1, annos c2) : listrel{(x,y). x \ y} \ strip c1 = strip c2" +apply(induct c1 c2 rule: le_acom.induct) +apply (auto simp: listrel.Nil listrel_Cons_iff listrel_app size_annos_same) +apply (metis listrel_app_same_size size_annos_same)+ +done + +lemma le_acom_subset_same_annos: + "(strict{(c,c'::'a::preord acom). c \ c'})^-1 \ + (strict(inv_image (listrel{(a,a'::'a). a \ a'} - {([],[])}) annos))^-1" +by(auto simp: le_iff_le_annos) + +lemma acc_acom: "acc {(a,a'::'a::preord). a \ a'} \ + acc {(c,c'::'a acom). c \ c'}" +apply(rule wf_subset[OF _ le_acom_subset_same_annos]) +apply(rule acc_inv_image[OF acc_listrel]) +apply(auto simp: refl_on_def trans_def intro: le_trans) +done + +text{* Termination of the fixed-point finders, assuming monotone functions: *} + +lemma pfp_termination: +fixes x0 :: "'a::preord" +assumes mono: "\x y. x \ y \ f x \ f y" and "acc {(x::'a,y). x \ y}" +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. x \ f x"]) + show "wf {(x, s). (s \ f s \ \ f s \ s) \ x = f s}" + by(rule wf_subset[OF assms(2)]) auto +next + show "x0 \ f x0" by(rule assms) +next + fix x assume "x \ f x" thus "f x \ f(f x)" by(rule mono) +qed + +lemma lpfpc_termination: + fixes f :: "(('a::SL_top)option acom \ 'a option acom)" + assumes "acc {(x::'a,y). x \ y}" and "\x y. x \ y \ f x \ f y" + and "\c. strip(f c) = strip c" + shows "\c'. lpfp\<^isub>c f c = Some c'" +unfolding lpfp\<^isub>c_def +apply(rule pfp_termination) + apply(erule assms(2)) + apply(rule acc_acom[OF acc_option[OF assms(1)]]) +apply(simp add: bot_acom assms(3)) +done + + end diff -r 3d518b508bbb -r 8b5f1f91ef38 src/HOL/IMP/Abs_Int0_const.thy --- a/src/HOL/IMP/Abs_Int0_const.thy Mon Jan 09 11:41:38 2012 +0100 +++ b/src/HOL/IMP/Abs_Int0_const.thy Mon Jan 09 13:48:14 2012 +0100 @@ -6,29 +6,29 @@ subsection "Constant Propagation" -datatype cval = Const val | Any +datatype const = Const val | Any -fun \_cval where -"\_cval (Const n) = {n}" | -"\_cval (Any) = UNIV" +fun \_const where +"\_const (Const n) = {n}" | +"\_const (Any) = UNIV" -fun plus_cval where -"plus_cval (Const m) (Const n) = Const(m+n)" | -"plus_cval _ _ = Any" +fun plus_const where +"plus_const (Const m) (Const n) = Const(m+n)" | +"plus_const _ _ = Any" -lemma plus_cval_cases: "plus_cval a1 a2 = +lemma plus_const_cases: "plus_const a1 a2 = (case (a1,a2) of (Const m, Const n) \ Const(m+n) | _ \ Any)" -by(auto split: prod.split cval.split) +by(auto split: prod.split const.split) -instantiation cval :: SL_top +instantiation const :: SL_top begin -fun le_cval where +fun le_const where "_ \ Any = True" | "Const n \ Const m = (n=m)" | "Any \ Const _ = False" -fun join_cval where +fun join_const where "Const m \ Const n = (if n=m then Const m else Any)" | "_ \ _ = Any" @@ -46,29 +46,29 @@ next case goal5 thus ?case by(cases z, cases y, cases x, simp_all) next - case goal6 thus ?case by(simp add: Top_cval_def) + case goal6 thus ?case by(simp add: Top_const_def) qed end interpretation Val_abs -where \ = \_cval and num' = Const and plus' = plus_cval +where \ = \_const and num' = Const and plus' = plus_const defines aval'_const is aval' proof case goal1 thus ?case by(cases a, cases b, simp, simp, cases b, simp, simp) next - case goal2 show ?case by(simp add: Top_cval_def) + case goal2 show ?case by(simp add: Top_const_def) next case goal3 show ?case by simp next case goal4 thus ?case - by(auto simp: plus_cval_cases split: cval.split) + by(auto simp: plus_const_cases split: const.split) qed interpretation Abs_Int -where \ = \_cval and num' = Const and plus' = plus_cval +where \ = \_const and num' = Const and plus' = plus_const defines AI_const is AI and step_const is step' proof qed @@ -77,12 +77,30 @@ text{* Monotonicity: *} interpretation Abs_Int_mono -where \ = \_cval and num' = Const and plus' = plus_cval +where \ = \_const and num' = Const and plus' = plus_const proof case goal1 thus ?case - by(auto simp: plus_cval_cases split: cval.split) + by(auto simp: plus_const_cases split: const.split) qed +text{* Termination: *} + +lemma measure_const: + "(strict{(x::const,y). x \ y})^-1 \ + measure(%x. case x of Const _ \ 1 | Any \ 0)" +by(auto split: const.splits) + +lemma measure_const_eq: + "\ x y::const. x \ y \ y \ x \ (%x. case x of Const _ \ 1 | Any \ 0) x = (%x. case x of Const _ \ 1 | Any \ 0) y" +by(auto split: const.splits) + +lemma acc_const_st: "Abs_Int0.acc{(x::const st,y). x \ y}" +by(rule wf_subset[OF wf_measure measure_st[OF measure_const measure_const_eq]]) + +lemma "EX c'. AI_const c = Some c'" +by(metis AI_def lpfpc_termination[OF acc_const_st, where f = "step_const \", + OF mono_step'[OF le_refl] strip_step']) + subsubsection "Tests"