Added termination to IMP Abs_Int
authornipkow
Mon, 09 Jan 2012 13:48:14 +0100
changeset 470298b5f1f91ef38
parent 47028 3d518b508bbb
child 47030 05a82dd869ed
Added termination to IMP Abs_Int
src/HOL/IMP/Abs_Int0.thy
src/HOL/IMP/Abs_Int0_const.thy
     1.1 --- a/src/HOL/IMP/Abs_Int0.thy	Mon Jan 09 11:41:38 2012 +0100
     1.2 +++ b/src/HOL/IMP/Abs_Int0.thy	Mon Jan 09 13:48:14 2012 +0100
     1.3 @@ -132,4 +132,260 @@
     1.4  
     1.5  end
     1.6  
     1.7 +
     1.8 +subsubsection "Ascending Chain Condition"
     1.9 +
    1.10 +hide_const acc
    1.11 +
    1.12 +abbreviation "strict r == r \<inter> -(r^-1)"
    1.13 +abbreviation "acc r == wf((strict r)^-1)"
    1.14 +
    1.15 +lemma strict_inv_image: "strict(inv_image r f) = inv_image (strict r) f"
    1.16 +by(auto simp: inv_image_def)
    1.17 +
    1.18 +lemma acc_inv_image:
    1.19 +  "acc r \<Longrightarrow> acc (inv_image r f)"
    1.20 +by (metis converse_inv_image strict_inv_image wf_inv_image)
    1.21 +
    1.22 +text{* ACC for option type: *}
    1.23 +
    1.24 +lemma acc_option: assumes "acc {(x,y::'a::preord). x \<sqsubseteq> y}"
    1.25 +shows "acc {(x,y::'a option). x \<sqsubseteq> y}"
    1.26 +proof(auto simp: wf_eq_minimal)
    1.27 +  fix xo :: "'a option" and Qo assume "xo : Qo"
    1.28 +  let ?Q = "{x. Some x \<in> Qo}"
    1.29 +  show "\<exists>yo\<in>Qo. \<forall>zo. yo \<sqsubseteq> zo \<and> ~ zo \<sqsubseteq> yo \<longrightarrow> zo \<notin> Qo" (is "\<exists>zo\<in>Qo. ?P zo")
    1.30 +  proof cases
    1.31 +    assume "?Q = {}"
    1.32 +    hence "?P None" by auto
    1.33 +    moreover have "None \<in> Qo" using `?Q = {}` `xo : Qo`
    1.34 +      by auto (metis not_Some_eq)
    1.35 +    ultimately show ?thesis by blast
    1.36 +  next
    1.37 +    assume "?Q \<noteq> {}"
    1.38 +    with assms show ?thesis
    1.39 +      apply(auto simp: wf_eq_minimal)
    1.40 +      apply(erule_tac x="?Q" in allE)
    1.41 +      apply auto
    1.42 +      apply(rule_tac x = "Some z" in bexI)
    1.43 +      by auto
    1.44 +  qed
    1.45 +qed
    1.46 +
    1.47 +text{* ACC for abstract states, via measure functions. *}
    1.48 +
    1.49 +(*FIXME mv*)
    1.50 +lemma setsum_strict_mono1:
    1.51 +fixes f :: "'a \<Rightarrow> 'b::{comm_monoid_add, ordered_cancel_ab_semigroup_add}"
    1.52 +assumes "finite A" and "ALL x:A. f x \<le> g x" and "EX a:A. f a < g a"
    1.53 +shows "setsum f A < setsum g A"
    1.54 +proof-
    1.55 +  from assms(3) obtain a where a: "a:A" "f a < g a" by blast
    1.56 +  have "setsum f A = setsum f ((A-{a}) \<union> {a})"
    1.57 +    by(simp add:insert_absorb[OF `a:A`])
    1.58 +  also have "\<dots> = setsum f (A-{a}) + setsum f {a}"
    1.59 +    using `finite A` by(subst setsum_Un_disjoint) auto
    1.60 +  also have "setsum f (A-{a}) \<le> setsum g (A-{a})"
    1.61 +    by(rule setsum_mono)(simp add: assms(2))
    1.62 +  also have "setsum f {a} < setsum g {a}" using a by simp
    1.63 +  also have "setsum g (A - {a}) + setsum g {a} = setsum g((A-{a}) \<union> {a})"
    1.64 +    using `finite A` by(subst setsum_Un_disjoint[symmetric]) auto
    1.65 +  also have "\<dots> = setsum g A" by(simp add:insert_absorb[OF `a:A`])
    1.66 +  finally show ?thesis by (metis add_right_mono add_strict_left_mono)
    1.67 +qed
    1.68 +
    1.69 +lemma measure_st: assumes "(strict{(x,y::'a::SL_top). x \<sqsubseteq> y})^-1 <= measure m"
    1.70 +and "\<forall>x y::'a. x \<sqsubseteq> y \<and> y \<sqsubseteq> x \<longrightarrow> m x = m y"
    1.71 +shows "(strict{(S,S'::'a st). S \<sqsubseteq> S'})^-1 \<subseteq>
    1.72 +  measure(%fd. \<Sum>x| x\<in>set(dom fd) \<and> ~ \<top> \<sqsubseteq> fun fd x. m(fun fd x)+1)"
    1.73 +proof-
    1.74 +  { fix S S' :: "'a st" assume "S \<sqsubseteq> S'" "~ S' \<sqsubseteq> S"
    1.75 +    let ?X = "set(dom S)" let ?Y = "set(dom S')"
    1.76 +    let ?f = "fun S" let ?g = "fun S'"
    1.77 +    let ?X' = "{x:?X. ~ \<top> \<sqsubseteq> ?f x}" let ?Y' = "{y:?Y. ~ \<top> \<sqsubseteq> ?g y}"
    1.78 +    from `S \<sqsubseteq> S'` have "ALL y:?Y'\<inter>?X. ?f y \<sqsubseteq> ?g y"
    1.79 +      by(auto simp: le_st_def lookup_def)
    1.80 +    hence 1: "ALL y:?Y'\<inter>?X. m(?g y)+1 \<le> m(?f y)+1"
    1.81 +      using assms(1,2) by(fastforce)
    1.82 +    from `~ S' \<sqsubseteq> S` obtain u where u: "u : ?X" "~ lookup S' u \<sqsubseteq> ?f u"
    1.83 +      by(auto simp: le_st_def)
    1.84 +    hence "u : ?X'" by simp (metis preord_class.le_trans top)
    1.85 +    have "?Y'-?X = {}" using `S \<sqsubseteq> S'` by(fastforce simp: le_st_def lookup_def)
    1.86 +    have "?Y'\<inter>?X <= ?X'" apply auto
    1.87 +      apply (metis `S \<sqsubseteq> S'` le_st_def lookup_def preord_class.le_trans)
    1.88 +      done
    1.89 +    have "(\<Sum>y\<in>?Y'. m(?g y)+1) = (\<Sum>y\<in>(?Y'-?X) \<union> (?Y'\<inter>?X). m(?g y)+1)"
    1.90 +      by (metis Un_Diff_Int)
    1.91 +    also have "\<dots> = (\<Sum>y\<in>?Y'\<inter>?X. m(?g y)+1)"
    1.92 +      using `?Y'-?X = {}` by (metis Un_empty_left)
    1.93 +    also have "\<dots> < (\<Sum>x\<in>?X'. m(?f x)+1)"
    1.94 +    proof cases
    1.95 +      assume "u \<in> ?Y'"
    1.96 +      hence "m(?g u) < m(?f u)" using assms(1) `S \<sqsubseteq> S'` u
    1.97 +        by (fastforce simp: le_st_def lookup_def)
    1.98 +      have "(\<Sum>y\<in>?Y'\<inter>?X. m(?g y)+1) < (\<Sum>y\<in>?Y'\<inter>?X. m(?f y)+1)"
    1.99 +        using `u:?X` `u:?Y'` `m(?g u) < m(?f u)`
   1.100 +        by(fastforce intro!: setsum_strict_mono1[OF _ 1])
   1.101 +      also have "\<dots> \<le> (\<Sum>y\<in>?X'. m(?f y)+1)"
   1.102 +        by(simp add: setsum_mono3[OF _ `?Y'\<inter>?X <= ?X'`])
   1.103 +      finally show ?thesis .
   1.104 +    next
   1.105 +      assume "u \<notin> ?Y'"
   1.106 +      with `?Y'\<inter>?X <= ?X'` have "?Y'\<inter>?X - {u} <= ?X' - {u}" by blast
   1.107 +      have "(\<Sum>y\<in>?Y'\<inter>?X. m(?g y)+1) = (\<Sum>y\<in>?Y'\<inter>?X - {u}. m(?g y)+1)"
   1.108 +      proof-
   1.109 +        have "?Y'\<inter>?X = ?Y'\<inter>?X - {u}" using `u \<notin> ?Y'` by auto
   1.110 +        thus ?thesis by metis
   1.111 +      qed
   1.112 +      also have "\<dots> < (\<Sum>y\<in>?Y'\<inter>?X-{u}. m(?g y)+1) + (\<Sum>y\<in>{u}. m(?f y)+1)" by simp
   1.113 +      also have "(\<Sum>y\<in>?Y'\<inter>?X-{u}. m(?g y)+1) \<le> (\<Sum>y\<in>?Y'\<inter>?X-{u}. m(?f y)+1)"
   1.114 +        using 1 by(blast intro: setsum_mono)
   1.115 +      also have "\<dots> \<le> (\<Sum>y\<in>?X'-{u}. m(?f y)+1)"
   1.116 +        by(simp add: setsum_mono3[OF _ `?Y'\<inter>?X-{u} <= ?X'-{u}`])
   1.117 +      also have "\<dots> + (\<Sum>y\<in>{u}. m(?f y)+1)= (\<Sum>y\<in>(?X'-{u}) \<union> {u}. m(?f y)+1)"
   1.118 +        using `u:?X'` by(subst setsum_Un_disjoint[symmetric]) auto
   1.119 +      also have "\<dots> = (\<Sum>x\<in>?X'. m(?f x)+1)"
   1.120 +        using `u : ?X'` by(simp add:insert_absorb)
   1.121 +      finally show ?thesis by (blast intro: add_right_mono)
   1.122 +    qed
   1.123 +    finally have "(\<Sum>y\<in>?Y'. m(?g y)+1) < (\<Sum>x\<in>?X'. m(?f x)+1)" .
   1.124 +  } thus ?thesis by(auto simp add: measure_def inv_image_def)
   1.125 +qed
   1.126 +
   1.127 +text{* ACC for acom. First the ordering on acom is related to an ordering on
   1.128 +lists of annotations. *}
   1.129 +
   1.130 +(* FIXME mv and add [simp] *)
   1.131 +lemma listrel_Cons_iff:
   1.132 +  "(x#xs, y#ys) : listrel r \<longleftrightarrow> (x,y) \<in> r \<and> (xs,ys) \<in> listrel r"
   1.133 +by (blast intro:listrel.Cons)
   1.134 +
   1.135 +lemma listrel_app: "(xs1,ys1) : listrel r \<Longrightarrow> (xs2,ys2) : listrel r
   1.136 +  \<Longrightarrow> (xs1@xs2, ys1@ys2) : listrel r"
   1.137 +by(auto simp add: listrel_iff_zip)
   1.138 +
   1.139 +lemma listrel_app_same_size: "size xs1 = size ys1 \<Longrightarrow> size xs2 = size ys2 \<Longrightarrow>
   1.140 +  (xs1@xs2, ys1@ys2) : listrel r \<longleftrightarrow>
   1.141 +  (xs1,ys1) : listrel r \<and> (xs2,ys2) : listrel r"
   1.142 +by(auto simp add: listrel_iff_zip)
   1.143 +
   1.144 +lemma listrel_converse: "listrel(r^-1) = (listrel r)^-1"
   1.145 +proof-
   1.146 +  { fix xs ys
   1.147 +    have "(xs,ys) : listrel(r^-1) \<longleftrightarrow> (ys,xs) : listrel r"
   1.148 +      apply(induct xs arbitrary: ys)
   1.149 +      apply (fastforce simp: listrel.Nil)
   1.150 +      apply (fastforce simp: listrel_Cons_iff)
   1.151 +      done
   1.152 +  } thus ?thesis by auto
   1.153 +qed
   1.154 +
   1.155 +(* It would be nice to get rid of refl & trans and build them into the proof *)
   1.156 +lemma acc_listrel: fixes r :: "('a*'a)set" assumes "refl r" and "trans r"
   1.157 +and "acc r" shows "acc (listrel r - {([],[])})"
   1.158 +proof-
   1.159 +  have refl: "!!x. (x,x) : r" using `refl r` unfolding refl_on_def by blast
   1.160 +  have trans: "!!x y z. (x,y) : r \<Longrightarrow> (y,z) : r \<Longrightarrow> (x,z) : r"
   1.161 +    using `trans r` unfolding trans_def by blast
   1.162 +  from assms(3) obtain mx :: "'a set \<Rightarrow> 'a" where
   1.163 +    mx: "!!S x. x:S \<Longrightarrow> mx S : S \<and> (\<forall>y. (mx S,y) : strict r \<longrightarrow> y \<notin> S)"
   1.164 +    by(simp add: wf_eq_minimal) metis
   1.165 +  let ?R = "listrel r - {([], [])}"
   1.166 +  { fix Q and xs :: "'a list"
   1.167 +    have "xs \<in> Q \<Longrightarrow> \<exists>ys. ys\<in>Q \<and> (\<forall>zs. (ys, zs) \<in> strict ?R \<longrightarrow> zs \<notin> Q)"
   1.168 +      (is "_ \<Longrightarrow> \<exists>ys. ?P Q ys")
   1.169 +    proof(induction xs arbitrary: Q rule: length_induct)
   1.170 +      case (1 xs)
   1.171 +      { have "!!ys Q. size ys < size xs \<Longrightarrow> ys : Q \<Longrightarrow> EX ms. ?P Q ms"
   1.172 +          using "1.IH" by blast
   1.173 +      } note IH = this
   1.174 +      show ?case
   1.175 +      proof(cases xs)
   1.176 +        case Nil with `xs : Q` have "?P Q []" by auto
   1.177 +        thus ?thesis by blast
   1.178 +      next
   1.179 +        case (Cons x ys)
   1.180 +        let ?Q1 = "{a. \<exists>bs. size bs = size ys \<and> a#bs : Q}"
   1.181 +        have "x : ?Q1" using `xs : Q` Cons by auto
   1.182 +        from mx[OF this] obtain m1 where
   1.183 +          1: "m1 \<in> ?Q1 \<and> (\<forall>y. (m1,y) \<in> strict r \<longrightarrow> y \<notin> ?Q1)" by blast
   1.184 +        then obtain ms1 where "size ms1 = size ys" "m1#ms1 : Q" by blast+
   1.185 +        hence "size ms1 < size xs" using Cons by auto
   1.186 +        let ?Q2 = "{bs. \<exists>m1'. (m1',m1):r \<and> (m1,m1'):r \<and> m1'#bs : Q \<and> size bs = size ms1}"
   1.187 +        have "ms1 : ?Q2" using `m1#ms1 : Q` by(blast intro: refl)
   1.188 +        from IH[OF `size ms1 < size xs` this]
   1.189 +        obtain ms where 2: "?P ?Q2 ms" by auto
   1.190 +        then obtain m1' where m1': "(m1',m1) : r \<and> (m1,m1') : r \<and> m1'#ms : Q"
   1.191 +          by blast
   1.192 +        hence "\<forall>ab. (m1'#ms,ab) : strict ?R \<longrightarrow> ab \<notin> Q" using 1 2
   1.193 +          apply (auto simp: listrel_Cons_iff)
   1.194 +          apply (metis `length ms1 = length ys` listrel_eq_len trans)
   1.195 +          by (metis `length ms1 = length ys` listrel_eq_len trans)
   1.196 +        with m1' show ?thesis by blast
   1.197 +      qed
   1.198 +    qed
   1.199 +  }
   1.200 +  thus ?thesis unfolding wf_eq_minimal by (metis converse_iff)
   1.201 +qed
   1.202 +
   1.203 +
   1.204 +fun annos :: "'a acom \<Rightarrow> 'a list" where
   1.205 +"annos (SKIP {a}) = [a]" |
   1.206 +"annos (x::=e {a}) = [a]" |
   1.207 +"annos (c1;c2) = annos c1 @ annos c2" |
   1.208 +"annos (IF b THEN c1 ELSE c2 {a}) = a #  annos c1 @ annos c2" |
   1.209 +"annos ({i} WHILE b DO c {a}) = i # a # annos c"
   1.210 +
   1.211 +lemma size_annos_same: "strip c1 = strip c2 \<Longrightarrow> size(annos c1) = size(annos c2)"
   1.212 +apply(induct c2 arbitrary: c1)
   1.213 +apply (auto simp: strip_eq_SKIP strip_eq_Assign strip_eq_Semi strip_eq_If strip_eq_While)
   1.214 +done
   1.215 +
   1.216 +lemma le_iff_le_annos: "c1 \<sqsubseteq> c2 \<longleftrightarrow>
   1.217 + (annos c1, annos c2) : listrel{(x,y). x \<sqsubseteq> y} \<and> strip c1 = strip c2"
   1.218 +apply(induct c1 c2 rule: le_acom.induct)
   1.219 +apply (auto simp: listrel.Nil listrel_Cons_iff listrel_app size_annos_same)
   1.220 +apply (metis listrel_app_same_size size_annos_same)+
   1.221 +done
   1.222 +
   1.223 +lemma le_acom_subset_same_annos:
   1.224 + "(strict{(c,c'::'a::preord acom). c \<sqsubseteq> c'})^-1 \<subseteq>
   1.225 +  (strict(inv_image (listrel{(a,a'::'a). a \<sqsubseteq> a'} - {([],[])}) annos))^-1"
   1.226 +by(auto simp: le_iff_le_annos)
   1.227 +
   1.228 +lemma acc_acom: "acc {(a,a'::'a::preord). a \<sqsubseteq> a'} \<Longrightarrow>
   1.229 +  acc {(c,c'::'a acom). c \<sqsubseteq> c'}"
   1.230 +apply(rule wf_subset[OF _ le_acom_subset_same_annos])
   1.231 +apply(rule acc_inv_image[OF acc_listrel])
   1.232 +apply(auto simp: refl_on_def trans_def intro: le_trans)
   1.233 +done
   1.234 +
   1.235 +text{* Termination of the fixed-point finders, assuming monotone functions: *}
   1.236 +
   1.237 +lemma pfp_termination:
   1.238 +fixes x0 :: "'a::preord"
   1.239 +assumes mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" and "acc {(x::'a,y). x \<sqsubseteq> y}"
   1.240 +and "x0 \<sqsubseteq> f x0" shows "EX x. pfp f x0 = Some x"
   1.241 +proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. x \<sqsubseteq> f x"])
   1.242 +  show "wf {(x, s). (s \<sqsubseteq> f s \<and> \<not> f s \<sqsubseteq> s) \<and> x = f s}"
   1.243 +    by(rule wf_subset[OF assms(2)]) auto
   1.244 +next
   1.245 +  show "x0 \<sqsubseteq> f x0" by(rule assms)
   1.246 +next
   1.247 +  fix x assume "x \<sqsubseteq> f x" thus "f x \<sqsubseteq> f(f x)" by(rule mono)
   1.248 +qed
   1.249 +
   1.250 +lemma lpfpc_termination:
   1.251 +  fixes f :: "(('a::SL_top)option acom \<Rightarrow> 'a option acom)"
   1.252 +  assumes "acc {(x::'a,y). x \<sqsubseteq> y}" and "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
   1.253 +  and "\<And>c. strip(f c) = strip c"
   1.254 +  shows "\<exists>c'. lpfp\<^isub>c f c = Some c'"
   1.255 +unfolding lpfp\<^isub>c_def
   1.256 +apply(rule pfp_termination)
   1.257 +  apply(erule assms(2))
   1.258 + apply(rule acc_acom[OF acc_option[OF assms(1)]])
   1.259 +apply(simp add: bot_acom assms(3))
   1.260 +done
   1.261 +
   1.262 +
   1.263  end
     2.1 --- a/src/HOL/IMP/Abs_Int0_const.thy	Mon Jan 09 11:41:38 2012 +0100
     2.2 +++ b/src/HOL/IMP/Abs_Int0_const.thy	Mon Jan 09 13:48:14 2012 +0100
     2.3 @@ -6,29 +6,29 @@
     2.4  
     2.5  subsection "Constant Propagation"
     2.6  
     2.7 -datatype cval = Const val | Any
     2.8 +datatype const = Const val | Any
     2.9  
    2.10 -fun \<gamma>_cval where
    2.11 -"\<gamma>_cval (Const n) = {n}" |
    2.12 -"\<gamma>_cval (Any) = UNIV"
    2.13 +fun \<gamma>_const where
    2.14 +"\<gamma>_const (Const n) = {n}" |
    2.15 +"\<gamma>_const (Any) = UNIV"
    2.16  
    2.17 -fun plus_cval where
    2.18 -"plus_cval (Const m) (Const n) = Const(m+n)" |
    2.19 -"plus_cval _ _ = Any"
    2.20 +fun plus_const where
    2.21 +"plus_const (Const m) (Const n) = Const(m+n)" |
    2.22 +"plus_const _ _ = Any"
    2.23  
    2.24 -lemma plus_cval_cases: "plus_cval a1 a2 =
    2.25 +lemma plus_const_cases: "plus_const a1 a2 =
    2.26    (case (a1,a2) of (Const m, Const n) \<Rightarrow> Const(m+n) | _ \<Rightarrow> Any)"
    2.27 -by(auto split: prod.split cval.split)
    2.28 +by(auto split: prod.split const.split)
    2.29  
    2.30 -instantiation cval :: SL_top
    2.31 +instantiation const :: SL_top
    2.32  begin
    2.33  
    2.34 -fun le_cval where
    2.35 +fun le_const where
    2.36  "_ \<sqsubseteq> Any = True" |
    2.37  "Const n \<sqsubseteq> Const m = (n=m)" |
    2.38  "Any \<sqsubseteq> Const _ = False"
    2.39  
    2.40 -fun join_cval where
    2.41 +fun join_const where
    2.42  "Const m \<squnion> Const n = (if n=m then Const m else Any)" |
    2.43  "_ \<squnion> _ = Any"
    2.44  
    2.45 @@ -46,29 +46,29 @@
    2.46  next
    2.47    case goal5 thus ?case by(cases z, cases y, cases x, simp_all)
    2.48  next
    2.49 -  case goal6 thus ?case by(simp add: Top_cval_def)
    2.50 +  case goal6 thus ?case by(simp add: Top_const_def)
    2.51  qed
    2.52  
    2.53  end
    2.54  
    2.55  
    2.56  interpretation Val_abs
    2.57 -where \<gamma> = \<gamma>_cval and num' = Const and plus' = plus_cval
    2.58 +where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const
    2.59  defines aval'_const is aval'
    2.60  proof
    2.61    case goal1 thus ?case
    2.62      by(cases a, cases b, simp, simp, cases b, simp, simp)
    2.63  next
    2.64 -  case goal2 show ?case by(simp add: Top_cval_def)
    2.65 +  case goal2 show ?case by(simp add: Top_const_def)
    2.66  next
    2.67    case goal3 show ?case by simp
    2.68  next
    2.69    case goal4 thus ?case
    2.70 -    by(auto simp: plus_cval_cases split: cval.split)
    2.71 +    by(auto simp: plus_const_cases split: const.split)
    2.72  qed
    2.73  
    2.74  interpretation Abs_Int
    2.75 -where \<gamma> = \<gamma>_cval and num' = Const and plus' = plus_cval
    2.76 +where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const
    2.77  defines AI_const is AI
    2.78  and step_const is step'
    2.79  proof qed
    2.80 @@ -77,12 +77,30 @@
    2.81  text{* Monotonicity: *}
    2.82  
    2.83  interpretation Abs_Int_mono
    2.84 -where \<gamma> = \<gamma>_cval and num' = Const and plus' = plus_cval
    2.85 +where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const
    2.86  proof
    2.87    case goal1 thus ?case
    2.88 -    by(auto simp: plus_cval_cases split: cval.split)
    2.89 +    by(auto simp: plus_const_cases split: const.split)
    2.90  qed
    2.91  
    2.92 +text{* Termination: *}
    2.93 +
    2.94 +lemma measure_const:
    2.95 +  "(strict{(x::const,y). x \<sqsubseteq> y})^-1 \<subseteq>
    2.96 +  measure(%x. case x of Const _ \<Rightarrow> 1 | Any \<Rightarrow> 0)"
    2.97 +by(auto split: const.splits)
    2.98 +
    2.99 +lemma measure_const_eq:
   2.100 +  "\<forall> x y::const. x \<sqsubseteq> y \<and> y \<sqsubseteq> x \<longrightarrow> (%x. case x of Const _ \<Rightarrow> 1 | Any \<Rightarrow> 0) x = (%x. case x of Const _ \<Rightarrow> 1 | Any \<Rightarrow> 0) y"
   2.101 +by(auto split: const.splits)
   2.102 +
   2.103 +lemma acc_const_st: "Abs_Int0.acc{(x::const st,y). x \<sqsubseteq> y}"
   2.104 +by(rule wf_subset[OF wf_measure measure_st[OF measure_const measure_const_eq]])
   2.105 +
   2.106 +lemma "EX c'. AI_const c = Some c'"
   2.107 +by(metis AI_def lpfpc_termination[OF acc_const_st, where f = "step_const \<top>",
   2.108 +  OF mono_step'[OF le_refl] strip_step'])
   2.109 +
   2.110  
   2.111  subsubsection "Tests"
   2.112