reorganised IMP
authornipkow
Thu, 19 Apr 2012 17:32:30 +0200
changeset 47602 3d44790b5ab0
parent 47583 f3f0e06549c2
child 47603 b716b16ab2ac
reorganised IMP
src/HOL/IMP/Abs_Int0.thy
src/HOL/IMP/Abs_Int0_const.thy
src/HOL/IMP/Abs_Int0_fun.thy
src/HOL/IMP/Abs_Int0_parity.thy
src/HOL/IMP/Abs_Int1.thy
src/HOL/IMP/Abs_Int1_ivl.thy
src/HOL/IMP/Abs_Int2.thy
src/HOL/IMP/Abs_Int_ITP/Abs_Int0_ITP.thy
src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy
src/HOL/IMP/Abs_Int_ITP/Abs_Int1_const_ITP.thy
src/HOL/IMP/Abs_Int_ITP/Abs_Int1_parity_ITP.thy
src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ITP.thy
src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ivl_ITP.thy
src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy
src/HOL/IMP/Abs_Int_ITP/Abs_State_ITP.thy
src/HOL/IMP/Abs_State.thy
src/HOL/IMP/ROOT.ML
src/HOL/IMP/document/root.tex
src/HOL/IsaMakefile
--- a/src/HOL/IMP/Abs_Int0.thy	Thu Apr 19 12:28:10 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,411 +0,0 @@
-(* Author: Tobias Nipkow *)
-
-theory Abs_Int0
-imports Abs_State
-begin
-
-subsection "Computable Abstract Interpretation"
-
-text{* Abstract interpretation over type @{text st} instead of
-functions. *}
-
-context Gamma
-begin
-
-fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where
-"aval' (N n) S = num' n" |
-"aval' (V x) S = lookup S x" |
-"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
-
-lemma aval'_sound: "s : \<gamma>\<^isub>f S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
-by (induction a) (auto simp: gamma_num' gamma_plus' \<gamma>_st_def lookup_def)
-
-end
-
-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}. *}
-
-locale Abs_Int = Gamma where \<gamma>=\<gamma> for \<gamma> :: "'av::SL_top \<Rightarrow> val set"
-begin
-
-fun step' :: "'av st option \<Rightarrow> 'av st option acom \<Rightarrow> 'av st option acom" where
-"step' S (SKIP {P}) = (SKIP {S})" |
-"step' S (x ::= e {P}) =
-  x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(update S x (aval' e S))}" |
-"step' S (c1; c2) = step' S c1; step' (post c1) c2" |
-"step' S (IF b THEN c1 ELSE c2 {P}) =
-  (let c1' = step' S c1; c2' = step' S c2
-   in IF b THEN c1' ELSE c2' {post c1 \<squnion> post c2})" |
-"step' S ({Inv} WHILE b DO c {P}) =
-   {S \<squnion> post c} WHILE b DO step' Inv c {Inv}"
-
-definition AI :: "com \<Rightarrow> 'av st option acom option" where
-"AI = lpfp\<^isub>c (step' \<top>)"
-
-
-lemma strip_step'[simp]: "strip(step' S c) = strip c"
-by(induct c arbitrary: S) (simp_all add: Let_def)
-
-
-text{* Soundness: *}
-
-lemma in_gamma_update:
-  "\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(update S x a)"
-by(simp add: \<gamma>_st_def lookup_update)
-
-text{* The soundness proofs are textually identical to the ones for the step
-function operating on states as functions. *}
-
-lemma step_preserves_le:
-  "\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; c \<le> \<gamma>\<^isub>c c' \<rbrakk> \<Longrightarrow> step S c \<le> \<gamma>\<^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 intro: aval'_sound in_gamma_update
-      split: option.splits del:subsetD)
-next
-  case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi)
-    by (metis le_post post_map_acom)
-next
-  case (If b c1 c2 P)
-  then obtain c1' c2' P' where
-      "c' = IF b THEN c1' ELSE c2' {P'}"
-      "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'" "c2 \<le> \<gamma>\<^isub>c c2'"
-    by (fastforce simp: If_le map_acom_If)
-  moreover have "post c1 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
-    by (metis (no_types) `c1 \<le> \<gamma>\<^isub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
-  moreover have "post c2 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
-    by (metis (no_types) `c2 \<le> \<gamma>\<^isub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
-  ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'` by (simp add: If.IH subset_iff)
-next
-  case (While I b c1 P)
-  then obtain c1' I' P' where
-    "c' = {I'} WHILE b DO c1' {P'}"
-    "I \<subseteq> \<gamma>\<^isub>o I'" "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'"
-    by (fastforce simp: map_acom_While While_le)
-  moreover have "S \<union> post c1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post c1')"
-    using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `c1 \<le> \<gamma>\<^isub>c c1'`, simplified]
-    by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans)
-  ultimately show ?case by (simp add: While.IH subset_iff)
-qed
-
-lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
-proof(simp add: CS_def AI_def)
-  assume 1: "lpfp\<^isub>c (step' \<top>) c = Some c'"
-  have 2: "step' \<top> c' \<sqsubseteq> c'" by(rule lpfpc_pfp[OF 1])
-  have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c"
-    by(simp add: strip_lpfpc[OF _ 1])
-  have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
-  proof(rule lfp_lowerbound[simplified,OF 3])
-    show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
-    proof(rule step_preserves_le[OF _ _])
-      show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
-      show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
-    qed
-  qed
-  from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
-    by (blast intro: mono_gamma_c order_trans)
-qed
-
-end
-
-
-subsubsection "Monotonicity"
-
-locale Abs_Int_mono = Abs_Int +
-assumes mono_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> plus' a1 a2 \<sqsubseteq> plus' b1 b2"
-begin
-
-lemma mono_aval': "S \<sqsubseteq> S' \<Longrightarrow> aval' e S \<sqsubseteq> aval' e S'"
-by(induction e) (auto simp: le_st_def lookup_def mono_plus')
-
-lemma mono_update: "a \<sqsubseteq> a' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> update S x a \<sqsubseteq> update S' x a'"
-by(auto simp add: le_st_def lookup_def update_def)
-
-lemma mono_step': "S \<sqsubseteq> S' \<Longrightarrow> c \<sqsubseteq> c' \<Longrightarrow> step' S c \<sqsubseteq> step' S' c'"
-apply(induction c c' arbitrary: S S' rule: le_acom.induct)
-apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj
-            split: option.split)
-done
-
-end
-
-
-subsubsection "Ascending Chain Condition"
-
-hide_const (open) acc
-
-abbreviation "strict r == r \<inter> -(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 \<Longrightarrow> 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 \<sqsubseteq> y}"
-shows "acc {(x,y::'a::preord option). x \<sqsubseteq> y}"
-proof(auto simp: wf_eq_minimal)
-  fix xo :: "'a option" and Qo assume "xo : Qo"
-  let ?Q = "{x. Some x \<in> Qo}"
-  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")
-  proof cases
-    assume "?Q = {}"
-    hence "?P None" by auto
-    moreover have "None \<in> Qo" using `?Q = {}` `xo : Qo`
-      by auto (metis not_Some_eq)
-    ultimately show ?thesis by blast
-  next
-    assume "?Q \<noteq> {}"
-    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 \<Rightarrow> 'b::{comm_monoid_add, ordered_cancel_ab_semigroup_add}"
-assumes "finite A" and "ALL x:A. f x \<le> 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}) \<union> {a})"
-    by(simp add:insert_absorb[OF `a:A`])
-  also have "\<dots> = setsum f (A-{a}) + setsum f {a}"
-    using `finite A` by(subst setsum_Un_disjoint) auto
-  also have "setsum f (A-{a}) \<le> 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}) \<union> {a})"
-    using `finite A` by(subst setsum_Un_disjoint[symmetric]) auto
-  also have "\<dots> = 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 \<sqsubseteq> y})^-1 <= measure m"
-and "\<forall>x y::'a::SL_top. x \<sqsubseteq> y \<and> y \<sqsubseteq> x \<longrightarrow> m x = m y"
-shows "(strict{(S,S'::'a::SL_top st). S \<sqsubseteq> S'})^-1 \<subseteq>
-  measure(%fd. \<Sum>x| x\<in>set(dom fd) \<and> ~ \<top> \<sqsubseteq> fun fd x. m(fun fd x)+1)"
-proof-
-  { fix S S' :: "'a st" assume "S \<sqsubseteq> S'" "~ S' \<sqsubseteq> S"
-    let ?X = "set(dom S)" let ?Y = "set(dom S')"
-    let ?f = "fun S" let ?g = "fun S'"
-    let ?X' = "{x:?X. ~ \<top> \<sqsubseteq> ?f x}" let ?Y' = "{y:?Y. ~ \<top> \<sqsubseteq> ?g y}"
-    from `S \<sqsubseteq> S'` have "ALL y:?Y'\<inter>?X. ?f y \<sqsubseteq> ?g y"
-      by(auto simp: le_st_def lookup_def)
-    hence 1: "ALL y:?Y'\<inter>?X. m(?g y)+1 \<le> m(?f y)+1"
-      using assms(1,2) by(fastforce)
-    from `~ S' \<sqsubseteq> S` obtain u where u: "u : ?X" "~ lookup S' u \<sqsubseteq> ?f u"
-      by(auto simp: le_st_def)
-    hence "u : ?X'" by simp (metis preord_class.le_trans top)
-    have "?Y'-?X = {}" using `S \<sqsubseteq> S'` by(fastforce simp: le_st_def lookup_def)
-    have "?Y'\<inter>?X <= ?X'" apply auto
-      apply (metis `S \<sqsubseteq> S'` le_st_def lookup_def preord_class.le_trans)
-      done
-    have "(\<Sum>y\<in>?Y'. m(?g y)+1) = (\<Sum>y\<in>(?Y'-?X) \<union> (?Y'\<inter>?X). m(?g y)+1)"
-      by (metis Un_Diff_Int)
-    also have "\<dots> = (\<Sum>y\<in>?Y'\<inter>?X. m(?g y)+1)"
-      using `?Y'-?X = {}` by (metis Un_empty_left)
-    also have "\<dots> < (\<Sum>x\<in>?X'. m(?f x)+1)"
-    proof cases
-      assume "u \<in> ?Y'"
-      hence "m(?g u) < m(?f u)" using assms(1) `S \<sqsubseteq> S'` u
-        by (fastforce simp: le_st_def lookup_def)
-      have "(\<Sum>y\<in>?Y'\<inter>?X. m(?g y)+1) < (\<Sum>y\<in>?Y'\<inter>?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 "\<dots> \<le> (\<Sum>y\<in>?X'. m(?f y)+1)"
-        by(simp add: setsum_mono3[OF _ `?Y'\<inter>?X <= ?X'`])
-      finally show ?thesis .
-    next
-      assume "u \<notin> ?Y'"
-      with `?Y'\<inter>?X <= ?X'` have "?Y'\<inter>?X - {u} <= ?X' - {u}" by blast
-      have "(\<Sum>y\<in>?Y'\<inter>?X. m(?g y)+1) = (\<Sum>y\<in>?Y'\<inter>?X - {u}. m(?g y)+1)"
-      proof-
-        have "?Y'\<inter>?X = ?Y'\<inter>?X - {u}" using `u \<notin> ?Y'` by auto
-        thus ?thesis by metis
-      qed
-      also have "\<dots> < (\<Sum>y\<in>?Y'\<inter>?X-{u}. m(?g y)+1) + (\<Sum>y\<in>{u}. m(?f y)+1)" by simp
-      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)"
-        using 1 by(blast intro: setsum_mono)
-      also have "\<dots> \<le> (\<Sum>y\<in>?X'-{u}. m(?f y)+1)"
-        by(simp add: setsum_mono3[OF _ `?Y'\<inter>?X-{u} <= ?X'-{u}`])
-      also have "\<dots> + (\<Sum>y\<in>{u}. m(?f y)+1)= (\<Sum>y\<in>(?X'-{u}) \<union> {u}. m(?f y)+1)"
-        using `u:?X'` by(subst setsum_Un_disjoint[symmetric]) auto
-      also have "\<dots> = (\<Sum>x\<in>?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 "(\<Sum>y\<in>?Y'. m(?g y)+1) < (\<Sum>x\<in>?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 \<longleftrightarrow> (x,y) \<in> r \<and> (xs,ys) \<in> listrel r"
-by (blast intro:listrel.Cons)
-
-lemma listrel_app: "(xs1,ys1) : listrel r \<Longrightarrow> (xs2,ys2) : listrel r
-  \<Longrightarrow> (xs1@xs2, ys1@ys2) : listrel r"
-by(auto simp add: listrel_iff_zip)
-
-lemma listrel_app_same_size: "size xs1 = size ys1 \<Longrightarrow> size xs2 = size ys2 \<Longrightarrow>
-  (xs1@xs2, ys1@ys2) : listrel r \<longleftrightarrow>
-  (xs1,ys1) : listrel r \<and> (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) \<longleftrightarrow> (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 \<Longrightarrow> (y,z) : r \<Longrightarrow> (x,z) : r"
-    using `trans r` unfolding trans_def by blast
-  from assms(3) obtain mx :: "'a set \<Rightarrow> 'a" where
-    mx: "!!S x. x:S \<Longrightarrow> mx S : S \<and> (\<forall>y. (mx S,y) : strict r \<longrightarrow> y \<notin> S)"
-    by(simp add: wf_eq_minimal) metis
-  let ?R = "listrel r - {([], [])}"
-  { fix Q and xs :: "'a list"
-    have "xs \<in> Q \<Longrightarrow> \<exists>ys. ys\<in>Q \<and> (\<forall>zs. (ys, zs) \<in> strict ?R \<longrightarrow> zs \<notin> Q)"
-      (is "_ \<Longrightarrow> \<exists>ys. ?P Q ys")
-    proof(induction xs arbitrary: Q rule: length_induct)
-      case (1 xs)
-      { have "!!ys Q. size ys < size xs \<Longrightarrow> ys : Q \<Longrightarrow> 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. \<exists>bs. size bs = size ys \<and> a#bs : Q}"
-        have "x : ?Q1" using `xs : Q` Cons by auto
-        from mx[OF this] obtain m1 where
-          1: "m1 \<in> ?Q1 \<and> (\<forall>y. (m1,y) \<in> strict r \<longrightarrow> y \<notin> ?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. \<exists>m1'. (m1',m1):r \<and> (m1,m1'):r \<and> m1'#bs : Q \<and> 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 \<and> (m1,m1') : r \<and> m1'#ms : Q"
-          by blast
-        hence "\<forall>ab. (m1'#ms,ab) : strict ?R \<longrightarrow> ab \<notin> 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 \<Rightarrow> '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 \<Longrightarrow> 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
-
-lemmas size_annos_same2 = eqTrueI[OF size_annos_same]
-
-lemma set_annos_anno: "set (annos (anno a c)) = {a}"
-by(induction c)(auto)
-
-lemma le_iff_le_annos: "c1 \<sqsubseteq> c2 \<longleftrightarrow>
- (annos c1, annos c2) : listrel{(x,y). x \<sqsubseteq> y} \<and> strip c1 = strip c2"
-apply(induct c1 c2 rule: le_acom.induct)
-apply (auto simp: listrel.Nil listrel_Cons_iff listrel_app size_annos_same2)
-apply (metis listrel_app_same_size size_annos_same)+
-done
-
-lemma le_acom_subset_same_annos:
- "(strict{(c,c'::'a::preord acom). c \<sqsubseteq> c'})^-1 \<subseteq>
-  (strict(inv_image (listrel{(a,a'::'a). a \<sqsubseteq> a'} - {([],[])}) annos))^-1"
-by(auto simp: le_iff_le_annos)
-
-lemma acc_acom: "acc {(a,a'::'a::preord). a \<sqsubseteq> a'} \<Longrightarrow>
-  acc {(c,c'::'a acom). c \<sqsubseteq> 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: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" and "acc {(x::'a,y). x \<sqsubseteq> y}"
-and "x0 \<sqsubseteq> f x0" shows "EX x. pfp f x0 = Some x"
-proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. x \<sqsubseteq> f x"])
-  show "wf {(x, s). (s \<sqsubseteq> f s \<and> \<not> f s \<sqsubseteq> s) \<and> x = f s}"
-    by(rule wf_subset[OF assms(2)]) auto
-next
-  show "x0 \<sqsubseteq> f x0" by(rule assms)
-next
-  fix x assume "x \<sqsubseteq> f x" thus "f x \<sqsubseteq> f(f x)" by(rule mono)
-qed
-
-lemma lpfpc_termination:
-  fixes f :: "(('a::SL_top)option acom \<Rightarrow> 'a option acom)"
-  assumes "acc {(x::'a,y). x \<sqsubseteq> y}" and "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
-  and "\<And>c. strip(f c) = strip c"
-  shows "\<exists>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
-
-context Abs_Int_mono
-begin
-
-lemma AI_Some_measure:
-assumes "(strict{(x,y::'a). x \<sqsubseteq> y})^-1 <= measure m"
-and "\<forall>x y::'a. x \<sqsubseteq> y \<and> y \<sqsubseteq> x \<longrightarrow> m x = m y"
-shows "\<exists>c'. AI c = Some c'"
-unfolding AI_def
-apply(rule lpfpc_termination)
-apply(rule wf_subset[OF wf_measure measure_st[OF assms]])
-apply(erule mono_step'[OF le_refl])
-apply(rule strip_step')
-done
-
-end
-
-end
--- a/src/HOL/IMP/Abs_Int0_const.thy	Thu Apr 19 12:28:10 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,139 +0,0 @@
-(* Author: Tobias Nipkow *)
-
-theory Abs_Int0_const
-imports Abs_Int0 Abs_Int_Tests
-begin
-
-subsection "Constant Propagation"
-
-datatype const = Const val | Any
-
-fun \<gamma>_const where
-"\<gamma>_const (Const n) = {n}" |
-"\<gamma>_const (Any) = UNIV"
-
-fun plus_const where
-"plus_const (Const m) (Const n) = Const(m+n)" |
-"plus_const _ _ = Any"
-
-lemma plus_const_cases: "plus_const a1 a2 =
-  (case (a1,a2) of (Const m, Const n) \<Rightarrow> Const(m+n) | _ \<Rightarrow> Any)"
-by(auto split: prod.split const.split)
-
-instantiation const :: SL_top
-begin
-
-fun le_const where
-"_ \<sqsubseteq> Any = True" |
-"Const n \<sqsubseteq> Const m = (n=m)" |
-"Any \<sqsubseteq> Const _ = False"
-
-fun join_const where
-"Const m \<squnion> Const n = (if n=m then Const m else Any)" |
-"_ \<squnion> _ = Any"
-
-definition "\<top> = Any"
-
-instance
-proof
-  case goal1 thus ?case by (cases x) simp_all
-next
-  case goal2 thus ?case by(cases z, cases y, cases x, simp_all)
-next
-  case goal3 thus ?case by(cases x, cases y, simp_all)
-next
-  case goal4 thus ?case by(cases y, cases x, simp_all)
-next
-  case goal5 thus ?case by(cases z, cases y, cases x, simp_all)
-next
-  case goal6 thus ?case by(simp add: Top_const_def)
-qed
-
-end
-
-
-interpretation Val_abs
-where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const
-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_const_def)
-next
-  case goal3 show ?case by simp
-next
-  case goal4 thus ?case
-    by(auto simp: plus_const_cases split: const.split)
-qed
-
-interpretation Abs_Int
-where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const
-defines AI_const is AI and step_const is step' and aval'_const is aval'
-..
-
-
-subsubsection "Tests"
-
-value "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test1_const))"
-value "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test1_const))"
-value "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test1_const))"
-value "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test1_const))"
-value "show_acom_opt (AI_const test1_const)"
-
-value "show_acom_opt (AI_const test2_const)"
-value "show_acom_opt (AI_const test3_const)"
-
-value "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test4_const))"
-value "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test4_const))"
-value "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test4_const))"
-value "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test4_const))"
-value "show_acom_opt (AI_const test4_const)"
-
-value "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test5_const))"
-value "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test5_const))"
-value "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test5_const))"
-value "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test5_const))"
-value "show_acom (((step_const \<top>)^^4) (\<bottom>\<^sub>c test5_const))"
-value "show_acom (((step_const \<top>)^^5) (\<bottom>\<^sub>c test5_const))"
-value "show_acom_opt (AI_const test5_const)"
-
-value "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test6_const))"
-value "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test6_const))"
-value "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test6_const))"
-value "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test6_const))"
-value "show_acom (((step_const \<top>)^^4) (\<bottom>\<^sub>c test6_const))"
-value "show_acom (((step_const \<top>)^^5) (\<bottom>\<^sub>c test6_const))"
-value "show_acom (((step_const \<top>)^^6) (\<bottom>\<^sub>c test6_const))"
-value "show_acom (((step_const \<top>)^^7) (\<bottom>\<^sub>c test6_const))"
-value "show_acom (((step_const \<top>)^^8) (\<bottom>\<^sub>c test6_const))"
-value "show_acom (((step_const \<top>)^^9) (\<bottom>\<^sub>c test6_const))"
-value "show_acom (((step_const \<top>)^^10) (\<bottom>\<^sub>c test6_const))"
-value "show_acom (((step_const \<top>)^^11) (\<bottom>\<^sub>c test6_const))"
-value "show_acom_opt (AI_const test6_const)"
-
-
-text{* Monotonicity: *}
-
-interpretation Abs_Int_mono
-where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const
-proof
-  case goal1 thus ?case
-    by(auto simp: plus_const_cases split: const.split)
-qed
-
-text{* Termination: *}
-
-definition "m_const x = (case x of Const _ \<Rightarrow> 1 | Any \<Rightarrow> 0)"
-
-lemma measure_const:
-  "(strict{(x::const,y). x \<sqsubseteq> y})^-1 \<subseteq> measure m_const"
-by(auto simp: m_const_def split: const.splits)
-
-lemma measure_const_eq:
-  "\<forall> x y::const. x \<sqsubseteq> y \<and> y \<sqsubseteq> x \<longrightarrow> m_const x = m_const y"
-by(auto simp: m_const_def split: const.splits)
-
-lemma "EX c'. AI_const c = Some c'"
-by(rule AI_Some_measure[OF measure_const measure_const_eq])
-
-end
--- a/src/HOL/IMP/Abs_Int0_fun.thy	Thu Apr 19 12:28:10 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,397 +0,0 @@
-(* Author: Tobias Nipkow *)
-
-theory Abs_Int0_fun
-imports "~~/src/HOL/ex/Interpretation_with_Defs"
-        "~~/src/HOL/Library/While_Combinator"
-        Collecting
-begin
-
-subsection "Orderings"
-
-class preord =
-fixes le :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubseteq>" 50)
-assumes le_refl[simp]: "x \<sqsubseteq> x"
-and le_trans: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z"
-begin
-
-definition mono where "mono f = (\<forall>x y. x \<sqsubseteq> y \<longrightarrow> f x \<sqsubseteq> f y)"
-
-lemma monoD: "mono f \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" by(simp add: mono_def)
-
-lemma mono_comp: "mono f \<Longrightarrow> mono g \<Longrightarrow> mono (g o f)"
-by(simp add: mono_def)
-
-declare le_trans[trans]
-
-end
-
-text{* Note: no antisymmetry. Allows implementations where some abstract
-element is implemented by two different values @{prop "x \<noteq> y"}
-such that @{prop"x \<sqsubseteq> y"} and @{prop"y \<sqsubseteq> x"}. Antisymmetry is not
-needed because we never compare elements for equality but only for @{text"\<sqsubseteq>"}.
-*}
-
-class SL_top = preord +
-fixes join :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<squnion>" 65)
-fixes Top :: "'a" ("\<top>")
-assumes join_ge1 [simp]: "x \<sqsubseteq> x \<squnion> y"
-and join_ge2 [simp]: "y \<sqsubseteq> x \<squnion> y"
-and join_least: "x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<squnion> y \<sqsubseteq> z"
-and top[simp]: "x \<sqsubseteq> \<top>"
-begin
-
-lemma join_le_iff[simp]: "x \<squnion> y \<sqsubseteq> z \<longleftrightarrow> x \<sqsubseteq> z \<and> y \<sqsubseteq> z"
-by (metis join_ge1 join_ge2 join_least le_trans)
-
-lemma le_join_disj: "x \<sqsubseteq> y \<or> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<squnion> z"
-by (metis join_ge1 join_ge2 le_trans)
-
-end
-
-instantiation "fun" :: (type, SL_top) SL_top
-begin
-
-definition "f \<sqsubseteq> g = (ALL x. f x \<sqsubseteq> g x)"
-definition "f \<squnion> g = (\<lambda>x. f x \<squnion> g x)"
-definition "\<top> = (\<lambda>x. \<top>)"
-
-lemma join_apply[simp]: "(f \<squnion> g) x = f x \<squnion> g x"
-by (simp add: join_fun_def)
-
-instance
-proof
-  case goal2 thus ?case by (metis le_fun_def preord_class.le_trans)
-qed (simp_all add: le_fun_def Top_fun_def)
-
-end
-
-
-instantiation acom :: (preord) preord
-begin
-
-fun le_acom :: "('a::preord)acom \<Rightarrow> 'a acom \<Rightarrow> bool" where
-"le_acom (SKIP {S}) (SKIP {S'}) = (S \<sqsubseteq> S')" |
-"le_acom (x ::= e {S}) (x' ::= e' {S'}) = (x=x' \<and> e=e' \<and> S \<sqsubseteq> S')" |
-"le_acom (c1;c2) (c1';c2') = (le_acom c1 c1' \<and> le_acom c2 c2')" |
-"le_acom (IF b THEN c1 ELSE c2 {S}) (IF b' THEN c1' ELSE c2' {S'}) =
-  (b=b' \<and> le_acom c1 c1' \<and> le_acom c2 c2' \<and> S \<sqsubseteq> S')" |
-"le_acom ({Inv} WHILE b DO c {P}) ({Inv'} WHILE b' DO c' {P'}) =
-  (b=b' \<and> le_acom c c' \<and> Inv \<sqsubseteq> Inv' \<and> P \<sqsubseteq> P')" |
-"le_acom _ _ = False"
-
-lemma [simp]: "SKIP {S} \<sqsubseteq> c \<longleftrightarrow> (\<exists>S'. c = SKIP {S'} \<and> S \<sqsubseteq> S')"
-by (cases c) auto
-
-lemma [simp]: "x ::= e {S} \<sqsubseteq> c \<longleftrightarrow> (\<exists>S'. c = x ::= e {S'} \<and> S \<sqsubseteq> S')"
-by (cases c) auto
-
-lemma [simp]: "c1;c2 \<sqsubseteq> c \<longleftrightarrow> (\<exists>c1' c2'. c = c1';c2' \<and> c1 \<sqsubseteq> c1' \<and> c2 \<sqsubseteq> c2')"
-by (cases c) auto
-
-lemma [simp]: "IF b THEN c1 ELSE c2 {S} \<sqsubseteq> c \<longleftrightarrow>
-  (\<exists>c1' c2' S'. c = IF b THEN c1' ELSE c2' {S'} \<and> c1 \<sqsubseteq> c1' \<and> c2 \<sqsubseteq> c2' \<and> S \<sqsubseteq> S')"
-by (cases c) auto
-
-lemma [simp]: "{Inv} WHILE b DO c {P} \<sqsubseteq> w \<longleftrightarrow>
-  (\<exists>Inv' c' P'. w = {Inv'} WHILE b DO c' {P'} \<and> c \<sqsubseteq> c' \<and> Inv \<sqsubseteq> Inv' \<and> P \<sqsubseteq> P')"
-by (cases w) auto
-
-instance
-proof
-  case goal1 thus ?case by (induct x) auto
-next
-  case goal2 thus ?case
-  apply(induct x y arbitrary: z rule: le_acom.induct)
-  apply (auto intro: le_trans)
-  done
-qed
-
-end
-
-
-subsubsection "Lifting"
-
-instantiation option :: (preord)preord
-begin
-
-fun le_option where
-"Some x \<sqsubseteq> Some y = (x \<sqsubseteq> y)" |
-"None \<sqsubseteq> y = True" |
-"Some _ \<sqsubseteq> None = False"
-
-lemma [simp]: "(x \<sqsubseteq> None) = (x = None)"
-by (cases x) simp_all
-
-lemma [simp]: "(Some x \<sqsubseteq> u) = (\<exists>y. u = Some y & x \<sqsubseteq> y)"
-by (cases u) auto
-
-instance proof
-  case goal1 show ?case by(cases x, simp_all)
-next
-  case goal2 thus ?case
-    by(cases z, simp, cases y, simp, cases x, auto intro: le_trans)
-qed
-
-end
-
-instantiation option :: (SL_top)SL_top
-begin
-
-fun join_option where
-"Some x \<squnion> Some y = Some(x \<squnion> y)" |
-"None \<squnion> y = y" |
-"x \<squnion> None = x"
-
-lemma join_None2[simp]: "x \<squnion> None = x"
-by (cases x) simp_all
-
-definition "\<top> = Some \<top>"
-
-instance proof
-  case goal1 thus ?case by(cases x, simp, cases y, simp_all)
-next
-  case goal2 thus ?case by(cases y, simp, cases x, simp_all)
-next
-  case goal3 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all)
-next
-  case goal4 thus ?case by(cases x, simp_all add: Top_option_def)
-qed
-
-end
-
-definition bot_acom :: "com \<Rightarrow> ('a::SL_top)option acom" ("\<bottom>\<^sub>c") where
-"\<bottom>\<^sub>c = anno None"
-
-lemma strip_bot_acom[simp]: "strip(\<bottom>\<^sub>c c) = c"
-by(simp add: bot_acom_def)
-
-lemma bot_acom[rule_format]: "strip c' = c \<longrightarrow> \<bottom>\<^sub>c c \<sqsubseteq> c'"
-apply(induct c arbitrary: c')
-apply (simp_all add: bot_acom_def)
- apply(induct_tac c')
-  apply simp_all
- apply(induct_tac c')
-  apply simp_all
- apply(induct_tac c')
-  apply simp_all
- apply(induct_tac c')
-  apply simp_all
- apply(induct_tac c')
-  apply simp_all
-done
-
-
-subsubsection "Post-fixed point iteration"
-
-definition
-  pfp :: "(('a::preord) \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a option" where
-"pfp f = while_option (\<lambda>x. \<not> f x \<sqsubseteq> x) f"
-
-lemma pfp_pfp: assumes "pfp f x0 = Some x" shows "f x \<sqsubseteq> x"
-using while_option_stop[OF assms[simplified pfp_def]] by simp
-
-lemma pfp_least:
-assumes mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
-and "f p \<sqsubseteq> p" and "x0 \<sqsubseteq> p" and "pfp f x0 = Some x" shows "x \<sqsubseteq> p"
-proof-
-  { fix x assume "x \<sqsubseteq> p"
-    hence  "f x \<sqsubseteq> f p" by(rule mono)
-    from this `f p \<sqsubseteq> p` have "f x \<sqsubseteq> p" by(rule le_trans)
-  }
-  thus "x \<sqsubseteq> p" using assms(2-) while_option_rule[where P = "%x. x \<sqsubseteq> p"]
-    unfolding pfp_def by blast
-qed
-
-definition
- lpfp\<^isub>c :: "(('a::SL_top)option acom \<Rightarrow> 'a option acom) \<Rightarrow> com \<Rightarrow> 'a option acom option" where
-"lpfp\<^isub>c f c = pfp f (\<bottom>\<^sub>c c)"
-
-lemma lpfpc_pfp: "lpfp\<^isub>c f c0 = Some c \<Longrightarrow> f c \<sqsubseteq> c"
-by(simp add: pfp_pfp lpfp\<^isub>c_def)
-
-lemma strip_pfp:
-assumes "\<And>x. g(f x) = g x" and "pfp f x0 = Some x" shows "g x = g x0"
-using assms while_option_rule[where P = "%x. g x = g x0" and c = f]
-unfolding pfp_def by metis
-
-lemma strip_lpfpc: assumes "\<And>c. strip(f c) = strip c" and "lpfp\<^isub>c f c = Some c'"
-shows "strip c' = c"
-using assms(1) strip_pfp[OF _ assms(2)[simplified lpfp\<^isub>c_def]]
-by(metis strip_bot_acom)
-
-lemma lpfpc_least:
-assumes mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
-and "strip p = c0" and "f p \<sqsubseteq> p" and lp: "lpfp\<^isub>c f c0 = Some c" shows "c \<sqsubseteq> p"
-using pfp_least[OF _ _ bot_acom[OF `strip p = c0`] lp[simplified lpfp\<^isub>c_def]]
-  mono `f p \<sqsubseteq> p`
-by blast
-
-
-subsection "Abstract Interpretation"
-
-definition \<gamma>_fun :: "('a \<Rightarrow> 'b set) \<Rightarrow> ('c \<Rightarrow> 'a) \<Rightarrow> ('c \<Rightarrow> 'b)set" where
-"\<gamma>_fun \<gamma> F = {f. \<forall>x. f x \<in> \<gamma>(F x)}"
-
-fun \<gamma>_option :: "('a \<Rightarrow> 'b set) \<Rightarrow> 'a option \<Rightarrow> 'b set" where
-"\<gamma>_option \<gamma> None = {}" |
-"\<gamma>_option \<gamma> (Some a) = \<gamma> a"
-
-text{* The interface for abstract values: *}
-
-locale Val_abs =
-fixes \<gamma> :: "'av::SL_top \<Rightarrow> val set"
-  assumes mono_gamma: "a \<sqsubseteq> b \<Longrightarrow> \<gamma> a \<subseteq> \<gamma> b"
-  and gamma_Top[simp]: "\<gamma> \<top> = UNIV"
-fixes num' :: "val \<Rightarrow> 'av"
-and plus' :: "'av \<Rightarrow> 'av \<Rightarrow> 'av"
-  assumes gamma_num': "n : \<gamma>(num' n)"
-  and gamma_plus':
- "n1 : \<gamma> a1 \<Longrightarrow> n2 : \<gamma> a2 \<Longrightarrow> n1+n2 : \<gamma>(plus' a1 a2)"
-
-type_synonym 'av st = "(vname \<Rightarrow> 'av)"
-
-locale Abs_Int_Fun = Val_abs \<gamma> for \<gamma> :: "'av::SL_top \<Rightarrow> val set"
-begin
-
-fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where
-"aval' (N n) S = num' n" |
-"aval' (V x) S = S x" |
-"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
-
-fun step' :: "'av st option \<Rightarrow> 'av st option acom \<Rightarrow> 'av st option acom"
- where
-"step' S (SKIP {P}) = (SKIP {S})" |
-"step' S (x ::= e {P}) =
-  x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(S(x := aval' e S))}" |
-"step' S (c1; c2) = step' S c1; step' (post c1) c2" |
-"step' S (IF b THEN c1 ELSE c2 {P}) =
-   IF b THEN step' S c1 ELSE step' S c2 {post c1 \<squnion> post c2}" |
-"step' S ({Inv} WHILE b DO c {P}) =
-  {S \<squnion> post c} WHILE b DO (step' Inv c) {Inv}"
-
-definition AI :: "com \<Rightarrow> 'av st option acom option" where
-"AI = lpfp\<^isub>c (step' \<top>)"
-
-
-lemma strip_step'[simp]: "strip(step' S c) = strip c"
-by(induct c arbitrary: S) (simp_all add: Let_def)
-
-
-abbreviation \<gamma>\<^isub>f :: "'av st \<Rightarrow> state set"
-where "\<gamma>\<^isub>f == \<gamma>_fun \<gamma>"
-
-abbreviation \<gamma>\<^isub>o :: "'av st option \<Rightarrow> state set"
-where "\<gamma>\<^isub>o == \<gamma>_option \<gamma>\<^isub>f"
-
-abbreviation \<gamma>\<^isub>c :: "'av st option acom \<Rightarrow> state set acom"
-where "\<gamma>\<^isub>c == map_acom \<gamma>\<^isub>o"
-
-lemma gamma_f_Top[simp]: "\<gamma>\<^isub>f Top = UNIV"
-by(simp add: Top_fun_def \<gamma>_fun_def)
-
-lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o Top = UNIV"
-by (simp add: Top_option_def)
-
-(* FIXME (maybe also le \<rightarrow> sqle?) *)
-
-lemma mono_gamma_f: "f \<sqsubseteq> g \<Longrightarrow> \<gamma>\<^isub>f f \<subseteq> \<gamma>\<^isub>f g"
-by(auto simp: le_fun_def \<gamma>_fun_def dest: mono_gamma)
-
-lemma mono_gamma_o:
-  "sa \<sqsubseteq> sa' \<Longrightarrow> \<gamma>\<^isub>o sa \<subseteq> \<gamma>\<^isub>o sa'"
-by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f)
-
-lemma mono_gamma_c: "ca \<sqsubseteq> ca' \<Longrightarrow> \<gamma>\<^isub>c ca \<le> \<gamma>\<^isub>c ca'"
-by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o)
-
-text{* Soundness: *}
-
-lemma aval'_sound: "s : \<gamma>\<^isub>f S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
-by (induct a) (auto simp: gamma_num' gamma_plus' \<gamma>_fun_def)
-
-lemma in_gamma_update:
-  "\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(S(x := a))"
-by(simp add: \<gamma>_fun_def)
-
-lemma step_preserves_le:
-  "\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; c \<le> \<gamma>\<^isub>c c' \<rbrakk> \<Longrightarrow> step S c \<le> \<gamma>\<^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 intro: aval'_sound in_gamma_update
-      split: option.splits del:subsetD)
-next
-  case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi)
-    by (metis le_post post_map_acom)
-next
-  case (If b c1 c2 P)
-  then obtain c1' c2' P' where
-      "c' = IF b THEN c1' ELSE c2' {P'}"
-      "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'" "c2 \<le> \<gamma>\<^isub>c c2'"
-    by (fastforce simp: If_le map_acom_If)
-  moreover have "post c1 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
-    by (metis (no_types) `c1 \<le> \<gamma>\<^isub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
-  moreover have "post c2 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
-    by (metis (no_types) `c2 \<le> \<gamma>\<^isub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
-  ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'` by (simp add: If.IH subset_iff)
-next
-  case (While I b c1 P)
-  then obtain c1' I' P' where
-    "c' = {I'} WHILE b DO c1' {P'}"
-    "I \<subseteq> \<gamma>\<^isub>o I'" "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'"
-    by (fastforce simp: map_acom_While While_le)
-  moreover have "S \<union> post c1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post c1')"
-    using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `c1 \<le> \<gamma>\<^isub>c c1'`, simplified]
-    by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans)
-  ultimately show ?case by (simp add: While.IH subset_iff)
-qed
-
-lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
-proof(simp add: CS_def AI_def)
-  assume 1: "lpfp\<^isub>c (step' \<top>) c = Some c'"
-  have 2: "step' \<top> c' \<sqsubseteq> c'" by(rule lpfpc_pfp[OF 1])
-  have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c"
-    by(simp add: strip_lpfpc[OF _ 1])
-  have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
-  proof(rule lfp_lowerbound[simplified,OF 3])
-    show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
-    proof(rule step_preserves_le[OF _ _])
-      show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
-      show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
-    qed
-  qed
-  with 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
-    by (blast intro: mono_gamma_c order_trans)
-qed
-
-end
-
-
-subsubsection "Monotonicity"
-
-lemma mono_post: "c \<sqsubseteq> c' \<Longrightarrow> post c \<sqsubseteq> post c'"
-by(induction c c' rule: le_acom.induct) (auto)
-
-locale Abs_Int_Fun_mono = Abs_Int_Fun +
-assumes mono_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> plus' a1 a2 \<sqsubseteq> plus' b1 b2"
-begin
-
-lemma mono_aval': "S \<sqsubseteq> S' \<Longrightarrow> aval' e S \<sqsubseteq> aval' e S'"
-by(induction e)(auto simp: le_fun_def mono_plus')
-
-lemma mono_update: "a \<sqsubseteq> a' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> S(x := a) \<sqsubseteq> S'(x := a')"
-by(simp add: le_fun_def)
-
-lemma mono_step': "S \<sqsubseteq> S' \<Longrightarrow> c \<sqsubseteq> c' \<Longrightarrow> step' S c \<sqsubseteq> step' S' c'"
-apply(induction c c' arbitrary: S S' rule: le_acom.induct)
-apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj
-            split: option.split)
-done
-
-end
-
-text{* Problem: not executable because of the comparison of abstract states,
-i.e. functions, in the post-fixedpoint computation. *}
-
-end
--- a/src/HOL/IMP/Abs_Int0_parity.thy	Thu Apr 19 12:28:10 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,166 +0,0 @@
-theory Abs_Int0_parity
-imports Abs_Int0
-begin
-
-subsection "Parity Analysis"
-
-datatype parity = Even | Odd | Either
-
-text{* Instantiation of class @{class preord} with type @{typ parity}: *}
-
-instantiation parity :: preord
-begin
-
-text{* First the definition of the interface function @{text"\<sqsubseteq>"}. Note that
-the header of the definition must refer to the ascii name @{const le} of the
-constants as @{text le_parity} and the definition is named @{text
-le_parity_def}.  Inside the definition the symbolic names can be used. *}
-
-definition le_parity where
-"x \<sqsubseteq> y = (y = Either \<or> x=y)"
-
-text{* Now the instance proof, i.e.\ the proof that the definition fulfills
-the axioms (assumptions) of the class. The initial proof-step generates the
-necessary proof obligations. *}
-
-instance
-proof
-  fix x::parity show "x \<sqsubseteq> x" by(auto simp: le_parity_def)
-next
-  fix x y z :: parity assume "x \<sqsubseteq> y" "y \<sqsubseteq> z" thus "x \<sqsubseteq> z"
-    by(auto simp: le_parity_def)
-qed
-
-end
-
-text{* Instantiation of class @{class SL_top} with type @{typ parity}: *}
-
-instantiation parity :: SL_top
-begin
-
-
-definition join_parity where
-"x \<squnion> y = (if x \<sqsubseteq> y then y else if y \<sqsubseteq> x then x else Either)"
-
-definition Top_parity where
-"\<top> = Either"
-
-text{* Now the instance proof. This time we take a lazy shortcut: we do not
-write out the proof obligations but use the @{text goali} primitive to refer
-to the assumptions of subgoal i and @{text "case?"} to refer to the
-conclusion of subgoal i. The class axioms are presented in the same order as
-in the class definition. *}
-
-instance
-proof
-  case goal1 (*join1*) show ?case by(auto simp: le_parity_def join_parity_def)
-next
-  case goal2 (*join2*) show ?case by(auto simp: le_parity_def join_parity_def)
-next
-  case goal3 (*join least*) thus ?case by(auto simp: le_parity_def join_parity_def)
-next
-  case goal4 (*Top*) show ?case by(auto simp: le_parity_def Top_parity_def)
-qed
-
-end
-
-
-text{* Now we define the functions used for instantiating the abstract
-interpretation locales. Note that the Isabelle terminology is
-\emph{interpretation}, not \emph{instantiation} of locales, but we use
-instantiation to avoid confusion with abstract interpretation.  *}
-
-fun \<gamma>_parity :: "parity \<Rightarrow> val set" where
-"\<gamma>_parity Even = {i. i mod 2 = 0}" |
-"\<gamma>_parity Odd  = {i. i mod 2 = 1}" |
-"\<gamma>_parity Either = UNIV"
-
-fun num_parity :: "val \<Rightarrow> parity" where
-"num_parity i = (if i mod 2 = 0 then Even else Odd)"
-
-fun plus_parity :: "parity \<Rightarrow> parity \<Rightarrow> parity" where
-"plus_parity Even Even = Even" |
-"plus_parity Odd  Odd  = Even" |
-"plus_parity Even Odd  = Odd" |
-"plus_parity Odd  Even = Odd" |
-"plus_parity Either y  = Either" |
-"plus_parity x Either  = Either"
-
-text{* First we instantiate the abstract value interface and prove that the
-functions on type @{typ parity} have all the necessary properties: *}
-
-interpretation Val_abs
-where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity
-proof txt{* of the locale axioms *}
-  fix a b :: parity
-  assume "a \<sqsubseteq> b" thus "\<gamma>_parity a \<subseteq> \<gamma>_parity b"
-    by(auto simp: le_parity_def)
-next txt{* The rest in the lazy, implicit way *}
-  case goal2 show ?case by(auto simp: Top_parity_def)
-next
-  case goal3 show ?case by auto
-next
-  txt{* Warning: this subproof refers to the names @{text a1} and @{text a2}
-  from the statement of the axiom. *}
-  case goal4 thus ?case
-  proof(cases a1 a2 rule: parity.exhaust[case_product parity.exhaust])
-  qed (auto simp add:mod_add_eq)
-qed
-
-text{* Instantiating the abstract interpretation locale requires no more
-proofs (they happened in the instatiation above) but delivers the
-instantiated abstract interpreter which we call AI: *}
-
-interpretation Abs_Int
-where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity
-defines aval_parity is aval' and step_parity is step' and AI_parity is AI
-..
-
-
-subsubsection "Tests"
-
-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)"
-
-definition "test2_parity =
-  ''x'' ::= N 1;
-  WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 3)"
-
-value "show_acom ((step_parity \<top> ^^1) (anno None test2_parity))"
-value "show_acom ((step_parity \<top> ^^2) (anno None test2_parity))"
-value "show_acom ((step_parity \<top> ^^3) (anno None test2_parity))"
-value "show_acom ((step_parity \<top> ^^4) (anno None test2_parity))"
-value "show_acom ((step_parity \<top> ^^5) (anno None test2_parity))"
-value "show_acom_opt (AI_parity test2_parity)"
-
-
-subsubsection "Termination"
-
-interpretation Abs_Int_mono
-where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity
-proof
-  case goal1 thus ?case
-  proof(cases a1 a2 b1 b2
-   rule: parity.exhaust[case_product parity.exhaust[case_product parity.exhaust[case_product parity.exhaust]]]) (* FIXME - UGLY! *)
-  qed (auto simp add:le_parity_def)
-qed
-
-
-definition m_parity :: "parity \<Rightarrow> nat" where
-"m_parity x = (if x=Either then 0 else 1)"
-
-lemma measure_parity:
-  "(strict{(x::parity,y). x \<sqsubseteq> y})^-1 \<subseteq> measure m_parity"
-by(auto simp add: m_parity_def le_parity_def)
-
-lemma measure_parity_eq:
-  "\<forall>x y::parity. x \<sqsubseteq> y \<and> y \<sqsubseteq> x \<longrightarrow> m_parity x = m_parity y"
-by(auto simp add: m_parity_def le_parity_def)
-
-lemma AI_parity_Some: "\<exists>c'. AI_parity c = Some c'"
-by(rule AI_Some_measure[OF measure_parity measure_parity_eq])
-
-end
--- a/src/HOL/IMP/Abs_Int1.thy	Thu Apr 19 12:28:10 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,358 +0,0 @@
-(* Author: Tobias Nipkow *)
-
-theory Abs_Int1
-imports Abs_Int0 Vars
-begin
-
-instantiation prod :: (preord,preord) preord
-begin
-
-definition "le_prod p1 p2 = (fst p1 \<sqsubseteq> fst p2 \<and> snd p1 \<sqsubseteq> snd p2)"
-
-instance
-proof
-  case goal1 show ?case by(simp add: le_prod_def)
-next
-  case goal2 thus ?case unfolding le_prod_def by(metis le_trans)
-qed
-
-end
-
-instantiation com :: vars
-begin
-
-fun vars_com :: "com \<Rightarrow> vname set" where
-"vars com.SKIP = {}" |
-"vars (x::=e) = {x} \<union> vars e" |
-"vars (c1;c2) = vars c1 \<union> vars c2" |
-"vars (IF b THEN c1 ELSE c2) = vars b \<union> vars c1 \<union> vars c2" |
-"vars (WHILE b DO c) = vars b \<union> vars c"
-
-instance ..
-
-end
-
-lemma finite_avars: "finite(vars(a::aexp))"
-by(induction a) simp_all
-
-lemma finite_bvars: "finite(vars(b::bexp))"
-by(induction b) (simp_all add: finite_avars)
-
-lemma finite_cvars: "finite(vars(c::com))"
-by(induction c) (simp_all add: finite_avars finite_bvars)
-
-
-subsection "Backward Analysis of Expressions"
-
-hide_const bot
-
-class L_top_bot = SL_top +
-fixes meet :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 65)
-and bot :: "'a" ("\<bottom>")
-assumes meet_le1 [simp]: "x \<sqinter> y \<sqsubseteq> x"
-and meet_le2 [simp]: "x \<sqinter> y \<sqsubseteq> y"
-and meet_greatest: "x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<sqinter> z"
-assumes bot[simp]: "\<bottom> \<sqsubseteq> x"
-begin
-
-lemma mono_meet: "x \<sqsubseteq> x' \<Longrightarrow> y \<sqsubseteq> y' \<Longrightarrow> x \<sqinter> y \<sqsubseteq> x' \<sqinter> y'"
-by (metis meet_greatest meet_le1 meet_le2 le_trans)
-
-end
-
-locale Val_abs1_gamma =
-  Gamma where \<gamma> = \<gamma> for \<gamma> :: "'av::L_top_bot \<Rightarrow> val set" +
-assumes inter_gamma_subset_gamma_meet:
-  "\<gamma> a1 \<inter> \<gamma> a2 \<subseteq> \<gamma>(a1 \<sqinter> a2)"
-and gamma_Bot[simp]: "\<gamma> \<bottom> = {}"
-begin
-
-lemma in_gamma_meet: "x : \<gamma> a1 \<Longrightarrow> x : \<gamma> a2 \<Longrightarrow> x : \<gamma>(a1 \<sqinter> a2)"
-by (metis IntI inter_gamma_subset_gamma_meet set_mp)
-
-lemma gamma_meet[simp]: "\<gamma>(a1 \<sqinter> a2) = \<gamma> a1 \<inter> \<gamma> a2"
-by (metis equalityI inter_gamma_subset_gamma_meet le_inf_iff mono_gamma meet_le1 meet_le2)
-
-end
-
-
-locale Val_abs1 =
-  Val_abs1_gamma where \<gamma> = \<gamma>
-   for \<gamma> :: "'av::L_top_bot \<Rightarrow> val set" +
-fixes test_num' :: "val \<Rightarrow> 'av \<Rightarrow> bool"
-and filter_plus' :: "'av \<Rightarrow> 'av \<Rightarrow> 'av \<Rightarrow> 'av * 'av"
-and filter_less' :: "bool \<Rightarrow> 'av \<Rightarrow> 'av \<Rightarrow> 'av * 'av"
-assumes test_num': "test_num' n a = (n : \<gamma> a)"
-and filter_plus': "filter_plus' a a1 a2 = (b1,b2) \<Longrightarrow>
-  n1 : \<gamma> a1 \<Longrightarrow> n2 : \<gamma> a2 \<Longrightarrow> n1+n2 : \<gamma> a \<Longrightarrow> n1 : \<gamma> b1 \<and> n2 : \<gamma> b2"
-and filter_less': "filter_less' (n1<n2) a1 a2 = (b1,b2) \<Longrightarrow>
-  n1 : \<gamma> a1 \<Longrightarrow> n2 : \<gamma> a2 \<Longrightarrow> n1 : \<gamma> b1 \<and> n2 : \<gamma> b2"
-
-
-locale Abs_Int1 =
-  Val_abs1 where \<gamma> = \<gamma> for \<gamma> :: "'av::L_top_bot \<Rightarrow> val set"
-begin
-
-lemma in_gamma_join_UpI: "s : \<gamma>\<^isub>o S1 \<or> s : \<gamma>\<^isub>o S2 \<Longrightarrow> s : \<gamma>\<^isub>o(S1 \<squnion> S2)"
-by (metis (no_types) join_ge1 join_ge2 mono_gamma_o set_rev_mp)
-
-fun aval'' :: "aexp \<Rightarrow> 'av st option \<Rightarrow> 'av" where
-"aval'' e None = \<bottom>" |
-"aval'' e (Some sa) = aval' e sa"
-
-lemma aval''_sound: "s : \<gamma>\<^isub>o S \<Longrightarrow> aval a s : \<gamma>(aval'' a S)"
-by(cases S)(simp add: aval'_sound)+
-
-subsubsection "Backward analysis"
-
-fun afilter :: "aexp \<Rightarrow> 'av \<Rightarrow> 'av st option \<Rightarrow> 'av st option" where
-"afilter (N n) a S = (if test_num' n a then S else None)" |
-"afilter (V x) a S = (case S of None \<Rightarrow> None | Some S \<Rightarrow>
-  let a' = lookup S x \<sqinter> a in
-  if a' \<sqsubseteq> \<bottom> then None else Some(update S x a'))" |
-"afilter (Plus e1 e2) a S =
- (let (a1,a2) = filter_plus' a (aval'' e1 S) (aval'' e2 S)
-  in afilter e1 a1 (afilter e2 a2 S))"
-
-text{* The test for @{const bot} in the @{const V}-case is important: @{const
-bot} indicates that a variable has no possible values, i.e.\ that the current
-program point is unreachable. But then the abstract state should collapse to
-@{const None}. Put differently, we maintain the invariant that in an abstract
-state of the form @{term"Some s"}, all variables are mapped to non-@{const
-bot} values. Otherwise the (pointwise) join of two abstract states, one of
-which contains @{const bot} values, may produce too large a result, thus
-making the analysis less precise. *}
-
-
-fun bfilter :: "bexp \<Rightarrow> bool \<Rightarrow> 'av st option \<Rightarrow> 'av st option" where
-"bfilter (Bc v) res S = (if v=res then S else None)" |
-"bfilter (Not b) res S = bfilter b (\<not> res) S" |
-"bfilter (And b1 b2) res S =
-  (if res then bfilter b1 True (bfilter b2 True S)
-   else bfilter b1 False S \<squnion> bfilter b2 False S)" |
-"bfilter (Less e1 e2) res S =
-  (let (res1,res2) = filter_less' res (aval'' e1 S) (aval'' e2 S)
-   in afilter e1 res1 (afilter e2 res2 S))"
-
-lemma afilter_sound: "s : \<gamma>\<^isub>o S \<Longrightarrow> aval e s : \<gamma> a \<Longrightarrow> s : \<gamma>\<^isub>o (afilter e a S)"
-proof(induction e arbitrary: a S)
-  case N thus ?case by simp (metis test_num')
-next
-  case (V x)
-  obtain S' where "S = Some S'" and "s : \<gamma>\<^isub>f S'" using `s : \<gamma>\<^isub>o S`
-    by(auto simp: in_gamma_option_iff)
-  moreover hence "s x : \<gamma> (lookup S' x)" by(simp add: \<gamma>_st_def)
-  moreover have "s x : \<gamma> a" using V by simp
-  ultimately show ?case using V(1)
-    by(simp add: lookup_update Let_def \<gamma>_st_def)
-      (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(3)] aval''_sound[OF Plus(3)]]
-    by (auto split: prod.split)
-qed
-
-lemma bfilter_sound: "s : \<gamma>\<^isub>o S \<Longrightarrow> bv = bval b s \<Longrightarrow> s : \<gamma>\<^isub>o(bfilter b bv S)"
-proof(induction b arbitrary: S bv)
-  case Bc thus ?case by simp
-next
-  case (Not b) thus ?case by simp
-next
-  case (And b1 b2) thus ?case by(fastforce simp: in_gamma_join_UpI)
-next
-  case (Less e1 e2) thus ?case
-    by (auto split: prod.split)
-       (metis afilter_sound filter_less' aval''_sound Less)
-qed
-
-
-fun step' :: "'av st option \<Rightarrow> 'av st option acom \<Rightarrow> 'av st option acom"
- where
-"step' S (SKIP {P}) = (SKIP {S})" |
-"step' S (x ::= e {P}) =
-  x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(update S x (aval' e S))}" |
-"step' S (c1; c2) = step' S c1; step' (post c1) c2" |
-"step' S (IF b THEN c1 ELSE c2 {P}) =
-  (let c1' = step' (bfilter b True S) c1; c2' = step' (bfilter b False S) c2
-   in IF b THEN c1' ELSE c2' {post c1 \<squnion> post c2})" |
-"step' S ({Inv} WHILE b DO c {P}) =
-   {S \<squnion> post c}
-   WHILE b DO step' (bfilter b True Inv) c
-   {bfilter b False Inv}"
-
-definition AI :: "com \<Rightarrow> 'av st option acom option" where
-"AI = lpfp\<^isub>c (step' \<top>)"
-
-lemma strip_step'[simp]: "strip(step' S c) = strip c"
-by(induct c arbitrary: S) (simp_all add: Let_def)
-
-
-subsubsection "Soundness"
-
-lemma in_gamma_update:
-  "\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(update S x a)"
-by(simp add: \<gamma>_st_def lookup_update)
-
-lemma step_preserves_le:
-  "\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; cs \<le> \<gamma>\<^isub>c ca \<rbrakk> \<Longrightarrow> step S cs \<le> \<gamma>\<^isub>c (step' S' ca)"
-proof(induction cs arbitrary: ca 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 intro: aval'_sound in_gamma_update
-      split: option.splits del:subsetD)
-next
-  case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi)
-    by (metis le_post post_map_acom)
-next
-  case (If b cs1 cs2 P)
-  then obtain ca1 ca2 Pa where
-      "ca= IF b THEN ca1 ELSE ca2 {Pa}"
-      "P \<subseteq> \<gamma>\<^isub>o Pa" "cs1 \<le> \<gamma>\<^isub>c ca1" "cs2 \<le> \<gamma>\<^isub>c ca2"
-    by (fastforce simp: If_le map_acom_If)
-  moreover have "post cs1 \<subseteq> \<gamma>\<^isub>o(post ca1 \<squnion> post ca2)"
-    by (metis (no_types) `cs1 \<le> \<gamma>\<^isub>c ca1` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
-  moreover have "post cs2 \<subseteq> \<gamma>\<^isub>o(post ca1 \<squnion> post ca2)"
-    by (metis (no_types) `cs2 \<le> \<gamma>\<^isub>c ca2` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
-  ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'`
-    by (simp add: If.IH subset_iff bfilter_sound)
-next
-  case (While I b cs1 P)
-  then obtain ca1 Ia Pa where
-    "ca = {Ia} WHILE b DO ca1 {Pa}"
-    "I \<subseteq> \<gamma>\<^isub>o Ia" "P \<subseteq> \<gamma>\<^isub>o Pa" "cs1 \<le> \<gamma>\<^isub>c ca1"
-    by (fastforce simp: map_acom_While While_le)
-  moreover have "S \<union> post cs1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post ca1)"
-    using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `cs1 \<le> \<gamma>\<^isub>c ca1`, simplified]
-    by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans)
-  ultimately show ?case by (simp add: While.IH subset_iff bfilter_sound)
-qed
-
-lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
-proof(simp add: CS_def AI_def)
-  assume 1: "lpfp\<^isub>c (step' \<top>) c = Some c'"
-  have 2: "step' \<top> c' \<sqsubseteq> c'" by(rule lpfpc_pfp[OF 1])
-  have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c"
-    by(simp add: strip_lpfpc[OF _ 1])
-  have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
-  proof(rule lfp_lowerbound[simplified,OF 3])
-    show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
-    proof(rule step_preserves_le[OF _ _])
-      show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
-      show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
-    qed
-  qed
-  from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
-    by (blast intro: mono_gamma_c order_trans)
-qed
-
-
-subsubsection "Commands over a set of variables"
-
-text{* Key invariant: the domains of all abstract states are subsets of the
-set of variables of the program. *}
-
-definition "domo S = (case S of None \<Rightarrow> {} | Some S' \<Rightarrow> set(dom S'))"
-
-definition Com :: "vname set \<Rightarrow> 'a st option acom set" where
-"Com X = {c. \<forall>S \<in> set(annos c). domo S \<subseteq> X}"
-
-lemma domo_Top[simp]: "domo \<top> = {}"
-by(simp add: domo_def Top_st_def Top_option_def)
-
-lemma bot_acom_Com[simp]: "\<bottom>\<^sub>c c \<in> Com X"
-by(simp add: bot_acom_def Com_def domo_def set_annos_anno)
-
-lemma post_in_annos: "post c : set(annos c)"
-by(induction c) simp_all
-
-lemma domo_join: "domo (S \<squnion> T) \<subseteq> domo S \<union> domo T"
-by(auto simp: domo_def join_st_def split: option.split)
-
-lemma domo_afilter: "vars a \<subseteq> X \<Longrightarrow> domo S \<subseteq> X \<Longrightarrow> domo(afilter a i S) \<subseteq> X"
-apply(induction a arbitrary: i S)
-apply(simp add: domo_def)
-apply(simp add: domo_def Let_def update_def lookup_def split: option.splits)
-apply blast
-apply(simp split: prod.split)
-done
-
-lemma domo_bfilter: "vars b \<subseteq> X \<Longrightarrow> domo S \<subseteq> X \<Longrightarrow> domo(bfilter b bv S) \<subseteq> X"
-apply(induction b arbitrary: bv S)
-apply(simp add: domo_def)
-apply(simp)
-apply(simp)
-apply rule
-apply (metis le_sup_iff subset_trans[OF domo_join])
-apply(simp split: prod.split)
-by (metis domo_afilter)
-
-lemma step'_Com:
-  "domo S \<subseteq> X \<Longrightarrow> vars(strip c) \<subseteq> X \<Longrightarrow> c : Com X \<Longrightarrow> step' S c : Com X"
-apply(induction c arbitrary: S)
-apply(simp add: Com_def)
-apply(simp add: Com_def domo_def update_def split: option.splits)
-apply(simp (no_asm_use) add: Com_def ball_Un)
-apply (metis post_in_annos)
-apply(simp (no_asm_use) add: Com_def ball_Un)
-apply rule
-apply (metis Un_assoc domo_join order_trans post_in_annos subset_Un_eq)
-apply (metis domo_bfilter)
-apply(simp (no_asm_use) add: Com_def)
-apply rule
-apply (metis domo_join le_sup_iff post_in_annos subset_trans)
-apply rule
-apply (metis domo_bfilter)
-by (metis domo_bfilter)
-
-end
-
-
-subsubsection "Monotonicity"
-
-locale Abs_Int1_mono = Abs_Int1 +
-assumes mono_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> plus' a1 a2 \<sqsubseteq> plus' b1 b2"
-and mono_filter_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> r \<sqsubseteq> r' \<Longrightarrow>
-  filter_plus' r a1 a2 \<sqsubseteq> filter_plus' r' b1 b2"
-and mono_filter_less': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow>
-  filter_less' bv a1 a2 \<sqsubseteq> filter_less' bv b1 b2"
-begin
-
-lemma mono_aval': "S \<sqsubseteq> S' \<Longrightarrow> aval' e S \<sqsubseteq> aval' e S'"
-by(induction e) (auto simp: le_st_def lookup_def mono_plus')
-
-lemma mono_aval'': "S \<sqsubseteq> S' \<Longrightarrow> aval'' e S \<sqsubseteq> aval'' e S'"
-apply(cases S)
- apply simp
-apply(cases S')
- apply simp
-by (simp add: mono_aval')
-
-lemma mono_afilter: "r \<sqsubseteq> r' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> afilter e r S \<sqsubseteq> afilter e r' S'"
-apply(induction e arbitrary: r r' S S')
-apply(auto simp: test_num' Let_def split: option.splits prod.splits)
-apply (metis mono_gamma subsetD)
-apply(drule_tac x = "list" in mono_lookup)
-apply (metis mono_meet le_trans)
-apply (metis mono_meet mono_lookup mono_update)
-apply(metis mono_aval'' mono_filter_plus'[simplified le_prod_def] fst_conv snd_conv)
-done
-
-lemma mono_bfilter: "S \<sqsubseteq> S' \<Longrightarrow> bfilter b r S \<sqsubseteq> bfilter b r S'"
-apply(induction b arbitrary: r S S')
-apply(auto simp: le_trans[OF _ join_ge1] le_trans[OF _ join_ge2] split: prod.splits)
-apply(metis mono_aval'' mono_afilter mono_filter_less'[simplified le_prod_def] fst_conv snd_conv)
-done
-
-lemma mono_step': "S \<sqsubseteq> S' \<Longrightarrow> c \<sqsubseteq> c' \<Longrightarrow> step' S c \<sqsubseteq> step' S' c'"
-apply(induction c c' arbitrary: S S' rule: le_acom.induct)
-apply (auto simp: mono_post mono_bfilter mono_update mono_aval' Let_def le_join_disj
-  split: option.split)
-done
-
-lemma mono_step'2: "mono (step' S)"
-by(simp add: mono_def mono_step'[OF le_refl])
-
-end
-
-end
--- a/src/HOL/IMP/Abs_Int1_ivl.thy	Thu Apr 19 12:28:10 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,285 +0,0 @@
-(* Author: Tobias Nipkow *)
-
-theory Abs_Int1_ivl
-imports Abs_Int1 Abs_Int_Tests
-begin
-
-subsection "Interval Analysis"
-
-datatype ivl = I "int option" "int option"
-
-definition "\<gamma>_ivl i = (case i of
-  I (Some l) (Some h) \<Rightarrow> {l..h} |
-  I (Some l) None \<Rightarrow> {l..} |
-  I None (Some h) \<Rightarrow> {..h} |
-  I None None \<Rightarrow> UNIV)"
-
-abbreviation I_Some_Some :: "int \<Rightarrow> int \<Rightarrow> ivl"  ("{_\<dots>_}") where
-"{lo\<dots>hi} == I (Some lo) (Some hi)"
-abbreviation I_Some_None :: "int \<Rightarrow> ivl"  ("{_\<dots>}") where
-"{lo\<dots>} == I (Some lo) None"
-abbreviation I_None_Some :: "int \<Rightarrow> ivl"  ("{\<dots>_}") where
-"{\<dots>hi} == I None (Some hi)"
-abbreviation I_None_None :: "ivl"  ("{\<dots>}") where
-"{\<dots>} == I None None"
-
-definition "num_ivl n = {n\<dots>n}"
-
-fun in_ivl :: "int \<Rightarrow> ivl \<Rightarrow> bool" where
-"in_ivl k (I (Some l) (Some h)) \<longleftrightarrow> l \<le> k \<and> k \<le> h" |
-"in_ivl k (I (Some l) None) \<longleftrightarrow> l \<le> k" |
-"in_ivl k (I None (Some h)) \<longleftrightarrow> k \<le> h" |
-"in_ivl k (I None None) \<longleftrightarrow> True"
-
-instantiation option :: (plus)plus
-begin
-
-fun plus_option where
-"Some x + Some y = Some(x+y)" |
-"_ + _ = None"
-
-instance ..
-
-end
-
-definition empty where "empty = {1\<dots>0}"
-
-fun is_empty where
-"is_empty {l\<dots>h} = (h<l)" |
-"is_empty _ = False"
-
-lemma [simp]: "is_empty(I l h) =
-  (case l of Some l \<Rightarrow> (case h of Some h \<Rightarrow> h<l | None \<Rightarrow> False) | None \<Rightarrow> False)"
-by(auto split:option.split)
-
-lemma [simp]: "is_empty i \<Longrightarrow> \<gamma>_ivl i = {}"
-by(auto simp add: \<gamma>_ivl_def split: ivl.split option.split)
-
-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) \<Rightarrow> I (l1+l2) (h1+h2))"
-
-instantiation ivl :: SL_top
-begin
-
-definition le_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> bool" where
-"le_option pos x y =
- (case x of (Some i) \<Rightarrow> (case y of Some j \<Rightarrow> i\<le>j | None \<Rightarrow> pos)
-  | None \<Rightarrow> (case y of Some j \<Rightarrow> \<not>pos | None \<Rightarrow> True))"
-
-fun le_aux where
-"le_aux (I l1 h1) (I l2 h2) = (le_option False l2 l1 & le_option True h1 h2)"
-
-definition le_ivl where
-"i1 \<sqsubseteq> i2 =
- (if is_empty i1 then True else
-  if is_empty i2 then False else le_aux i1 i2)"
-
-definition min_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> int option" where
-"min_option pos o1 o2 = (if le_option pos o1 o2 then o1 else o2)"
-
-definition max_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> int option" where
-"max_option pos o1 o2 = (if le_option pos o1 o2 then o2 else o1)"
-
-definition "i1 \<squnion> i2 =
- (if is_empty i1 then i2 else if is_empty i2 then i1
-  else case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow>
-          I (min_option False l1 l2) (max_option True h1 h2))"
-
-definition "\<top> = {\<dots>}"
-
-instance
-proof
-  case goal1 thus ?case
-    by(cases x, simp add: le_ivl_def le_option_def split: option.split)
-next
-  case goal2 thus ?case
-    by(cases x, cases y, cases z, auto simp: le_ivl_def le_option_def split: option.splits if_splits)
-next
-  case goal3 thus ?case
-    by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits)
-next
-  case goal4 thus ?case
-    by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits)
-next
-  case goal5 thus ?case
-    by(cases x, cases y, cases z, auto simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits if_splits)
-next
-  case goal6 thus ?case
-    by(cases x, simp add: Top_ivl_def le_ivl_def le_option_def split: option.split)
-qed
-
-end
-
-
-instantiation ivl :: L_top_bot
-begin
-
-definition "i1 \<sqinter> i2 = (if is_empty i1 \<or> is_empty i2 then empty else
-  case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow>
-    I (max_option False l1 l2) (min_option True h1 h2))"
-
-definition "\<bottom> = empty"
-
-instance
-proof
-  case goal1 thus ?case
-    by (simp add:meet_ivl_def empty_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits)
-next
-  case goal2 thus ?case
-    by (simp add: empty_def meet_ivl_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits)
-next
-  case goal3 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 goal4 show ?case by(cases x, simp add: bot_ivl_def empty_def le_ivl_def)
-qed
-
-end
-
-instantiation option :: (minus)minus
-begin
-
-fun minus_option where
-"Some x - Some y = Some(x-y)" |
-"_ - _ = None"
-
-instance ..
-
-end
-
-definition "minus_ivl i1 i2 = (if is_empty i1 | is_empty i2 then empty else
-  case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow> I (l1-h2) (h1-l2))"
-
-lemma gamma_minus_ivl:
-  "n1 : \<gamma>_ivl i1 \<Longrightarrow> n2 : \<gamma>_ivl i2 \<Longrightarrow> n1-n2 : \<gamma>_ivl(minus_ivl i1 i2)"
-by(auto simp add: minus_ivl_def \<gamma>_ivl_def split: ivl.splits option.splits)
-
-definition "filter_plus_ivl i i1 i2 = ((*if is_empty i then empty else*)
-  i1 \<sqinter> minus_ivl i i2, i2 \<sqinter> minus_ivl i i1)"
-
-fun filter_less_ivl :: "bool \<Rightarrow> ivl \<Rightarrow> ivl \<Rightarrow> ivl * ivl" where
-"filter_less_ivl res (I l1 h1) (I l2 h2) =
-  (if is_empty(I l1 h1) \<or> is_empty(I l2 h2) then (empty, empty) else
-   if res
-   then (I l1 (min_option True h1 (h2 - Some 1)),
-         I (max_option False (l1 + Some 1) l2) h2)
-   else (I (max_option False l1 l2) h1, I l2 (min_option True h1 h2)))"
-
-interpretation Val_abs
-where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = plus_ivl
-proof
-  case goal1 thus ?case
-    by(auto simp: \<gamma>_ivl_def le_ivl_def le_option_def split: ivl.split option.split if_splits)
-next
-  case goal2 show ?case by(simp add: \<gamma>_ivl_def Top_ivl_def)
-next
-  case goal3 thus ?case by(simp add: \<gamma>_ivl_def num_ivl_def)
-next
-  case goal4 thus ?case
-    by(auto simp add: \<gamma>_ivl_def plus_ivl_def split: ivl.split option.splits)
-qed
-
-interpretation Val_abs1_gamma
-where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = plus_ivl
-defines aval_ivl is aval'
-proof
-  case goal1 thus ?case
-    by(auto simp add: \<gamma>_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 \<gamma>_ivl_def empty_def)
-qed
-
-lemma mono_minus_ivl:
-  "i1 \<sqsubseteq> i1' \<Longrightarrow> i2 \<sqsubseteq> i2' \<Longrightarrow> minus_ivl i1 i2 \<sqsubseteq> minus_ivl i1' i2'"
-apply(auto simp add: minus_ivl_def empty_def le_ivl_def le_option_def split: ivl.splits)
-  apply(simp split: option.splits)
- apply(simp split: option.splits)
-apply(simp split: option.splits)
-done
-
-
-interpretation Val_abs1
-where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = plus_ivl
-and test_num' = in_ivl
-and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl
-proof
-  case goal1 thus ?case
-    by (simp add: \<gamma>_ivl_def split: ivl.split option.split)
-next
-  case goal2 thus ?case
-    by(auto simp add: filter_plus_ivl_def)
-      (metis gamma_minus_ivl add_diff_cancel add_commute)+
-next
-  case goal3 thus ?case
-    by(cases a1, cases a2,
-      auto simp: \<gamma>_ivl_def min_option_def max_option_def le_option_def split: if_splits option.splits)
-qed
-
-interpretation Abs_Int1
-where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = plus_ivl
-and test_num' = in_ivl
-and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl
-defines afilter_ivl is afilter
-and bfilter_ivl is bfilter
-and step_ivl is step'
-and AI_ivl is AI
-and aval_ivl' is aval''
-..
-
-
-text{* Monotonicity: *}
-
-interpretation Abs_Int1_mono
-where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = plus_ivl
-and test_num' = in_ivl
-and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl
-proof
-  case goal1 thus ?case
-    by(auto simp: plus_ivl_def le_ivl_def le_option_def empty_def split: if_splits ivl.splits option.splits)
-next
-  case goal2 thus ?case
-    by(auto simp: filter_plus_ivl_def le_prod_def mono_meet mono_minus_ivl)
-next
-  case goal3 thus ?case
-    apply(cases a1, cases b1, cases a2, cases b2, auto simp: le_prod_def)
-    by(auto simp add: empty_def le_ivl_def le_option_def min_option_def max_option_def split: option.splits)
-qed
-
-
-subsubsection "Tests"
-
-value "show_acom_opt (AI_ivl test1_ivl)"
-
-text{* Better than @{text AI_const}: *}
-value "show_acom_opt (AI_ivl test3_const)"
-value "show_acom_opt (AI_ivl test4_const)"
-value "show_acom_opt (AI_ivl test6_const)"
-
-value "show_acom_opt (AI_ivl test2_ivl)"
-value "show_acom (((step_ivl \<top>)^^0) (\<bottom>\<^sub>c test2_ivl))"
-value "show_acom (((step_ivl \<top>)^^1) (\<bottom>\<^sub>c test2_ivl))"
-value "show_acom (((step_ivl \<top>)^^2) (\<bottom>\<^sub>c test2_ivl))"
-
-text{* Fixed point reached in 2 steps. Not so if the start value of x is known: *}
-
-value "show_acom_opt (AI_ivl test3_ivl)"
-value "show_acom (((step_ivl \<top>)^^0) (\<bottom>\<^sub>c test3_ivl))"
-value "show_acom (((step_ivl \<top>)^^1) (\<bottom>\<^sub>c test3_ivl))"
-value "show_acom (((step_ivl \<top>)^^2) (\<bottom>\<^sub>c test3_ivl))"
-value "show_acom (((step_ivl \<top>)^^3) (\<bottom>\<^sub>c test3_ivl))"
-value "show_acom (((step_ivl \<top>)^^4) (\<bottom>\<^sub>c test3_ivl))"
-
-text{* Takes as many iterations as the actual execution. Would diverge if
-loop did not terminate. Worse still, as the following example shows: even if
-the actual execution terminates, the analysis may not. The value of y keeps
-decreasing as the analysis is iterated, no matter how long: *}
-
-value "show_acom (((step_ivl \<top>)^^50) (\<bottom>\<^sub>c test4_ivl))"
-
-text{* Relationships between variables are NOT captured: *}
-value "show_acom_opt (AI_ivl test5_ivl)"
-
-text{* Again, the analysis would not terminate: *}
-value "show_acom (((step_ivl \<top>)^^50) (\<bottom>\<^sub>c test6_ivl))"
-
-end
--- a/src/HOL/IMP/Abs_Int2.thy	Thu Apr 19 12:28:10 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,683 +0,0 @@
-(* Author: Tobias Nipkow *)
-
-theory Abs_Int2
-imports Abs_Int1_ivl
-begin
-
-subsection "Widening and Narrowing"
-
-class WN = SL_top +
-fixes widen :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<nabla>" 65)
-assumes widen1: "x \<sqsubseteq> x \<nabla> y"
-assumes widen2: "y \<sqsubseteq> x \<nabla> y"
-fixes narrow :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<triangle>" 65)
-assumes narrow1: "y \<sqsubseteq> x \<Longrightarrow> y \<sqsubseteq> x \<triangle> y"
-assumes narrow2: "y \<sqsubseteq> x \<Longrightarrow> x \<triangle> y \<sqsubseteq> x"
-
-
-subsubsection "Intervals"
-
-instantiation ivl :: WN
-begin
-
-definition "widen_ivl ivl1 ivl2 =
-  ((*if is_empty ivl1 then ivl2 else
-   if is_empty ivl2 then ivl1 else*)
-     case (ivl1,ivl2) of (I l1 h1, I l2 h2) \<Rightarrow>
-       I (if le_option False l2 l1 \<and> l2 \<noteq> l1 then None else l1)
-         (if le_option True h1 h2 \<and> h1 \<noteq> h2 then None else h1))"
-
-definition "narrow_ivl ivl1 ivl2 =
-  ((*if is_empty ivl1 \<or> is_empty ivl2 then empty else*)
-     case (ivl1,ivl2) of (I l1 h1, I l2 h2) \<Rightarrow>
-       I (if l1 = None then l2 else l1)
-         (if h1 = None then h2 else h1))"
-
-instance
-proof qed
-  (auto simp add: widen_ivl_def narrow_ivl_def le_option_def le_ivl_def empty_def split: ivl.split option.split if_splits)
-
-end
-
-
-subsubsection "Abstract State"
-
-instantiation st :: (WN)WN
-begin
-
-definition "widen_st F1 F2 =
-  FunDom (\<lambda>x. fun F1 x \<nabla> fun F2 x) (inter_list (dom F1) (dom F2))"
-
-definition "narrow_st F1 F2 =
-  FunDom (\<lambda>x. fun F1 x \<triangle> fun F2 x) (inter_list (dom F1) (dom F2))"
-
-instance
-proof
-  case goal1 thus ?case
-    by(simp add: widen_st_def le_st_def lookup_def widen1)
-next
-  case goal2 thus ?case
-    by(simp add: widen_st_def le_st_def lookup_def widen2)
-next
-  case goal3 thus ?case
-    by(auto simp: narrow_st_def le_st_def lookup_def narrow1)
-next
-  case goal4 thus ?case
-    by(auto simp: narrow_st_def le_st_def lookup_def narrow2)
-qed
-
-end
-
-
-subsubsection "Option"
-
-instantiation option :: (WN)WN
-begin
-
-fun widen_option where
-"None \<nabla> x = x" |
-"x \<nabla> None = x" |
-"(Some x) \<nabla> (Some y) = Some(x \<nabla> y)"
-
-fun narrow_option where
-"None \<triangle> x = None" |
-"x \<triangle> None = None" |
-"(Some x) \<triangle> (Some y) = Some(x \<triangle> y)"
-
-instance
-proof
-  case goal1 show ?case
-    by(induct x y rule: widen_option.induct) (simp_all add: widen1)
-next
-  case goal2 show ?case
-    by(induct x y rule: widen_option.induct) (simp_all add: widen2)
-next
-  case goal3 thus ?case
-    by(induct x y rule: narrow_option.induct) (simp_all add: narrow1)
-next
-  case goal4 thus ?case
-    by(induct x y rule: narrow_option.induct) (simp_all add: narrow2)
-qed
-
-end
-
-
-subsubsection "Annotated commands"
-
-fun map2_acom :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a acom \<Rightarrow> 'a acom \<Rightarrow> 'a acom" where
-"map2_acom f (SKIP {a1}) (SKIP {a2}) = (SKIP {f a1 a2})" |
-"map2_acom f (x ::= e {a1}) (x' ::= e' {a2}) = (x ::= e {f a1 a2})" |
-"map2_acom f (c1;c2) (c1';c2') = (map2_acom f c1 c1'; map2_acom f c2 c2')" |
-"map2_acom f (IF b THEN c1 ELSE c2 {a1}) (IF b' THEN c1' ELSE c2' {a2}) =
-  (IF b THEN map2_acom f c1 c1' ELSE map2_acom f c2 c2' {f a1 a2})" |
-"map2_acom f ({a1} WHILE b DO c {a2}) ({a3} WHILE b' DO c' {a4}) =
-  ({f a1 a3} WHILE b DO map2_acom f c c' {f a2 a4})"
-
-abbreviation widen_acom :: "('a::WN)acom \<Rightarrow> 'a acom \<Rightarrow> 'a acom" (infix "\<nabla>\<^sub>c" 65)
-where "widen_acom == map2_acom (op \<nabla>)"
-
-abbreviation narrow_acom :: "('a::WN)acom \<Rightarrow> 'a acom \<Rightarrow> 'a acom" (infix "\<triangle>\<^sub>c" 65)
-where "narrow_acom == map2_acom (op \<triangle>)"
-
-lemma widen1_acom: "strip c = strip c' \<Longrightarrow> c \<sqsubseteq> c \<nabla>\<^sub>c c'"
-by(induct c c' rule: le_acom.induct)(simp_all add: widen1)
-
-lemma widen2_acom: "strip c = strip c' \<Longrightarrow> c' \<sqsubseteq> c \<nabla>\<^sub>c c'"
-by(induct c c' rule: le_acom.induct)(simp_all add: widen2)
-
-lemma narrow1_acom: "y \<sqsubseteq> x \<Longrightarrow> y \<sqsubseteq> x \<triangle>\<^sub>c y"
-by(induct y x rule: le_acom.induct) (simp_all add: narrow1)
-
-lemma narrow2_acom: "y \<sqsubseteq> x \<Longrightarrow> x \<triangle>\<^sub>c y \<sqsubseteq> x"
-by(induct y x rule: le_acom.induct) (simp_all add: narrow2)
-
-
-subsubsection "Post-fixed point computation"
-
-definition iter_widen :: "('a acom \<Rightarrow> 'a acom) \<Rightarrow> 'a acom \<Rightarrow> ('a::WN)acom option"
-where "iter_widen f = while_option (\<lambda>c. \<not> f c \<sqsubseteq> c) (\<lambda>c. c \<nabla>\<^sub>c f c)"
-
-definition iter_narrow :: "('a acom \<Rightarrow> 'a acom) \<Rightarrow> 'a acom \<Rightarrow> 'a::WN acom option"
-where "iter_narrow f = while_option (\<lambda>c. \<not> c \<sqsubseteq> c \<triangle>\<^sub>c f c) (\<lambda>c. c \<triangle>\<^sub>c f c)"
-
-definition pfp_wn ::
-  "(('a::WN)option acom \<Rightarrow> 'a option acom) \<Rightarrow> com \<Rightarrow> 'a option acom option"
-where "pfp_wn f c = (case iter_widen f (\<bottom>\<^sub>c c) of None \<Rightarrow> None
-                     | Some c' \<Rightarrow> iter_narrow f c')"
-
-lemma strip_map2_acom:
- "strip c1 = strip c2 \<Longrightarrow> strip(map2_acom f c1 c2) = strip c1"
-by(induct f c1 c2 rule: map2_acom.induct) simp_all
-
-lemma iter_widen_pfp: "iter_widen f c = Some c' \<Longrightarrow> f c' \<sqsubseteq> c'"
-by(auto simp add: iter_widen_def dest: while_option_stop)
-
-lemma strip_while: fixes f :: "'a acom \<Rightarrow> 'a acom"
-assumes "\<forall>c. strip (f c) = strip c" and "while_option P f c = Some c'"
-shows "strip c' = strip c"
-using while_option_rule[where P = "\<lambda>c'. strip c' = strip c", OF _ assms(2)]
-by (metis assms(1))
-
-lemma strip_iter_widen: fixes f :: "'a::WN acom \<Rightarrow> 'a acom"
-assumes "\<forall>c. strip (f c) = strip c" and "iter_widen f c = Some c'"
-shows "strip c' = strip c"
-proof-
-  have "\<forall>c. strip(c \<nabla>\<^sub>c f c) = strip c" by (metis assms(1) strip_map2_acom)
-  from strip_while[OF this] assms(2) show ?thesis by(simp add: iter_widen_def)
-qed
-
-lemma iter_narrow_pfp: assumes "mono f" and "f c0 \<sqsubseteq> c0"
-and "iter_narrow f c0 = Some c"
-shows "f c \<sqsubseteq> c \<and> c \<sqsubseteq> c0" (is "?P c")
-proof-
-  { fix c assume "?P c"
-    note 1 = conjunct1[OF this] and 2 = conjunct2[OF this]
-    let ?c' = "c \<triangle>\<^sub>c f c"
-    have "?P ?c'"
-    proof
-      have "f ?c' \<sqsubseteq> f c"  by(rule monoD[OF `mono f` narrow2_acom[OF 1]])
-      also have "\<dots> \<sqsubseteq> ?c'" by(rule narrow1_acom[OF 1])
-      finally show "f ?c' \<sqsubseteq> ?c'" .
-      have "?c' \<sqsubseteq> c" by (rule narrow2_acom[OF 1])
-      also have "c \<sqsubseteq> c0" by(rule 2)
-      finally show "?c' \<sqsubseteq> c0" .
-    qed
-  }
-  with while_option_rule[where P = ?P, OF _ assms(3)[simplified iter_narrow_def]]
-    assms(2) le_refl
-  show ?thesis by blast
-qed
-
-lemma pfp_wn_pfp:
-  "\<lbrakk> mono f;  pfp_wn f c = Some c' \<rbrakk> \<Longrightarrow> f c' \<sqsubseteq> c'"
-unfolding pfp_wn_def
-by (auto dest: iter_widen_pfp iter_narrow_pfp split: option.splits)
-
-lemma strip_pfp_wn:
-  "\<lbrakk> \<forall>c. strip(f c) = strip c; pfp_wn f c = Some c' \<rbrakk> \<Longrightarrow> strip c' = c"
-apply(auto simp add: pfp_wn_def iter_narrow_def split: option.splits)
-by (metis (no_types) strip_map2_acom strip_while strip_bot_acom strip_iter_widen)
-
-locale Abs_Int2 = Abs_Int1_mono
-where \<gamma>=\<gamma> for \<gamma> :: "'av::{WN,L_top_bot} \<Rightarrow> val set"
-begin
-
-definition AI_wn :: "com \<Rightarrow> 'av st option acom option" where
-"AI_wn = pfp_wn (step' \<top>)"
-
-lemma AI_wn_sound: "AI_wn c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
-proof(simp add: CS_def AI_wn_def)
-  assume 1: "pfp_wn (step' \<top>) c = Some c'"
-  from pfp_wn_pfp[OF mono_step'2 1]
-  have 2: "step' \<top> c' \<sqsubseteq> c'" .
-  have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c" by(simp add: strip_pfp_wn[OF _ 1])
-  have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
-  proof(rule lfp_lowerbound[simplified,OF 3])
-    show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
-    proof(rule step_preserves_le[OF _ _])
-      show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
-      show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
-    qed
-  qed
-  from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
-    by (blast intro: mono_gamma_c order_trans)
-qed
-
-end
-
-interpretation Abs_Int2
-where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = plus_ivl
-and test_num' = in_ivl
-and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl
-defines AI_ivl' is AI_wn
-..
-
-
-subsubsection "Tests"
-
-definition "step_up_ivl n = ((\<lambda>c. c \<nabla>\<^sub>c step_ivl \<top> c)^^n)"
-definition "step_down_ivl n = ((\<lambda>c. c \<triangle>\<^sub>c step_ivl \<top> c)^^n)"
-
-text{* For @{const test3_ivl}, @{const AI_ivl} needed as many iterations as
-the loop took to execute. In contrast, @{const AI_ivl'} converges in a
-constant number of steps: *}
-
-value "show_acom (step_up_ivl 1 (\<bottom>\<^sub>c test3_ivl))"
-value "show_acom (step_up_ivl 2 (\<bottom>\<^sub>c test3_ivl))"
-value "show_acom (step_up_ivl 3 (\<bottom>\<^sub>c test3_ivl))"
-value "show_acom (step_up_ivl 4 (\<bottom>\<^sub>c test3_ivl))"
-value "show_acom (step_up_ivl 5 (\<bottom>\<^sub>c test3_ivl))"
-value "show_acom (step_down_ivl 1 (step_up_ivl 5 (\<bottom>\<^sub>c test3_ivl)))"
-value "show_acom (step_down_ivl 2 (step_up_ivl 5 (\<bottom>\<^sub>c test3_ivl)))"
-value "show_acom (step_down_ivl 3 (step_up_ivl 5 (\<bottom>\<^sub>c test3_ivl)))"
-
-text{* Now all the analyses terminate: *}
-
-value "show_acom_opt (AI_ivl' test4_ivl)"
-value "show_acom_opt (AI_ivl' test5_ivl)"
-value "show_acom_opt (AI_ivl' test6_ivl)"
-
-
-subsubsection "Termination: Intervals"
-
-definition m_ivl :: "ivl \<Rightarrow> nat" where
-"m_ivl ivl = (case ivl of I l h \<Rightarrow>
-     (case l of None \<Rightarrow> 0 | Some _ \<Rightarrow> 1) + (case h of None \<Rightarrow> 0 | Some _ \<Rightarrow> 1))"
-
-lemma m_ivl_height: "m_ivl ivl \<le> 2"
-by(simp add: m_ivl_def split: ivl.split option.split)
-
-lemma m_ivl_anti_mono: "(y::ivl) \<sqsubseteq> x \<Longrightarrow> m_ivl x \<le> m_ivl y"
-by(auto simp: m_ivl_def le_option_def le_ivl_def
-        split: ivl.split option.split if_splits)
-
-lemma m_ivl_widen:
-  "~ y \<sqsubseteq> x \<Longrightarrow> m_ivl(x \<nabla> y) < m_ivl x"
-by(auto simp: m_ivl_def widen_ivl_def le_option_def le_ivl_def
-        split: ivl.splits option.splits if_splits)
-
-lemma Top_less_ivl: "\<top> \<sqsubseteq> x \<Longrightarrow> m_ivl x = 0"
-by(auto simp: m_ivl_def le_option_def le_ivl_def empty_def Top_ivl_def
-        split: ivl.split option.split if_splits)
-
-
-definition n_ivl :: "ivl \<Rightarrow> nat" where
-"n_ivl ivl = 2 - m_ivl ivl"
-
-lemma n_ivl_mono: "(x::ivl) \<sqsubseteq> y \<Longrightarrow> n_ivl x \<le> n_ivl y"
-unfolding n_ivl_def by (metis diff_le_mono2 m_ivl_anti_mono)
-
-lemma n_ivl_narrow:
-  "~ x \<sqsubseteq> x \<triangle> y \<Longrightarrow> n_ivl(x \<triangle> y) < n_ivl x"
-by(auto simp: n_ivl_def m_ivl_def narrow_ivl_def le_option_def le_ivl_def
-        split: ivl.splits option.splits if_splits)
-
-
-subsubsection "Termination: Abstract State"
-
-definition "m_st m st = (\<Sum>x\<in>set(dom st). m(fun st x))"
-
-lemma m_st_height: assumes "finite X" and "set (dom S) \<subseteq> X"
-shows "m_st m_ivl S \<le> 2 * card X"
-proof(auto simp: m_st_def)
-  have "(\<Sum>x\<in>set(dom S). m_ivl (fun S x)) \<le> (\<Sum>x\<in>set(dom S). 2)" (is "?L \<le> _")
-    by(rule setsum_mono)(simp add:m_ivl_height)
-  also have "\<dots> \<le> (\<Sum>x\<in>X. 2)"
-    by(rule setsum_mono3[OF assms]) simp
-  also have "\<dots> = 2 * card X" by(simp add: setsum_constant)
-  finally show "?L \<le> \<dots>" .
-qed
-
-lemma m_st_anti_mono:
-  "S1 \<sqsubseteq> S2 \<Longrightarrow> m_st m_ivl S2 \<le> m_st m_ivl S1"
-proof(auto simp: m_st_def le_st_def lookup_def split: if_splits)
-  let ?X = "set(dom S1)" let ?Y = "set(dom S2)"
-  let ?f = "fun S1" let ?g = "fun S2"
-  assume asm: "\<forall>x\<in>?Y. (x \<in> ?X \<longrightarrow> ?f x \<sqsubseteq> ?g x) \<and> (x \<in> ?X \<or> \<top> \<sqsubseteq> ?g x)"
-  hence 1: "\<forall>y\<in>?Y\<inter>?X. m_ivl(?g y) \<le> m_ivl(?f y)" by(simp add: m_ivl_anti_mono)
-  have 0: "\<forall>x\<in>?Y-?X. m_ivl(?g x) = 0" using asm by (auto simp: Top_less_ivl)
-  have "(\<Sum>y\<in>?Y. m_ivl(?g y)) = (\<Sum>y\<in>(?Y-?X) \<union> (?Y\<inter>?X). m_ivl(?g y))"
-    by (metis Un_Diff_Int)
-  also have "\<dots> = (\<Sum>y\<in>?Y-?X. m_ivl(?g y)) + (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?g y))"
-    by(subst setsum_Un_disjoint) auto
-  also have "(\<Sum>y\<in>?Y-?X. m_ivl(?g y)) = 0" using 0 by simp
-  also have "0 + (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?g y)) = (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?g y))" by simp
-  also have "\<dots> \<le> (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?f y))"
-    by(rule setsum_mono)(simp add: 1)
-  also have "\<dots> \<le> (\<Sum>y\<in>?X. m_ivl(?f y))"
-    by(simp add: setsum_mono3[of "?X" "?Y Int ?X", OF _ Int_lower2])
-  finally show "(\<Sum>y\<in>?Y. m_ivl(?g y)) \<le> (\<Sum>x\<in>?X. m_ivl(?f x))"
-    by (metis add_less_cancel_left)
-qed
-
-lemma m_st_widen:
-assumes "\<not> S2 \<sqsubseteq> S1" shows "m_st m_ivl (S1 \<nabla> S2) < m_st m_ivl S1"
-proof-
-  { let ?X = "set(dom S1)" let ?Y = "set(dom S2)"
-    let ?f = "fun S1" let ?g = "fun S2"
-    fix x assume "x \<in> ?X" "\<not> lookup S2 x \<sqsubseteq> ?f x"
-    have "(\<Sum>x\<in>?X\<inter>?Y. m_ivl(?f x \<nabla> ?g x)) < (\<Sum>x\<in>?X. m_ivl(?f x))" (is "?L < ?R")
-    proof cases
-      assume "x : ?Y"
-      have "?L < (\<Sum>x\<in>?X\<inter>?Y. m_ivl(?f x))"
-      proof(rule setsum_strict_mono1, simp)
-        show "\<forall>x\<in>?X\<inter>?Y. m_ivl(?f x \<nabla> ?g x) \<le> m_ivl (?f x)"
-          by (metis m_ivl_anti_mono widen1)
-      next
-        show "\<exists>x\<in>?X\<inter>?Y. m_ivl(?f x \<nabla> ?g x) < m_ivl(?f x)"
-          using `x:?X` `x:?Y` `\<not> lookup S2 x \<sqsubseteq> ?f x`
-          by (metis IntI m_ivl_widen lookup_def)
-      qed
-      also have "\<dots> \<le> ?R" by(simp add: setsum_mono3[OF _ Int_lower1])
-      finally show ?thesis .
-    next
-      assume "x ~: ?Y"
-      have "?L \<le> (\<Sum>x\<in>?X\<inter>?Y. m_ivl(?f x))"
-      proof(rule setsum_mono, simp)
-        fix x assume "x:?X \<and> x:?Y" show "m_ivl(?f x \<nabla> ?g x) \<le> m_ivl (?f x)"
-          by (metis m_ivl_anti_mono widen1)
-      qed
-      also have "\<dots> < m_ivl(?f x) + \<dots>"
-        using m_ivl_widen[OF `\<not> lookup S2 x \<sqsubseteq> ?f x`]
-        by (metis Nat.le_refl add_strict_increasing gr0I not_less0)
-      also have "\<dots> = (\<Sum>y\<in>insert x (?X\<inter>?Y). m_ivl(?f y))"
-        using `x ~: ?Y` by simp
-      also have "\<dots> \<le> (\<Sum>x\<in>?X. m_ivl(?f x))"
-        by(rule setsum_mono3)(insert `x:?X`, auto)
-      finally show ?thesis .
-    qed
-  } with assms show ?thesis
-    by(auto simp: le_st_def widen_st_def m_st_def Int_def)
-qed
-
-definition "n_st m X st = (\<Sum>x\<in>X. m(lookup st x))"
-
-lemma n_st_mono: assumes "set(dom S1) \<subseteq> X" "set(dom S2) \<subseteq> X" "S1 \<sqsubseteq> S2"
-shows "n_st n_ivl X S1 \<le> n_st n_ivl X S2"
-proof-
-  have "(\<Sum>x\<in>X. n_ivl(lookup S1 x)) \<le> (\<Sum>x\<in>X. n_ivl(lookup S2 x))"
-    apply(rule setsum_mono) using assms
-    by(auto simp: le_st_def lookup_def n_ivl_mono split: if_splits)
-  thus ?thesis by(simp add: n_st_def)
-qed
-
-lemma n_st_narrow:
-assumes "finite X" and "set(dom S1) \<subseteq> X" "set(dom S2) \<subseteq> X"
-and "S2 \<sqsubseteq> S1" "\<not> S1 \<sqsubseteq> S1 \<triangle> S2"
-shows "n_st n_ivl X (S1 \<triangle> S2) < n_st n_ivl X S1"
-proof-
-  have 1: "\<forall>x\<in>X. n_ivl (lookup (S1 \<triangle> S2) x) \<le> n_ivl (lookup S1 x)"
-    using assms(2-4)
-    by(auto simp: le_st_def narrow_st_def lookup_def n_ivl_mono narrow2
-            split: if_splits)
-  have 2: "\<exists>x\<in>X. n_ivl (lookup (S1 \<triangle> S2) x) < n_ivl (lookup S1 x)"
-    using assms(2-5)
-    by(auto simp: le_st_def narrow_st_def lookup_def intro: n_ivl_narrow
-            split: if_splits)
-  have "(\<Sum>x\<in>X. n_ivl(lookup (S1 \<triangle> S2) x)) < (\<Sum>x\<in>X. n_ivl(lookup S1 x))"
-    apply(rule setsum_strict_mono1[OF `finite X`]) using 1 2 by blast+
-  thus ?thesis by(simp add: n_st_def)
-qed
-
-
-subsubsection "Termination: Option"
-
-definition "m_o m n opt = (case opt of None \<Rightarrow> n+1 | Some x \<Rightarrow> m x)"
-
-lemma m_o_anti_mono: "finite X \<Longrightarrow> domo S2 \<subseteq> X \<Longrightarrow> S1 \<sqsubseteq> S2 \<Longrightarrow>
-  m_o (m_st m_ivl) (2 * card X) S2 \<le> m_o (m_st m_ivl) (2 * card X) S1"
-apply(induction S1 S2 rule: le_option.induct)
-apply(auto simp: domo_def m_o_def m_st_anti_mono le_SucI m_st_height
-           split: option.splits)
-done
-
-lemma m_o_widen: "\<lbrakk> finite X; domo S2 \<subseteq> X; \<not> S2 \<sqsubseteq> S1 \<rbrakk> \<Longrightarrow>
-  m_o (m_st m_ivl) (2 * card X) (S1 \<nabla> S2) < m_o (m_st m_ivl) (2 * card X) S1"
-by(auto simp: m_o_def domo_def m_st_height less_Suc_eq_le m_st_widen
-        split: option.split)
-
-definition "n_o n opt = (case opt of None \<Rightarrow> 0 | Some x \<Rightarrow> n x + 1)"
-
-lemma n_o_mono: "domo S1 \<subseteq> X \<Longrightarrow> domo S2 \<subseteq> X \<Longrightarrow> S1 \<sqsubseteq> S2 \<Longrightarrow>
-  n_o (n_st n_ivl X) S1 \<le> n_o (n_st n_ivl X) S2"
-apply(induction S1 S2 rule: le_option.induct)
-apply(auto simp: domo_def n_o_def n_st_mono
-           split: option.splits)
-done
-
-lemma n_o_narrow:
-  "\<lbrakk> finite X; domo S1 \<subseteq> X; domo S2 \<subseteq> X; S2 \<sqsubseteq> S1; \<not> S1 \<sqsubseteq> S1 \<triangle> S2 \<rbrakk>
-  \<Longrightarrow> n_o (n_st n_ivl X) (S1 \<triangle> S2) < n_o (n_st n_ivl X) S1"
-apply(induction S1 S2 rule: narrow_option.induct)
-apply(auto simp: n_o_def domo_def n_st_narrow)
-done
-
-lemma domo_widen_subset: "domo (S1 \<nabla> S2) \<subseteq> domo S1 \<union> domo S2"
-apply(induction S1 S2 rule: widen_option.induct)
-apply (auto simp: domo_def widen_st_def)
-done
-
-lemma domo_narrow_subset: "domo (S1 \<triangle> S2) \<subseteq> domo S1 \<union> domo S2"
-apply(induction S1 S2 rule: narrow_option.induct)
-apply (auto simp: domo_def narrow_st_def)
-done
-
-subsubsection "Termination: Commands"
-
-lemma strip_widen_acom[simp]:
-  "strip c' = strip (c::'a::WN acom) \<Longrightarrow>  strip (c \<nabla>\<^sub>c c') = strip c"
-by(induction "widen::'a\<Rightarrow>'a\<Rightarrow>'a" c c' rule: map2_acom.induct) simp_all
-
-lemma strip_narrow_acom[simp]:
-  "strip c' = strip (c::'a::WN acom) \<Longrightarrow>  strip (c \<triangle>\<^sub>c c') = strip c"
-by(induction "narrow::'a\<Rightarrow>'a\<Rightarrow>'a" c c' rule: map2_acom.induct) simp_all
-
-lemma annos_widen_acom[simp]: "strip c1 = strip (c2::'a::WN acom) \<Longrightarrow>
-  annos(c1 \<nabla>\<^sub>c c2) = map (%(x,y).x\<nabla>y) (zip (annos c1) (annos(c2::'a::WN acom)))"
-by(induction "widen::'a\<Rightarrow>'a\<Rightarrow>'a" c1 c2 rule: map2_acom.induct)
-  (simp_all add: size_annos_same2)
-
-lemma annos_narrow_acom[simp]: "strip c1 = strip (c2::'a::WN acom) \<Longrightarrow>
-  annos(c1 \<triangle>\<^sub>c c2) = map (%(x,y).x\<triangle>y) (zip (annos c1) (annos(c2::'a::WN acom)))"
-by(induction "narrow::'a\<Rightarrow>'a\<Rightarrow>'a" c1 c2 rule: map2_acom.induct)
-  (simp_all add: size_annos_same2)
-
-lemma widen_acom_Com[simp]: "strip c2 = strip c1 \<Longrightarrow>
-  c1 : Com X \<Longrightarrow> c2 : Com X \<Longrightarrow> (c1 \<nabla>\<^sub>c c2) : Com X"
-apply(auto simp add: Com_def)
-apply(rename_tac S S' x)
-apply(erule in_set_zipE)
-apply(auto simp: domo_def split: option.splits)
-apply(case_tac S)
-apply(case_tac S')
-apply simp
-apply fastforce
-apply(case_tac S')
-apply fastforce
-apply (fastforce simp: widen_st_def)
-done
-
-lemma narrow_acom_Com[simp]: "strip c2 = strip c1 \<Longrightarrow>
-  c1 : Com X \<Longrightarrow> c2 : Com X \<Longrightarrow> (c1 \<triangle>\<^sub>c c2) : Com X"
-apply(auto simp add: Com_def)
-apply(rename_tac S S' x)
-apply(erule in_set_zipE)
-apply(auto simp: domo_def split: option.splits)
-apply(case_tac S)
-apply(case_tac S')
-apply simp
-apply fastforce
-apply(case_tac S')
-apply fastforce
-apply (fastforce simp: narrow_st_def)
-done
-
-definition "m_c m c = (let as = annos c in \<Sum>i=0..<size as. m(as!i))"
-
-lemma measure_m_c: "finite X \<Longrightarrow> {(c, c \<nabla>\<^sub>c c') |c c'\<Colon>ivl st option acom.
-     strip c' = strip c \<and> c : Com X \<and> c' : Com X \<and> \<not> c' \<sqsubseteq> c}\<inverse>
-    \<subseteq> measure(m_c(m_o (m_st m_ivl) (2*card(X))))"
-apply(auto simp: m_c_def Let_def Com_def)
-apply(subgoal_tac "length(annos c') = length(annos c)")
-prefer 2 apply (simp add: size_annos_same2)
-apply (auto)
-apply(rule setsum_strict_mono1)
-apply simp
-apply (clarsimp)
-apply(erule m_o_anti_mono)
-apply(rule subset_trans[OF domo_widen_subset])
-apply fastforce
-apply(rule widen1)
-apply(auto simp: le_iff_le_annos listrel_iff_nth)
-apply(rule_tac x=n in bexI)
-prefer 2 apply simp
-apply(erule m_o_widen)
-apply (simp)+
-done
-
-lemma measure_n_c: "finite X \<Longrightarrow> {(c, c \<triangle>\<^sub>c c') |c c'.
-  strip c = strip c' \<and> c \<in> Com X \<and> c' \<in> Com X \<and> c' \<sqsubseteq> c \<and> \<not> c \<sqsubseteq> c \<triangle>\<^sub>c c'}\<inverse>
-  \<subseteq> measure(m_c(n_o (n_st n_ivl X)))"
-apply(auto simp: m_c_def Let_def Com_def)
-apply(subgoal_tac "length(annos c') = length(annos c)")
-prefer 2 apply (simp add: size_annos_same2)
-apply (auto)
-apply(rule setsum_strict_mono1)
-apply simp
-apply (clarsimp)
-apply(rule n_o_mono)
-using domo_narrow_subset apply fastforce
-apply fastforce
-apply(rule narrow2)
-apply(fastforce simp: le_iff_le_annos listrel_iff_nth)
-apply(auto simp: le_iff_le_annos listrel_iff_nth strip_narrow_acom)
-apply(rule_tac x=n in bexI)
-prefer 2 apply simp
-apply(erule n_o_narrow)
-apply (simp)+
-done
-
-
-subsubsection "Termination: Post-Fixed Point Iterations"
-
-lemma iter_widen_termination:
-fixes c0 :: "'a::WN acom"
-assumes P_f: "\<And>c. P c \<Longrightarrow> P(f c)"
-assumes P_widen: "\<And>c c'. P c \<Longrightarrow> P c' \<Longrightarrow> P(c \<nabla>\<^sub>c c')"
-and "wf({(c::'a acom,c \<nabla>\<^sub>c c')|c c'. P c \<and> P c' \<and> ~ c' \<sqsubseteq> c}^-1)"
-and "P c0" and "c0 \<sqsubseteq> f c0" shows "EX c. iter_widen f c0 = Some c"
-proof(simp add: iter_widen_def, rule wf_while_option_Some[where P = "P"])
-  show "wf {(cc', c). (P c \<and> \<not> f c \<sqsubseteq> c) \<and> cc' = c \<nabla>\<^sub>c f c}"
-    apply(rule wf_subset[OF assms(3)]) by(blast intro: P_f)
-next
-  show "P c0" by(rule `P c0`)
-next
-  fix c assume "P c" thus "P (c \<nabla>\<^sub>c f c)" by(simp add: P_f P_widen)
-qed
-
-lemma iter_narrow_termination:
-assumes P_f: "\<And>c. P c \<Longrightarrow> P(c \<triangle>\<^sub>c f c)"
-and wf: "wf({(c, c \<triangle>\<^sub>c f c)|c c'. P c \<and> ~ c \<sqsubseteq> c \<triangle>\<^sub>c f c}^-1)"
-and "P c0" shows "EX c. iter_narrow f c0 = Some c"
-proof(simp add: iter_narrow_def, rule wf_while_option_Some[where P = "P"])
-  show "wf {(c', c). (P c \<and> \<not> c \<sqsubseteq> c \<triangle>\<^sub>c f c) \<and> c' = c \<triangle>\<^sub>c f c}"
-    apply(rule wf_subset[OF wf]) by(blast intro: P_f)
-next
-  show "P c0" by(rule `P c0`)
-next
-  fix c assume "P c" thus "P (c \<triangle>\<^sub>c f c)" by(simp add: P_f)
-qed
-
-lemma iter_winden_step_ivl_termination:
-  "EX c. iter_widen (step_ivl \<top>) (\<bottom>\<^sub>c c0) = Some c"
-apply(rule iter_widen_termination[where
-  P = "%c. strip c = c0 \<and> c : Com(vars c0)"])
-apply (simp_all add: step'_Com bot_acom)
-apply(rule wf_subset)
-apply(rule wf_measure)
-apply(rule subset_trans)
-prefer 2
-apply(rule measure_m_c[where X = "vars c0", OF finite_cvars])
-apply blast
-done
-
-lemma iter_narrow_step_ivl_termination:
-  "c0 \<in> Com (vars(strip c0)) \<Longrightarrow> step_ivl \<top> c0 \<sqsubseteq> c0 \<Longrightarrow>
-  EX c. iter_narrow (step_ivl \<top>) c0 = Some c"
-apply(rule iter_narrow_termination[where
-  P = "%c. strip c = strip c0 \<and> c : Com(vars(strip c0)) \<and> step_ivl \<top> c \<sqsubseteq> c"])
-apply (simp_all add: step'_Com)
-apply(clarify)
-apply(frule narrow2_acom, drule mono_step'[OF le_refl], erule le_trans[OF _ narrow1_acom])
-apply assumption
-apply(rule wf_subset)
-apply(rule wf_measure)
-apply(rule subset_trans)
-prefer 2
-apply(rule measure_n_c[where X = "vars(strip c0)", OF finite_cvars])
-apply auto
-by (metis bot_least domo_Top order_refl step'_Com strip_step')
-
-(* FIXME: simplify type system: Combine Com(X) and vars <= X?? *)
-lemma while_Com:
-fixes c :: "'a st option acom"
-assumes "while_option P f c = Some c'"
-and "!!c. strip(f c) = strip c"
-and "\<forall>c::'a st option acom. c : Com(X) \<longrightarrow> vars(strip c) \<subseteq> X \<longrightarrow> f c : Com(X)"
-and "c : Com(X)" and "vars(strip c) \<subseteq> X" shows "c' : Com(X)"
-using while_option_rule[where P = "\<lambda>c'. c' : Com(X) \<and> vars(strip c') \<subseteq> X", OF _ assms(1)]
-by(simp add: assms(2-))
-
-lemma iter_widen_Com: fixes f :: "'a::WN st option acom \<Rightarrow> 'a st option acom"
-assumes "iter_widen f c = Some c'"
-and "\<forall>c. c : Com(X) \<longrightarrow> vars(strip c) \<subseteq> X \<longrightarrow> f c : Com(X)"
-and "!!c. strip(f c) = strip c"
-and "c : Com(X)" and "vars (strip c) \<subseteq> X" shows "c' : Com(X)"
-proof-
-  have "\<forall>c. c : Com(X) \<longrightarrow> vars(strip c) \<subseteq> X \<longrightarrow> c \<nabla>\<^sub>c f c : Com(X)"
-    by (metis (full_types) widen_acom_Com assms(2,3))
-  from while_Com[OF assms(1)[simplified iter_widen_def] _ this assms(4,5)]
-  show ?thesis using assms(3) by(simp)
-qed
-
-
-context Abs_Int2
-begin
-
-lemma iter_widen_step'_Com:
-  "iter_widen (step' \<top>) c = Some c' \<Longrightarrow> vars(strip c) \<subseteq> X \<Longrightarrow> c : Com(X)
-   \<Longrightarrow> c' : Com(X)"
-apply(subgoal_tac "strip c'= strip c")
-prefer 2 apply (metis strip_iter_widen strip_step')
-apply(drule iter_widen_Com)
-prefer 3 apply assumption
-prefer 3 apply assumption
-apply (auto simp: step'_Com)
-done
-
-end
-
-theorem AI_ivl'_termination:
-  "EX c'. AI_ivl' c = Some c'"
-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 (metis bot_acom_Com iter_widen_step'_Com[OF _ subset_refl] strip_iter_widen strip_step')
-apply(erule iter_widen_pfp)
-done
-
-end
-
-
-(* interesting(?) relic
-lemma widen_assoc:
-  "~ (y::ivl) \<sqsubseteq> x \<Longrightarrow> ~ z \<sqsubseteq> x \<nabla> y \<Longrightarrow> ((x::ivl) \<nabla> y) \<nabla> z = x \<nabla> (y \<nabla> z)"
-apply(cases x)
-apply(cases y)
-apply(cases z)
-apply(rename_tac x1 x2 y1 y2 z1 z2)
-apply(simp add: le_ivl_def)
-apply(case_tac x1)
-apply(case_tac x2)
-apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits)
-apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits)
-apply(case_tac x2)
-apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits)
-apply(case_tac y1)
-apply(case_tac y2)
-apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits)
-apply(case_tac z1)
-apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1]
-apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1]
-apply(case_tac y2)
-apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1]
-apply(case_tac z1)
-apply(auto simp add:le_option_def widen_ivl_def split: if_splits ivl.splits option.splits)[1]
-apply(case_tac z2)
-apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits)[1]
-apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits)[1]
-done
-
-lemma widen_step_trans:
-  "~ (y::ivl) \<sqsubseteq> x \<Longrightarrow> ~ z \<sqsubseteq> x \<nabla> y \<Longrightarrow> EX u. (x \<nabla> y) \<nabla> z = x \<nabla> u \<and> ~ u \<sqsubseteq> x"
-by (metis widen_assoc preord_class.le_trans widen1)
-*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int0_ITP.thy	Thu Apr 19 17:32:30 2012 +0200
@@ -0,0 +1,397 @@
+(* Author: Tobias Nipkow *)
+
+theory Abs_Int0_ITP
+imports "~~/src/HOL/ex/Interpretation_with_Defs"
+        "~~/src/HOL/Library/While_Combinator"
+        Collecting
+begin
+
+subsection "Orderings"
+
+class preord =
+fixes le :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubseteq>" 50)
+assumes le_refl[simp]: "x \<sqsubseteq> x"
+and le_trans: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z"
+begin
+
+definition mono where "mono f = (\<forall>x y. x \<sqsubseteq> y \<longrightarrow> f x \<sqsubseteq> f y)"
+
+lemma monoD: "mono f \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" by(simp add: mono_def)
+
+lemma mono_comp: "mono f \<Longrightarrow> mono g \<Longrightarrow> mono (g o f)"
+by(simp add: mono_def)
+
+declare le_trans[trans]
+
+end
+
+text{* Note: no antisymmetry. Allows implementations where some abstract
+element is implemented by two different values @{prop "x \<noteq> y"}
+such that @{prop"x \<sqsubseteq> y"} and @{prop"y \<sqsubseteq> x"}. Antisymmetry is not
+needed because we never compare elements for equality but only for @{text"\<sqsubseteq>"}.
+*}
+
+class SL_top = preord +
+fixes join :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<squnion>" 65)
+fixes Top :: "'a" ("\<top>")
+assumes join_ge1 [simp]: "x \<sqsubseteq> x \<squnion> y"
+and join_ge2 [simp]: "y \<sqsubseteq> x \<squnion> y"
+and join_least: "x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<squnion> y \<sqsubseteq> z"
+and top[simp]: "x \<sqsubseteq> \<top>"
+begin
+
+lemma join_le_iff[simp]: "x \<squnion> y \<sqsubseteq> z \<longleftrightarrow> x \<sqsubseteq> z \<and> y \<sqsubseteq> z"
+by (metis join_ge1 join_ge2 join_least le_trans)
+
+lemma le_join_disj: "x \<sqsubseteq> y \<or> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<squnion> z"
+by (metis join_ge1 join_ge2 le_trans)
+
+end
+
+instantiation "fun" :: (type, SL_top) SL_top
+begin
+
+definition "f \<sqsubseteq> g = (ALL x. f x \<sqsubseteq> g x)"
+definition "f \<squnion> g = (\<lambda>x. f x \<squnion> g x)"
+definition "\<top> = (\<lambda>x. \<top>)"
+
+lemma join_apply[simp]: "(f \<squnion> g) x = f x \<squnion> g x"
+by (simp add: join_fun_def)
+
+instance
+proof
+  case goal2 thus ?case by (metis le_fun_def preord_class.le_trans)
+qed (simp_all add: le_fun_def Top_fun_def)
+
+end
+
+
+instantiation acom :: (preord) preord
+begin
+
+fun le_acom :: "('a::preord)acom \<Rightarrow> 'a acom \<Rightarrow> bool" where
+"le_acom (SKIP {S}) (SKIP {S'}) = (S \<sqsubseteq> S')" |
+"le_acom (x ::= e {S}) (x' ::= e' {S'}) = (x=x' \<and> e=e' \<and> S \<sqsubseteq> S')" |
+"le_acom (c1;c2) (c1';c2') = (le_acom c1 c1' \<and> le_acom c2 c2')" |
+"le_acom (IF b THEN c1 ELSE c2 {S}) (IF b' THEN c1' ELSE c2' {S'}) =
+  (b=b' \<and> le_acom c1 c1' \<and> le_acom c2 c2' \<and> S \<sqsubseteq> S')" |
+"le_acom ({Inv} WHILE b DO c {P}) ({Inv'} WHILE b' DO c' {P'}) =
+  (b=b' \<and> le_acom c c' \<and> Inv \<sqsubseteq> Inv' \<and> P \<sqsubseteq> P')" |
+"le_acom _ _ = False"
+
+lemma [simp]: "SKIP {S} \<sqsubseteq> c \<longleftrightarrow> (\<exists>S'. c = SKIP {S'} \<and> S \<sqsubseteq> S')"
+by (cases c) auto
+
+lemma [simp]: "x ::= e {S} \<sqsubseteq> c \<longleftrightarrow> (\<exists>S'. c = x ::= e {S'} \<and> S \<sqsubseteq> S')"
+by (cases c) auto
+
+lemma [simp]: "c1;c2 \<sqsubseteq> c \<longleftrightarrow> (\<exists>c1' c2'. c = c1';c2' \<and> c1 \<sqsubseteq> c1' \<and> c2 \<sqsubseteq> c2')"
+by (cases c) auto
+
+lemma [simp]: "IF b THEN c1 ELSE c2 {S} \<sqsubseteq> c \<longleftrightarrow>
+  (\<exists>c1' c2' S'. c = IF b THEN c1' ELSE c2' {S'} \<and> c1 \<sqsubseteq> c1' \<and> c2 \<sqsubseteq> c2' \<and> S \<sqsubseteq> S')"
+by (cases c) auto
+
+lemma [simp]: "{Inv} WHILE b DO c {P} \<sqsubseteq> w \<longleftrightarrow>
+  (\<exists>Inv' c' P'. w = {Inv'} WHILE b DO c' {P'} \<and> c \<sqsubseteq> c' \<and> Inv \<sqsubseteq> Inv' \<and> P \<sqsubseteq> P')"
+by (cases w) auto
+
+instance
+proof
+  case goal1 thus ?case by (induct x) auto
+next
+  case goal2 thus ?case
+  apply(induct x y arbitrary: z rule: le_acom.induct)
+  apply (auto intro: le_trans)
+  done
+qed
+
+end
+
+
+subsubsection "Lifting"
+
+instantiation option :: (preord)preord
+begin
+
+fun le_option where
+"Some x \<sqsubseteq> Some y = (x \<sqsubseteq> y)" |
+"None \<sqsubseteq> y = True" |
+"Some _ \<sqsubseteq> None = False"
+
+lemma [simp]: "(x \<sqsubseteq> None) = (x = None)"
+by (cases x) simp_all
+
+lemma [simp]: "(Some x \<sqsubseteq> u) = (\<exists>y. u = Some y & x \<sqsubseteq> y)"
+by (cases u) auto
+
+instance proof
+  case goal1 show ?case by(cases x, simp_all)
+next
+  case goal2 thus ?case
+    by(cases z, simp, cases y, simp, cases x, auto intro: le_trans)
+qed
+
+end
+
+instantiation option :: (SL_top)SL_top
+begin
+
+fun join_option where
+"Some x \<squnion> Some y = Some(x \<squnion> y)" |
+"None \<squnion> y = y" |
+"x \<squnion> None = x"
+
+lemma join_None2[simp]: "x \<squnion> None = x"
+by (cases x) simp_all
+
+definition "\<top> = Some \<top>"
+
+instance proof
+  case goal1 thus ?case by(cases x, simp, cases y, simp_all)
+next
+  case goal2 thus ?case by(cases y, simp, cases x, simp_all)
+next
+  case goal3 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all)
+next
+  case goal4 thus ?case by(cases x, simp_all add: Top_option_def)
+qed
+
+end
+
+definition bot_acom :: "com \<Rightarrow> ('a::SL_top)option acom" ("\<bottom>\<^sub>c") where
+"\<bottom>\<^sub>c = anno None"
+
+lemma strip_bot_acom[simp]: "strip(\<bottom>\<^sub>c c) = c"
+by(simp add: bot_acom_def)
+
+lemma bot_acom[rule_format]: "strip c' = c \<longrightarrow> \<bottom>\<^sub>c c \<sqsubseteq> c'"
+apply(induct c arbitrary: c')
+apply (simp_all add: bot_acom_def)
+ apply(induct_tac c')
+  apply simp_all
+ apply(induct_tac c')
+  apply simp_all
+ apply(induct_tac c')
+  apply simp_all
+ apply(induct_tac c')
+  apply simp_all
+ apply(induct_tac c')
+  apply simp_all
+done
+
+
+subsubsection "Post-fixed point iteration"
+
+definition
+  pfp :: "(('a::preord) \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a option" where
+"pfp f = while_option (\<lambda>x. \<not> f x \<sqsubseteq> x) f"
+
+lemma pfp_pfp: assumes "pfp f x0 = Some x" shows "f x \<sqsubseteq> x"
+using while_option_stop[OF assms[simplified pfp_def]] by simp
+
+lemma pfp_least:
+assumes mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
+and "f p \<sqsubseteq> p" and "x0 \<sqsubseteq> p" and "pfp f x0 = Some x" shows "x \<sqsubseteq> p"
+proof-
+  { fix x assume "x \<sqsubseteq> p"
+    hence  "f x \<sqsubseteq> f p" by(rule mono)
+    from this `f p \<sqsubseteq> p` have "f x \<sqsubseteq> p" by(rule le_trans)
+  }
+  thus "x \<sqsubseteq> p" using assms(2-) while_option_rule[where P = "%x. x \<sqsubseteq> p"]
+    unfolding pfp_def by blast
+qed
+
+definition
+ lpfp\<^isub>c :: "(('a::SL_top)option acom \<Rightarrow> 'a option acom) \<Rightarrow> com \<Rightarrow> 'a option acom option" where
+"lpfp\<^isub>c f c = pfp f (\<bottom>\<^sub>c c)"
+
+lemma lpfpc_pfp: "lpfp\<^isub>c f c0 = Some c \<Longrightarrow> f c \<sqsubseteq> c"
+by(simp add: pfp_pfp lpfp\<^isub>c_def)
+
+lemma strip_pfp:
+assumes "\<And>x. g(f x) = g x" and "pfp f x0 = Some x" shows "g x = g x0"
+using assms while_option_rule[where P = "%x. g x = g x0" and c = f]
+unfolding pfp_def by metis
+
+lemma strip_lpfpc: assumes "\<And>c. strip(f c) = strip c" and "lpfp\<^isub>c f c = Some c'"
+shows "strip c' = c"
+using assms(1) strip_pfp[OF _ assms(2)[simplified lpfp\<^isub>c_def]]
+by(metis strip_bot_acom)
+
+lemma lpfpc_least:
+assumes mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
+and "strip p = c0" and "f p \<sqsubseteq> p" and lp: "lpfp\<^isub>c f c0 = Some c" shows "c \<sqsubseteq> p"
+using pfp_least[OF _ _ bot_acom[OF `strip p = c0`] lp[simplified lpfp\<^isub>c_def]]
+  mono `f p \<sqsubseteq> p`
+by blast
+
+
+subsection "Abstract Interpretation"
+
+definition \<gamma>_fun :: "('a \<Rightarrow> 'b set) \<Rightarrow> ('c \<Rightarrow> 'a) \<Rightarrow> ('c \<Rightarrow> 'b)set" where
+"\<gamma>_fun \<gamma> F = {f. \<forall>x. f x \<in> \<gamma>(F x)}"
+
+fun \<gamma>_option :: "('a \<Rightarrow> 'b set) \<Rightarrow> 'a option \<Rightarrow> 'b set" where
+"\<gamma>_option \<gamma> None = {}" |
+"\<gamma>_option \<gamma> (Some a) = \<gamma> a"
+
+text{* The interface for abstract values: *}
+
+locale Val_abs =
+fixes \<gamma> :: "'av::SL_top \<Rightarrow> val set"
+  assumes mono_gamma: "a \<sqsubseteq> b \<Longrightarrow> \<gamma> a \<subseteq> \<gamma> b"
+  and gamma_Top[simp]: "\<gamma> \<top> = UNIV"
+fixes num' :: "val \<Rightarrow> 'av"
+and plus' :: "'av \<Rightarrow> 'av \<Rightarrow> 'av"
+  assumes gamma_num': "n : \<gamma>(num' n)"
+  and gamma_plus':
+ "n1 : \<gamma> a1 \<Longrightarrow> n2 : \<gamma> a2 \<Longrightarrow> n1+n2 : \<gamma>(plus' a1 a2)"
+
+type_synonym 'av st = "(vname \<Rightarrow> 'av)"
+
+locale Abs_Int_Fun = Val_abs \<gamma> for \<gamma> :: "'av::SL_top \<Rightarrow> val set"
+begin
+
+fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where
+"aval' (N n) S = num' n" |
+"aval' (V x) S = S x" |
+"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
+
+fun step' :: "'av st option \<Rightarrow> 'av st option acom \<Rightarrow> 'av st option acom"
+ where
+"step' S (SKIP {P}) = (SKIP {S})" |
+"step' S (x ::= e {P}) =
+  x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(S(x := aval' e S))}" |
+"step' S (c1; c2) = step' S c1; step' (post c1) c2" |
+"step' S (IF b THEN c1 ELSE c2 {P}) =
+   IF b THEN step' S c1 ELSE step' S c2 {post c1 \<squnion> post c2}" |
+"step' S ({Inv} WHILE b DO c {P}) =
+  {S \<squnion> post c} WHILE b DO (step' Inv c) {Inv}"
+
+definition AI :: "com \<Rightarrow> 'av st option acom option" where
+"AI = lpfp\<^isub>c (step' \<top>)"
+
+
+lemma strip_step'[simp]: "strip(step' S c) = strip c"
+by(induct c arbitrary: S) (simp_all add: Let_def)
+
+
+abbreviation \<gamma>\<^isub>f :: "'av st \<Rightarrow> state set"
+where "\<gamma>\<^isub>f == \<gamma>_fun \<gamma>"
+
+abbreviation \<gamma>\<^isub>o :: "'av st option \<Rightarrow> state set"
+where "\<gamma>\<^isub>o == \<gamma>_option \<gamma>\<^isub>f"
+
+abbreviation \<gamma>\<^isub>c :: "'av st option acom \<Rightarrow> state set acom"
+where "\<gamma>\<^isub>c == map_acom \<gamma>\<^isub>o"
+
+lemma gamma_f_Top[simp]: "\<gamma>\<^isub>f Top = UNIV"
+by(simp add: Top_fun_def \<gamma>_fun_def)
+
+lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o Top = UNIV"
+by (simp add: Top_option_def)
+
+(* FIXME (maybe also le \<rightarrow> sqle?) *)
+
+lemma mono_gamma_f: "f \<sqsubseteq> g \<Longrightarrow> \<gamma>\<^isub>f f \<subseteq> \<gamma>\<^isub>f g"
+by(auto simp: le_fun_def \<gamma>_fun_def dest: mono_gamma)
+
+lemma mono_gamma_o:
+  "sa \<sqsubseteq> sa' \<Longrightarrow> \<gamma>\<^isub>o sa \<subseteq> \<gamma>\<^isub>o sa'"
+by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f)
+
+lemma mono_gamma_c: "ca \<sqsubseteq> ca' \<Longrightarrow> \<gamma>\<^isub>c ca \<le> \<gamma>\<^isub>c ca'"
+by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o)
+
+text{* Soundness: *}
+
+lemma aval'_sound: "s : \<gamma>\<^isub>f S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
+by (induct a) (auto simp: gamma_num' gamma_plus' \<gamma>_fun_def)
+
+lemma in_gamma_update:
+  "\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(S(x := a))"
+by(simp add: \<gamma>_fun_def)
+
+lemma step_preserves_le:
+  "\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; c \<le> \<gamma>\<^isub>c c' \<rbrakk> \<Longrightarrow> step S c \<le> \<gamma>\<^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 intro: aval'_sound in_gamma_update
+      split: option.splits del:subsetD)
+next
+  case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi)
+    by (metis le_post post_map_acom)
+next
+  case (If b c1 c2 P)
+  then obtain c1' c2' P' where
+      "c' = IF b THEN c1' ELSE c2' {P'}"
+      "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'" "c2 \<le> \<gamma>\<^isub>c c2'"
+    by (fastforce simp: If_le map_acom_If)
+  moreover have "post c1 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
+    by (metis (no_types) `c1 \<le> \<gamma>\<^isub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
+  moreover have "post c2 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
+    by (metis (no_types) `c2 \<le> \<gamma>\<^isub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
+  ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'` by (simp add: If.IH subset_iff)
+next
+  case (While I b c1 P)
+  then obtain c1' I' P' where
+    "c' = {I'} WHILE b DO c1' {P'}"
+    "I \<subseteq> \<gamma>\<^isub>o I'" "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'"
+    by (fastforce simp: map_acom_While While_le)
+  moreover have "S \<union> post c1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post c1')"
+    using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `c1 \<le> \<gamma>\<^isub>c c1'`, simplified]
+    by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans)
+  ultimately show ?case by (simp add: While.IH subset_iff)
+qed
+
+lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
+proof(simp add: CS_def AI_def)
+  assume 1: "lpfp\<^isub>c (step' \<top>) c = Some c'"
+  have 2: "step' \<top> c' \<sqsubseteq> c'" by(rule lpfpc_pfp[OF 1])
+  have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c"
+    by(simp add: strip_lpfpc[OF _ 1])
+  have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
+  proof(rule lfp_lowerbound[simplified,OF 3])
+    show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
+    proof(rule step_preserves_le[OF _ _])
+      show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
+      show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
+    qed
+  qed
+  with 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
+    by (blast intro: mono_gamma_c order_trans)
+qed
+
+end
+
+
+subsubsection "Monotonicity"
+
+lemma mono_post: "c \<sqsubseteq> c' \<Longrightarrow> post c \<sqsubseteq> post c'"
+by(induction c c' rule: le_acom.induct) (auto)
+
+locale Abs_Int_Fun_mono = Abs_Int_Fun +
+assumes mono_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> plus' a1 a2 \<sqsubseteq> plus' b1 b2"
+begin
+
+lemma mono_aval': "S \<sqsubseteq> S' \<Longrightarrow> aval' e S \<sqsubseteq> aval' e S'"
+by(induction e)(auto simp: le_fun_def mono_plus')
+
+lemma mono_update: "a \<sqsubseteq> a' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> S(x := a) \<sqsubseteq> S'(x := a')"
+by(simp add: le_fun_def)
+
+lemma mono_step': "S \<sqsubseteq> S' \<Longrightarrow> c \<sqsubseteq> c' \<Longrightarrow> step' S c \<sqsubseteq> step' S' c'"
+apply(induction c c' arbitrary: S S' rule: le_acom.induct)
+apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj
+            split: option.split)
+done
+
+end
+
+text{* Problem: not executable because of the comparison of abstract states,
+i.e. functions, in the post-fixedpoint computation. *}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy	Thu Apr 19 17:32:30 2012 +0200
@@ -0,0 +1,411 @@
+(* Author: Tobias Nipkow *)
+
+theory Abs_Int1_ITP
+imports Abs_State_ITP
+begin
+
+subsection "Computable Abstract Interpretation"
+
+text{* Abstract interpretation over type @{text st} instead of
+functions. *}
+
+context Gamma
+begin
+
+fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where
+"aval' (N n) S = num' n" |
+"aval' (V x) S = lookup S x" |
+"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
+
+lemma aval'_sound: "s : \<gamma>\<^isub>f S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
+by (induction a) (auto simp: gamma_num' gamma_plus' \<gamma>_st_def lookup_def)
+
+end
+
+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}. *}
+
+locale Abs_Int = Gamma where \<gamma>=\<gamma> for \<gamma> :: "'av::SL_top \<Rightarrow> val set"
+begin
+
+fun step' :: "'av st option \<Rightarrow> 'av st option acom \<Rightarrow> 'av st option acom" where
+"step' S (SKIP {P}) = (SKIP {S})" |
+"step' S (x ::= e {P}) =
+  x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(update S x (aval' e S))}" |
+"step' S (c1; c2) = step' S c1; step' (post c1) c2" |
+"step' S (IF b THEN c1 ELSE c2 {P}) =
+  (let c1' = step' S c1; c2' = step' S c2
+   in IF b THEN c1' ELSE c2' {post c1 \<squnion> post c2})" |
+"step' S ({Inv} WHILE b DO c {P}) =
+   {S \<squnion> post c} WHILE b DO step' Inv c {Inv}"
+
+definition AI :: "com \<Rightarrow> 'av st option acom option" where
+"AI = lpfp\<^isub>c (step' \<top>)"
+
+
+lemma strip_step'[simp]: "strip(step' S c) = strip c"
+by(induct c arbitrary: S) (simp_all add: Let_def)
+
+
+text{* Soundness: *}
+
+lemma in_gamma_update:
+  "\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(update S x a)"
+by(simp add: \<gamma>_st_def lookup_update)
+
+text{* The soundness proofs are textually identical to the ones for the step
+function operating on states as functions. *}
+
+lemma step_preserves_le:
+  "\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; c \<le> \<gamma>\<^isub>c c' \<rbrakk> \<Longrightarrow> step S c \<le> \<gamma>\<^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 intro: aval'_sound in_gamma_update
+      split: option.splits del:subsetD)
+next
+  case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi)
+    by (metis le_post post_map_acom)
+next
+  case (If b c1 c2 P)
+  then obtain c1' c2' P' where
+      "c' = IF b THEN c1' ELSE c2' {P'}"
+      "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'" "c2 \<le> \<gamma>\<^isub>c c2'"
+    by (fastforce simp: If_le map_acom_If)
+  moreover have "post c1 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
+    by (metis (no_types) `c1 \<le> \<gamma>\<^isub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
+  moreover have "post c2 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
+    by (metis (no_types) `c2 \<le> \<gamma>\<^isub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
+  ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'` by (simp add: If.IH subset_iff)
+next
+  case (While I b c1 P)
+  then obtain c1' I' P' where
+    "c' = {I'} WHILE b DO c1' {P'}"
+    "I \<subseteq> \<gamma>\<^isub>o I'" "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'"
+    by (fastforce simp: map_acom_While While_le)
+  moreover have "S \<union> post c1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post c1')"
+    using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `c1 \<le> \<gamma>\<^isub>c c1'`, simplified]
+    by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans)
+  ultimately show ?case by (simp add: While.IH subset_iff)
+qed
+
+lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
+proof(simp add: CS_def AI_def)
+  assume 1: "lpfp\<^isub>c (step' \<top>) c = Some c'"
+  have 2: "step' \<top> c' \<sqsubseteq> c'" by(rule lpfpc_pfp[OF 1])
+  have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c"
+    by(simp add: strip_lpfpc[OF _ 1])
+  have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
+  proof(rule lfp_lowerbound[simplified,OF 3])
+    show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
+    proof(rule step_preserves_le[OF _ _])
+      show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
+      show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
+    qed
+  qed
+  from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
+    by (blast intro: mono_gamma_c order_trans)
+qed
+
+end
+
+
+subsubsection "Monotonicity"
+
+locale Abs_Int_mono = Abs_Int +
+assumes mono_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> plus' a1 a2 \<sqsubseteq> plus' b1 b2"
+begin
+
+lemma mono_aval': "S \<sqsubseteq> S' \<Longrightarrow> aval' e S \<sqsubseteq> aval' e S'"
+by(induction e) (auto simp: le_st_def lookup_def mono_plus')
+
+lemma mono_update: "a \<sqsubseteq> a' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> update S x a \<sqsubseteq> update S' x a'"
+by(auto simp add: le_st_def lookup_def update_def)
+
+lemma mono_step': "S \<sqsubseteq> S' \<Longrightarrow> c \<sqsubseteq> c' \<Longrightarrow> step' S c \<sqsubseteq> step' S' c'"
+apply(induction c c' arbitrary: S S' rule: le_acom.induct)
+apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj
+            split: option.split)
+done
+
+end
+
+
+subsubsection "Ascending Chain Condition"
+
+hide_const (open) acc
+
+abbreviation "strict r == r \<inter> -(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 \<Longrightarrow> 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 \<sqsubseteq> y}"
+shows "acc {(x,y::'a::preord option). x \<sqsubseteq> y}"
+proof(auto simp: wf_eq_minimal)
+  fix xo :: "'a option" and Qo assume "xo : Qo"
+  let ?Q = "{x. Some x \<in> Qo}"
+  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")
+  proof cases
+    assume "?Q = {}"
+    hence "?P None" by auto
+    moreover have "None \<in> Qo" using `?Q = {}` `xo : Qo`
+      by auto (metis not_Some_eq)
+    ultimately show ?thesis by blast
+  next
+    assume "?Q \<noteq> {}"
+    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 \<Rightarrow> 'b::{comm_monoid_add, ordered_cancel_ab_semigroup_add}"
+assumes "finite A" and "ALL x:A. f x \<le> 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}) \<union> {a})"
+    by(simp add:insert_absorb[OF `a:A`])
+  also have "\<dots> = setsum f (A-{a}) + setsum f {a}"
+    using `finite A` by(subst setsum_Un_disjoint) auto
+  also have "setsum f (A-{a}) \<le> 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}) \<union> {a})"
+    using `finite A` by(subst setsum_Un_disjoint[symmetric]) auto
+  also have "\<dots> = 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 \<sqsubseteq> y})^-1 <= measure m"
+and "\<forall>x y::'a::SL_top. x \<sqsubseteq> y \<and> y \<sqsubseteq> x \<longrightarrow> m x = m y"
+shows "(strict{(S,S'::'a::SL_top st). S \<sqsubseteq> S'})^-1 \<subseteq>
+  measure(%fd. \<Sum>x| x\<in>set(dom fd) \<and> ~ \<top> \<sqsubseteq> fun fd x. m(fun fd x)+1)"
+proof-
+  { fix S S' :: "'a st" assume "S \<sqsubseteq> S'" "~ S' \<sqsubseteq> S"
+    let ?X = "set(dom S)" let ?Y = "set(dom S')"
+    let ?f = "fun S" let ?g = "fun S'"
+    let ?X' = "{x:?X. ~ \<top> \<sqsubseteq> ?f x}" let ?Y' = "{y:?Y. ~ \<top> \<sqsubseteq> ?g y}"
+    from `S \<sqsubseteq> S'` have "ALL y:?Y'\<inter>?X. ?f y \<sqsubseteq> ?g y"
+      by(auto simp: le_st_def lookup_def)
+    hence 1: "ALL y:?Y'\<inter>?X. m(?g y)+1 \<le> m(?f y)+1"
+      using assms(1,2) by(fastforce)
+    from `~ S' \<sqsubseteq> S` obtain u where u: "u : ?X" "~ lookup S' u \<sqsubseteq> ?f u"
+      by(auto simp: le_st_def)
+    hence "u : ?X'" by simp (metis preord_class.le_trans top)
+    have "?Y'-?X = {}" using `S \<sqsubseteq> S'` by(fastforce simp: le_st_def lookup_def)
+    have "?Y'\<inter>?X <= ?X'" apply auto
+      apply (metis `S \<sqsubseteq> S'` le_st_def lookup_def preord_class.le_trans)
+      done
+    have "(\<Sum>y\<in>?Y'. m(?g y)+1) = (\<Sum>y\<in>(?Y'-?X) \<union> (?Y'\<inter>?X). m(?g y)+1)"
+      by (metis Un_Diff_Int)
+    also have "\<dots> = (\<Sum>y\<in>?Y'\<inter>?X. m(?g y)+1)"
+      using `?Y'-?X = {}` by (metis Un_empty_left)
+    also have "\<dots> < (\<Sum>x\<in>?X'. m(?f x)+1)"
+    proof cases
+      assume "u \<in> ?Y'"
+      hence "m(?g u) < m(?f u)" using assms(1) `S \<sqsubseteq> S'` u
+        by (fastforce simp: le_st_def lookup_def)
+      have "(\<Sum>y\<in>?Y'\<inter>?X. m(?g y)+1) < (\<Sum>y\<in>?Y'\<inter>?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 "\<dots> \<le> (\<Sum>y\<in>?X'. m(?f y)+1)"
+        by(simp add: setsum_mono3[OF _ `?Y'\<inter>?X <= ?X'`])
+      finally show ?thesis .
+    next
+      assume "u \<notin> ?Y'"
+      with `?Y'\<inter>?X <= ?X'` have "?Y'\<inter>?X - {u} <= ?X' - {u}" by blast
+      have "(\<Sum>y\<in>?Y'\<inter>?X. m(?g y)+1) = (\<Sum>y\<in>?Y'\<inter>?X - {u}. m(?g y)+1)"
+      proof-
+        have "?Y'\<inter>?X = ?Y'\<inter>?X - {u}" using `u \<notin> ?Y'` by auto
+        thus ?thesis by metis
+      qed
+      also have "\<dots> < (\<Sum>y\<in>?Y'\<inter>?X-{u}. m(?g y)+1) + (\<Sum>y\<in>{u}. m(?f y)+1)" by simp
+      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)"
+        using 1 by(blast intro: setsum_mono)
+      also have "\<dots> \<le> (\<Sum>y\<in>?X'-{u}. m(?f y)+1)"
+        by(simp add: setsum_mono3[OF _ `?Y'\<inter>?X-{u} <= ?X'-{u}`])
+      also have "\<dots> + (\<Sum>y\<in>{u}. m(?f y)+1)= (\<Sum>y\<in>(?X'-{u}) \<union> {u}. m(?f y)+1)"
+        using `u:?X'` by(subst setsum_Un_disjoint[symmetric]) auto
+      also have "\<dots> = (\<Sum>x\<in>?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 "(\<Sum>y\<in>?Y'. m(?g y)+1) < (\<Sum>x\<in>?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 \<longleftrightarrow> (x,y) \<in> r \<and> (xs,ys) \<in> listrel r"
+by (blast intro:listrel.Cons)
+
+lemma listrel_app: "(xs1,ys1) : listrel r \<Longrightarrow> (xs2,ys2) : listrel r
+  \<Longrightarrow> (xs1@xs2, ys1@ys2) : listrel r"
+by(auto simp add: listrel_iff_zip)
+
+lemma listrel_app_same_size: "size xs1 = size ys1 \<Longrightarrow> size xs2 = size ys2 \<Longrightarrow>
+  (xs1@xs2, ys1@ys2) : listrel r \<longleftrightarrow>
+  (xs1,ys1) : listrel r \<and> (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) \<longleftrightarrow> (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 \<Longrightarrow> (y,z) : r \<Longrightarrow> (x,z) : r"
+    using `trans r` unfolding trans_def by blast
+  from assms(3) obtain mx :: "'a set \<Rightarrow> 'a" where
+    mx: "!!S x. x:S \<Longrightarrow> mx S : S \<and> (\<forall>y. (mx S,y) : strict r \<longrightarrow> y \<notin> S)"
+    by(simp add: wf_eq_minimal) metis
+  let ?R = "listrel r - {([], [])}"
+  { fix Q and xs :: "'a list"
+    have "xs \<in> Q \<Longrightarrow> \<exists>ys. ys\<in>Q \<and> (\<forall>zs. (ys, zs) \<in> strict ?R \<longrightarrow> zs \<notin> Q)"
+      (is "_ \<Longrightarrow> \<exists>ys. ?P Q ys")
+    proof(induction xs arbitrary: Q rule: length_induct)
+      case (1 xs)
+      { have "!!ys Q. size ys < size xs \<Longrightarrow> ys : Q \<Longrightarrow> 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. \<exists>bs. size bs = size ys \<and> a#bs : Q}"
+        have "x : ?Q1" using `xs : Q` Cons by auto
+        from mx[OF this] obtain m1 where
+          1: "m1 \<in> ?Q1 \<and> (\<forall>y. (m1,y) \<in> strict r \<longrightarrow> y \<notin> ?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. \<exists>m1'. (m1',m1):r \<and> (m1,m1'):r \<and> m1'#bs : Q \<and> 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 \<and> (m1,m1') : r \<and> m1'#ms : Q"
+          by blast
+        hence "\<forall>ab. (m1'#ms,ab) : strict ?R \<longrightarrow> ab \<notin> 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 \<Rightarrow> '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 \<Longrightarrow> 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
+
+lemmas size_annos_same2 = eqTrueI[OF size_annos_same]
+
+lemma set_annos_anno: "set (annos (anno a c)) = {a}"
+by(induction c)(auto)
+
+lemma le_iff_le_annos: "c1 \<sqsubseteq> c2 \<longleftrightarrow>
+ (annos c1, annos c2) : listrel{(x,y). x \<sqsubseteq> y} \<and> strip c1 = strip c2"
+apply(induct c1 c2 rule: le_acom.induct)
+apply (auto simp: listrel.Nil listrel_Cons_iff listrel_app size_annos_same2)
+apply (metis listrel_app_same_size size_annos_same)+
+done
+
+lemma le_acom_subset_same_annos:
+ "(strict{(c,c'::'a::preord acom). c \<sqsubseteq> c'})^-1 \<subseteq>
+  (strict(inv_image (listrel{(a,a'::'a). a \<sqsubseteq> a'} - {([],[])}) annos))^-1"
+by(auto simp: le_iff_le_annos)
+
+lemma acc_acom: "acc {(a,a'::'a::preord). a \<sqsubseteq> a'} \<Longrightarrow>
+  acc {(c,c'::'a acom). c \<sqsubseteq> 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: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" and "acc {(x::'a,y). x \<sqsubseteq> y}"
+and "x0 \<sqsubseteq> f x0" shows "EX x. pfp f x0 = Some x"
+proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. x \<sqsubseteq> f x"])
+  show "wf {(x, s). (s \<sqsubseteq> f s \<and> \<not> f s \<sqsubseteq> s) \<and> x = f s}"
+    by(rule wf_subset[OF assms(2)]) auto
+next
+  show "x0 \<sqsubseteq> f x0" by(rule assms)
+next
+  fix x assume "x \<sqsubseteq> f x" thus "f x \<sqsubseteq> f(f x)" by(rule mono)
+qed
+
+lemma lpfpc_termination:
+  fixes f :: "(('a::SL_top)option acom \<Rightarrow> 'a option acom)"
+  assumes "acc {(x::'a,y). x \<sqsubseteq> y}" and "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
+  and "\<And>c. strip(f c) = strip c"
+  shows "\<exists>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
+
+context Abs_Int_mono
+begin
+
+lemma AI_Some_measure:
+assumes "(strict{(x,y::'a). x \<sqsubseteq> y})^-1 <= measure m"
+and "\<forall>x y::'a. x \<sqsubseteq> y \<and> y \<sqsubseteq> x \<longrightarrow> m x = m y"
+shows "\<exists>c'. AI c = Some c'"
+unfolding AI_def
+apply(rule lpfpc_termination)
+apply(rule wf_subset[OF wf_measure measure_st[OF assms]])
+apply(erule mono_step'[OF le_refl])
+apply(rule strip_step')
+done
+
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_const_ITP.thy	Thu Apr 19 17:32:30 2012 +0200
@@ -0,0 +1,139 @@
+(* Author: Tobias Nipkow *)
+
+theory Abs_Int1_const_ITP
+imports Abs_Int1_ITP Abs_Int_Tests
+begin
+
+subsection "Constant Propagation"
+
+datatype const = Const val | Any
+
+fun \<gamma>_const where
+"\<gamma>_const (Const n) = {n}" |
+"\<gamma>_const (Any) = UNIV"
+
+fun plus_const where
+"plus_const (Const m) (Const n) = Const(m+n)" |
+"plus_const _ _ = Any"
+
+lemma plus_const_cases: "plus_const a1 a2 =
+  (case (a1,a2) of (Const m, Const n) \<Rightarrow> Const(m+n) | _ \<Rightarrow> Any)"
+by(auto split: prod.split const.split)
+
+instantiation const :: SL_top
+begin
+
+fun le_const where
+"_ \<sqsubseteq> Any = True" |
+"Const n \<sqsubseteq> Const m = (n=m)" |
+"Any \<sqsubseteq> Const _ = False"
+
+fun join_const where
+"Const m \<squnion> Const n = (if n=m then Const m else Any)" |
+"_ \<squnion> _ = Any"
+
+definition "\<top> = Any"
+
+instance
+proof
+  case goal1 thus ?case by (cases x) simp_all
+next
+  case goal2 thus ?case by(cases z, cases y, cases x, simp_all)
+next
+  case goal3 thus ?case by(cases x, cases y, simp_all)
+next
+  case goal4 thus ?case by(cases y, cases x, simp_all)
+next
+  case goal5 thus ?case by(cases z, cases y, cases x, simp_all)
+next
+  case goal6 thus ?case by(simp add: Top_const_def)
+qed
+
+end
+
+
+interpretation Val_abs
+where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const
+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_const_def)
+next
+  case goal3 show ?case by simp
+next
+  case goal4 thus ?case
+    by(auto simp: plus_const_cases split: const.split)
+qed
+
+interpretation Abs_Int
+where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const
+defines AI_const is AI and step_const is step' and aval'_const is aval'
+..
+
+
+subsubsection "Tests"
+
+value "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test1_const))"
+value "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test1_const))"
+value "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test1_const))"
+value "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test1_const))"
+value "show_acom_opt (AI_const test1_const)"
+
+value "show_acom_opt (AI_const test2_const)"
+value "show_acom_opt (AI_const test3_const)"
+
+value "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test4_const))"
+value "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test4_const))"
+value "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test4_const))"
+value "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test4_const))"
+value "show_acom_opt (AI_const test4_const)"
+
+value "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test5_const))"
+value "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test5_const))"
+value "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test5_const))"
+value "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test5_const))"
+value "show_acom (((step_const \<top>)^^4) (\<bottom>\<^sub>c test5_const))"
+value "show_acom (((step_const \<top>)^^5) (\<bottom>\<^sub>c test5_const))"
+value "show_acom_opt (AI_const test5_const)"
+
+value "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test6_const))"
+value "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test6_const))"
+value "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test6_const))"
+value "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test6_const))"
+value "show_acom (((step_const \<top>)^^4) (\<bottom>\<^sub>c test6_const))"
+value "show_acom (((step_const \<top>)^^5) (\<bottom>\<^sub>c test6_const))"
+value "show_acom (((step_const \<top>)^^6) (\<bottom>\<^sub>c test6_const))"
+value "show_acom (((step_const \<top>)^^7) (\<bottom>\<^sub>c test6_const))"
+value "show_acom (((step_const \<top>)^^8) (\<bottom>\<^sub>c test6_const))"
+value "show_acom (((step_const \<top>)^^9) (\<bottom>\<^sub>c test6_const))"
+value "show_acom (((step_const \<top>)^^10) (\<bottom>\<^sub>c test6_const))"
+value "show_acom (((step_const \<top>)^^11) (\<bottom>\<^sub>c test6_const))"
+value "show_acom_opt (AI_const test6_const)"
+
+
+text{* Monotonicity: *}
+
+interpretation Abs_Int_mono
+where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const
+proof
+  case goal1 thus ?case
+    by(auto simp: plus_const_cases split: const.split)
+qed
+
+text{* Termination: *}
+
+definition "m_const x = (case x of Const _ \<Rightarrow> 1 | Any \<Rightarrow> 0)"
+
+lemma measure_const:
+  "(strict{(x::const,y). x \<sqsubseteq> y})^-1 \<subseteq> measure m_const"
+by(auto simp: m_const_def split: const.splits)
+
+lemma measure_const_eq:
+  "\<forall> x y::const. x \<sqsubseteq> y \<and> y \<sqsubseteq> x \<longrightarrow> m_const x = m_const y"
+by(auto simp: m_const_def split: const.splits)
+
+lemma "EX c'. AI_const c = Some c'"
+by(rule AI_Some_measure[OF measure_const measure_const_eq])
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_parity_ITP.thy	Thu Apr 19 17:32:30 2012 +0200
@@ -0,0 +1,168 @@
+(* Author: Tobias Nipkow *)
+
+theory Abs_Int1_parity_ITP
+imports Abs_Int1_ITP
+begin
+
+subsection "Parity Analysis"
+
+datatype parity = Even | Odd | Either
+
+text{* Instantiation of class @{class preord} with type @{typ parity}: *}
+
+instantiation parity :: preord
+begin
+
+text{* First the definition of the interface function @{text"\<sqsubseteq>"}. Note that
+the header of the definition must refer to the ascii name @{const le} of the
+constants as @{text le_parity} and the definition is named @{text
+le_parity_def}.  Inside the definition the symbolic names can be used. *}
+
+definition le_parity where
+"x \<sqsubseteq> y = (y = Either \<or> x=y)"
+
+text{* Now the instance proof, i.e.\ the proof that the definition fulfills
+the axioms (assumptions) of the class. The initial proof-step generates the
+necessary proof obligations. *}
+
+instance
+proof
+  fix x::parity show "x \<sqsubseteq> x" by(auto simp: le_parity_def)
+next
+  fix x y z :: parity assume "x \<sqsubseteq> y" "y \<sqsubseteq> z" thus "x \<sqsubseteq> z"
+    by(auto simp: le_parity_def)
+qed
+
+end
+
+text{* Instantiation of class @{class SL_top} with type @{typ parity}: *}
+
+instantiation parity :: SL_top
+begin
+
+
+definition join_parity where
+"x \<squnion> y = (if x \<sqsubseteq> y then y else if y \<sqsubseteq> x then x else Either)"
+
+definition Top_parity where
+"\<top> = Either"
+
+text{* Now the instance proof. This time we take a lazy shortcut: we do not
+write out the proof obligations but use the @{text goali} primitive to refer
+to the assumptions of subgoal i and @{text "case?"} to refer to the
+conclusion of subgoal i. The class axioms are presented in the same order as
+in the class definition. *}
+
+instance
+proof
+  case goal1 (*join1*) show ?case by(auto simp: le_parity_def join_parity_def)
+next
+  case goal2 (*join2*) show ?case by(auto simp: le_parity_def join_parity_def)
+next
+  case goal3 (*join least*) thus ?case by(auto simp: le_parity_def join_parity_def)
+next
+  case goal4 (*Top*) show ?case by(auto simp: le_parity_def Top_parity_def)
+qed
+
+end
+
+
+text{* Now we define the functions used for instantiating the abstract
+interpretation locales. Note that the Isabelle terminology is
+\emph{interpretation}, not \emph{instantiation} of locales, but we use
+instantiation to avoid confusion with abstract interpretation.  *}
+
+fun \<gamma>_parity :: "parity \<Rightarrow> val set" where
+"\<gamma>_parity Even = {i. i mod 2 = 0}" |
+"\<gamma>_parity Odd  = {i. i mod 2 = 1}" |
+"\<gamma>_parity Either = UNIV"
+
+fun num_parity :: "val \<Rightarrow> parity" where
+"num_parity i = (if i mod 2 = 0 then Even else Odd)"
+
+fun plus_parity :: "parity \<Rightarrow> parity \<Rightarrow> parity" where
+"plus_parity Even Even = Even" |
+"plus_parity Odd  Odd  = Even" |
+"plus_parity Even Odd  = Odd" |
+"plus_parity Odd  Even = Odd" |
+"plus_parity Either y  = Either" |
+"plus_parity x Either  = Either"
+
+text{* First we instantiate the abstract value interface and prove that the
+functions on type @{typ parity} have all the necessary properties: *}
+
+interpretation Val_abs
+where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity
+proof txt{* of the locale axioms *}
+  fix a b :: parity
+  assume "a \<sqsubseteq> b" thus "\<gamma>_parity a \<subseteq> \<gamma>_parity b"
+    by(auto simp: le_parity_def)
+next txt{* The rest in the lazy, implicit way *}
+  case goal2 show ?case by(auto simp: Top_parity_def)
+next
+  case goal3 show ?case by auto
+next
+  txt{* Warning: this subproof refers to the names @{text a1} and @{text a2}
+  from the statement of the axiom. *}
+  case goal4 thus ?case
+  proof(cases a1 a2 rule: parity.exhaust[case_product parity.exhaust])
+  qed (auto simp add:mod_add_eq)
+qed
+
+text{* Instantiating the abstract interpretation locale requires no more
+proofs (they happened in the instatiation above) but delivers the
+instantiated abstract interpreter which we call AI: *}
+
+interpretation Abs_Int
+where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity
+defines aval_parity is aval' and step_parity is step' and AI_parity is AI
+..
+
+
+subsubsection "Tests"
+
+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)"
+
+definition "test2_parity =
+  ''x'' ::= N 1;
+  WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 3)"
+
+value "show_acom ((step_parity \<top> ^^1) (anno None test2_parity))"
+value "show_acom ((step_parity \<top> ^^2) (anno None test2_parity))"
+value "show_acom ((step_parity \<top> ^^3) (anno None test2_parity))"
+value "show_acom ((step_parity \<top> ^^4) (anno None test2_parity))"
+value "show_acom ((step_parity \<top> ^^5) (anno None test2_parity))"
+value "show_acom_opt (AI_parity test2_parity)"
+
+
+subsubsection "Termination"
+
+interpretation Abs_Int_mono
+where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity
+proof
+  case goal1 thus ?case
+  proof(cases a1 a2 b1 b2
+   rule: parity.exhaust[case_product parity.exhaust[case_product parity.exhaust[case_product parity.exhaust]]]) (* FIXME - UGLY! *)
+  qed (auto simp add:le_parity_def)
+qed
+
+
+definition m_parity :: "parity \<Rightarrow> nat" where
+"m_parity x = (if x=Either then 0 else 1)"
+
+lemma measure_parity:
+  "(strict{(x::parity,y). x \<sqsubseteq> y})^-1 \<subseteq> measure m_parity"
+by(auto simp add: m_parity_def le_parity_def)
+
+lemma measure_parity_eq:
+  "\<forall>x y::parity. x \<sqsubseteq> y \<and> y \<sqsubseteq> x \<longrightarrow> m_parity x = m_parity y"
+by(auto simp add: m_parity_def le_parity_def)
+
+lemma AI_parity_Some: "\<exists>c'. AI_parity c = Some c'"
+by(rule AI_Some_measure[OF measure_parity measure_parity_eq])
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ITP.thy	Thu Apr 19 17:32:30 2012 +0200
@@ -0,0 +1,358 @@
+(* Author: Tobias Nipkow *)
+
+theory Abs_Int2_ITP
+imports Abs_Int1_ITP Vars
+begin
+
+instantiation prod :: (preord,preord) preord
+begin
+
+definition "le_prod p1 p2 = (fst p1 \<sqsubseteq> fst p2 \<and> snd p1 \<sqsubseteq> snd p2)"
+
+instance
+proof
+  case goal1 show ?case by(simp add: le_prod_def)
+next
+  case goal2 thus ?case unfolding le_prod_def by(metis le_trans)
+qed
+
+end
+
+instantiation com :: vars
+begin
+
+fun vars_com :: "com \<Rightarrow> vname set" where
+"vars com.SKIP = {}" |
+"vars (x::=e) = {x} \<union> vars e" |
+"vars (c1;c2) = vars c1 \<union> vars c2" |
+"vars (IF b THEN c1 ELSE c2) = vars b \<union> vars c1 \<union> vars c2" |
+"vars (WHILE b DO c) = vars b \<union> vars c"
+
+instance ..
+
+end
+
+lemma finite_avars: "finite(vars(a::aexp))"
+by(induction a) simp_all
+
+lemma finite_bvars: "finite(vars(b::bexp))"
+by(induction b) (simp_all add: finite_avars)
+
+lemma finite_cvars: "finite(vars(c::com))"
+by(induction c) (simp_all add: finite_avars finite_bvars)
+
+
+subsection "Backward Analysis of Expressions"
+
+hide_const bot
+
+class L_top_bot = SL_top +
+fixes meet :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 65)
+and bot :: "'a" ("\<bottom>")
+assumes meet_le1 [simp]: "x \<sqinter> y \<sqsubseteq> x"
+and meet_le2 [simp]: "x \<sqinter> y \<sqsubseteq> y"
+and meet_greatest: "x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<sqinter> z"
+assumes bot[simp]: "\<bottom> \<sqsubseteq> x"
+begin
+
+lemma mono_meet: "x \<sqsubseteq> x' \<Longrightarrow> y \<sqsubseteq> y' \<Longrightarrow> x \<sqinter> y \<sqsubseteq> x' \<sqinter> y'"
+by (metis meet_greatest meet_le1 meet_le2 le_trans)
+
+end
+
+locale Val_abs1_gamma =
+  Gamma where \<gamma> = \<gamma> for \<gamma> :: "'av::L_top_bot \<Rightarrow> val set" +
+assumes inter_gamma_subset_gamma_meet:
+  "\<gamma> a1 \<inter> \<gamma> a2 \<subseteq> \<gamma>(a1 \<sqinter> a2)"
+and gamma_Bot[simp]: "\<gamma> \<bottom> = {}"
+begin
+
+lemma in_gamma_meet: "x : \<gamma> a1 \<Longrightarrow> x : \<gamma> a2 \<Longrightarrow> x : \<gamma>(a1 \<sqinter> a2)"
+by (metis IntI inter_gamma_subset_gamma_meet set_mp)
+
+lemma gamma_meet[simp]: "\<gamma>(a1 \<sqinter> a2) = \<gamma> a1 \<inter> \<gamma> a2"
+by (metis equalityI inter_gamma_subset_gamma_meet le_inf_iff mono_gamma meet_le1 meet_le2)
+
+end
+
+
+locale Val_abs1 =
+  Val_abs1_gamma where \<gamma> = \<gamma>
+   for \<gamma> :: "'av::L_top_bot \<Rightarrow> val set" +
+fixes test_num' :: "val \<Rightarrow> 'av \<Rightarrow> bool"
+and filter_plus' :: "'av \<Rightarrow> 'av \<Rightarrow> 'av \<Rightarrow> 'av * 'av"
+and filter_less' :: "bool \<Rightarrow> 'av \<Rightarrow> 'av \<Rightarrow> 'av * 'av"
+assumes test_num': "test_num' n a = (n : \<gamma> a)"
+and filter_plus': "filter_plus' a a1 a2 = (b1,b2) \<Longrightarrow>
+  n1 : \<gamma> a1 \<Longrightarrow> n2 : \<gamma> a2 \<Longrightarrow> n1+n2 : \<gamma> a \<Longrightarrow> n1 : \<gamma> b1 \<and> n2 : \<gamma> b2"
+and filter_less': "filter_less' (n1<n2) a1 a2 = (b1,b2) \<Longrightarrow>
+  n1 : \<gamma> a1 \<Longrightarrow> n2 : \<gamma> a2 \<Longrightarrow> n1 : \<gamma> b1 \<and> n2 : \<gamma> b2"
+
+
+locale Abs_Int1 =
+  Val_abs1 where \<gamma> = \<gamma> for \<gamma> :: "'av::L_top_bot \<Rightarrow> val set"
+begin
+
+lemma in_gamma_join_UpI: "s : \<gamma>\<^isub>o S1 \<or> s : \<gamma>\<^isub>o S2 \<Longrightarrow> s : \<gamma>\<^isub>o(S1 \<squnion> S2)"
+by (metis (no_types) join_ge1 join_ge2 mono_gamma_o set_rev_mp)
+
+fun aval'' :: "aexp \<Rightarrow> 'av st option \<Rightarrow> 'av" where
+"aval'' e None = \<bottom>" |
+"aval'' e (Some sa) = aval' e sa"
+
+lemma aval''_sound: "s : \<gamma>\<^isub>o S \<Longrightarrow> aval a s : \<gamma>(aval'' a S)"
+by(cases S)(simp add: aval'_sound)+
+
+subsubsection "Backward analysis"
+
+fun afilter :: "aexp \<Rightarrow> 'av \<Rightarrow> 'av st option \<Rightarrow> 'av st option" where
+"afilter (N n) a S = (if test_num' n a then S else None)" |
+"afilter (V x) a S = (case S of None \<Rightarrow> None | Some S \<Rightarrow>
+  let a' = lookup S x \<sqinter> a in
+  if a' \<sqsubseteq> \<bottom> then None else Some(update S x a'))" |
+"afilter (Plus e1 e2) a S =
+ (let (a1,a2) = filter_plus' a (aval'' e1 S) (aval'' e2 S)
+  in afilter e1 a1 (afilter e2 a2 S))"
+
+text{* The test for @{const bot} in the @{const V}-case is important: @{const
+bot} indicates that a variable has no possible values, i.e.\ that the current
+program point is unreachable. But then the abstract state should collapse to
+@{const None}. Put differently, we maintain the invariant that in an abstract
+state of the form @{term"Some s"}, all variables are mapped to non-@{const
+bot} values. Otherwise the (pointwise) join of two abstract states, one of
+which contains @{const bot} values, may produce too large a result, thus
+making the analysis less precise. *}
+
+
+fun bfilter :: "bexp \<Rightarrow> bool \<Rightarrow> 'av st option \<Rightarrow> 'av st option" where
+"bfilter (Bc v) res S = (if v=res then S else None)" |
+"bfilter (Not b) res S = bfilter b (\<not> res) S" |
+"bfilter (And b1 b2) res S =
+  (if res then bfilter b1 True (bfilter b2 True S)
+   else bfilter b1 False S \<squnion> bfilter b2 False S)" |
+"bfilter (Less e1 e2) res S =
+  (let (res1,res2) = filter_less' res (aval'' e1 S) (aval'' e2 S)
+   in afilter e1 res1 (afilter e2 res2 S))"
+
+lemma afilter_sound: "s : \<gamma>\<^isub>o S \<Longrightarrow> aval e s : \<gamma> a \<Longrightarrow> s : \<gamma>\<^isub>o (afilter e a S)"
+proof(induction e arbitrary: a S)
+  case N thus ?case by simp (metis test_num')
+next
+  case (V x)
+  obtain S' where "S = Some S'" and "s : \<gamma>\<^isub>f S'" using `s : \<gamma>\<^isub>o S`
+    by(auto simp: in_gamma_option_iff)
+  moreover hence "s x : \<gamma> (lookup S' x)" by(simp add: \<gamma>_st_def)
+  moreover have "s x : \<gamma> a" using V by simp
+  ultimately show ?case using V(1)
+    by(simp add: lookup_update Let_def \<gamma>_st_def)
+      (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(3)] aval''_sound[OF Plus(3)]]
+    by (auto split: prod.split)
+qed
+
+lemma bfilter_sound: "s : \<gamma>\<^isub>o S \<Longrightarrow> bv = bval b s \<Longrightarrow> s : \<gamma>\<^isub>o(bfilter b bv S)"
+proof(induction b arbitrary: S bv)
+  case Bc thus ?case by simp
+next
+  case (Not b) thus ?case by simp
+next
+  case (And b1 b2) thus ?case by(fastforce simp: in_gamma_join_UpI)
+next
+  case (Less e1 e2) thus ?case
+    by (auto split: prod.split)
+       (metis afilter_sound filter_less' aval''_sound Less)
+qed
+
+
+fun step' :: "'av st option \<Rightarrow> 'av st option acom \<Rightarrow> 'av st option acom"
+ where
+"step' S (SKIP {P}) = (SKIP {S})" |
+"step' S (x ::= e {P}) =
+  x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(update S x (aval' e S))}" |
+"step' S (c1; c2) = step' S c1; step' (post c1) c2" |
+"step' S (IF b THEN c1 ELSE c2 {P}) =
+  (let c1' = step' (bfilter b True S) c1; c2' = step' (bfilter b False S) c2
+   in IF b THEN c1' ELSE c2' {post c1 \<squnion> post c2})" |
+"step' S ({Inv} WHILE b DO c {P}) =
+   {S \<squnion> post c}
+   WHILE b DO step' (bfilter b True Inv) c
+   {bfilter b False Inv}"
+
+definition AI :: "com \<Rightarrow> 'av st option acom option" where
+"AI = lpfp\<^isub>c (step' \<top>)"
+
+lemma strip_step'[simp]: "strip(step' S c) = strip c"
+by(induct c arbitrary: S) (simp_all add: Let_def)
+
+
+subsubsection "Soundness"
+
+lemma in_gamma_update:
+  "\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(update S x a)"
+by(simp add: \<gamma>_st_def lookup_update)
+
+lemma step_preserves_le:
+  "\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; cs \<le> \<gamma>\<^isub>c ca \<rbrakk> \<Longrightarrow> step S cs \<le> \<gamma>\<^isub>c (step' S' ca)"
+proof(induction cs arbitrary: ca 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 intro: aval'_sound in_gamma_update
+      split: option.splits del:subsetD)
+next
+  case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi)
+    by (metis le_post post_map_acom)
+next
+  case (If b cs1 cs2 P)
+  then obtain ca1 ca2 Pa where
+      "ca= IF b THEN ca1 ELSE ca2 {Pa}"
+      "P \<subseteq> \<gamma>\<^isub>o Pa" "cs1 \<le> \<gamma>\<^isub>c ca1" "cs2 \<le> \<gamma>\<^isub>c ca2"
+    by (fastforce simp: If_le map_acom_If)
+  moreover have "post cs1 \<subseteq> \<gamma>\<^isub>o(post ca1 \<squnion> post ca2)"
+    by (metis (no_types) `cs1 \<le> \<gamma>\<^isub>c ca1` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
+  moreover have "post cs2 \<subseteq> \<gamma>\<^isub>o(post ca1 \<squnion> post ca2)"
+    by (metis (no_types) `cs2 \<le> \<gamma>\<^isub>c ca2` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
+  ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'`
+    by (simp add: If.IH subset_iff bfilter_sound)
+next
+  case (While I b cs1 P)
+  then obtain ca1 Ia Pa where
+    "ca = {Ia} WHILE b DO ca1 {Pa}"
+    "I \<subseteq> \<gamma>\<^isub>o Ia" "P \<subseteq> \<gamma>\<^isub>o Pa" "cs1 \<le> \<gamma>\<^isub>c ca1"
+    by (fastforce simp: map_acom_While While_le)
+  moreover have "S \<union> post cs1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post ca1)"
+    using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `cs1 \<le> \<gamma>\<^isub>c ca1`, simplified]
+    by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans)
+  ultimately show ?case by (simp add: While.IH subset_iff bfilter_sound)
+qed
+
+lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
+proof(simp add: CS_def AI_def)
+  assume 1: "lpfp\<^isub>c (step' \<top>) c = Some c'"
+  have 2: "step' \<top> c' \<sqsubseteq> c'" by(rule lpfpc_pfp[OF 1])
+  have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c"
+    by(simp add: strip_lpfpc[OF _ 1])
+  have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
+  proof(rule lfp_lowerbound[simplified,OF 3])
+    show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
+    proof(rule step_preserves_le[OF _ _])
+      show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
+      show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
+    qed
+  qed
+  from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
+    by (blast intro: mono_gamma_c order_trans)
+qed
+
+
+subsubsection "Commands over a set of variables"
+
+text{* Key invariant: the domains of all abstract states are subsets of the
+set of variables of the program. *}
+
+definition "domo S = (case S of None \<Rightarrow> {} | Some S' \<Rightarrow> set(dom S'))"
+
+definition Com :: "vname set \<Rightarrow> 'a st option acom set" where
+"Com X = {c. \<forall>S \<in> set(annos c). domo S \<subseteq> X}"
+
+lemma domo_Top[simp]: "domo \<top> = {}"
+by(simp add: domo_def Top_st_def Top_option_def)
+
+lemma bot_acom_Com[simp]: "\<bottom>\<^sub>c c \<in> Com X"
+by(simp add: bot_acom_def Com_def domo_def set_annos_anno)
+
+lemma post_in_annos: "post c : set(annos c)"
+by(induction c) simp_all
+
+lemma domo_join: "domo (S \<squnion> T) \<subseteq> domo S \<union> domo T"
+by(auto simp: domo_def join_st_def split: option.split)
+
+lemma domo_afilter: "vars a \<subseteq> X \<Longrightarrow> domo S \<subseteq> X \<Longrightarrow> domo(afilter a i S) \<subseteq> X"
+apply(induction a arbitrary: i S)
+apply(simp add: domo_def)
+apply(simp add: domo_def Let_def update_def lookup_def split: option.splits)
+apply blast
+apply(simp split: prod.split)
+done
+
+lemma domo_bfilter: "vars b \<subseteq> X \<Longrightarrow> domo S \<subseteq> X \<Longrightarrow> domo(bfilter b bv S) \<subseteq> X"
+apply(induction b arbitrary: bv S)
+apply(simp add: domo_def)
+apply(simp)
+apply(simp)
+apply rule
+apply (metis le_sup_iff subset_trans[OF domo_join])
+apply(simp split: prod.split)
+by (metis domo_afilter)
+
+lemma step'_Com:
+  "domo S \<subseteq> X \<Longrightarrow> vars(strip c) \<subseteq> X \<Longrightarrow> c : Com X \<Longrightarrow> step' S c : Com X"
+apply(induction c arbitrary: S)
+apply(simp add: Com_def)
+apply(simp add: Com_def domo_def update_def split: option.splits)
+apply(simp (no_asm_use) add: Com_def ball_Un)
+apply (metis post_in_annos)
+apply(simp (no_asm_use) add: Com_def ball_Un)
+apply rule
+apply (metis Un_assoc domo_join order_trans post_in_annos subset_Un_eq)
+apply (metis domo_bfilter)
+apply(simp (no_asm_use) add: Com_def)
+apply rule
+apply (metis domo_join le_sup_iff post_in_annos subset_trans)
+apply rule
+apply (metis domo_bfilter)
+by (metis domo_bfilter)
+
+end
+
+
+subsubsection "Monotonicity"
+
+locale Abs_Int1_mono = Abs_Int1 +
+assumes mono_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> plus' a1 a2 \<sqsubseteq> plus' b1 b2"
+and mono_filter_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> r \<sqsubseteq> r' \<Longrightarrow>
+  filter_plus' r a1 a2 \<sqsubseteq> filter_plus' r' b1 b2"
+and mono_filter_less': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow>
+  filter_less' bv a1 a2 \<sqsubseteq> filter_less' bv b1 b2"
+begin
+
+lemma mono_aval': "S \<sqsubseteq> S' \<Longrightarrow> aval' e S \<sqsubseteq> aval' e S'"
+by(induction e) (auto simp: le_st_def lookup_def mono_plus')
+
+lemma mono_aval'': "S \<sqsubseteq> S' \<Longrightarrow> aval'' e S \<sqsubseteq> aval'' e S'"
+apply(cases S)
+ apply simp
+apply(cases S')
+ apply simp
+by (simp add: mono_aval')
+
+lemma mono_afilter: "r \<sqsubseteq> r' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> afilter e r S \<sqsubseteq> afilter e r' S'"
+apply(induction e arbitrary: r r' S S')
+apply(auto simp: test_num' Let_def split: option.splits prod.splits)
+apply (metis mono_gamma subsetD)
+apply(drule_tac x = "list" in mono_lookup)
+apply (metis mono_meet le_trans)
+apply (metis mono_meet mono_lookup mono_update)
+apply(metis mono_aval'' mono_filter_plus'[simplified le_prod_def] fst_conv snd_conv)
+done
+
+lemma mono_bfilter: "S \<sqsubseteq> S' \<Longrightarrow> bfilter b r S \<sqsubseteq> bfilter b r S'"
+apply(induction b arbitrary: r S S')
+apply(auto simp: le_trans[OF _ join_ge1] le_trans[OF _ join_ge2] split: prod.splits)
+apply(metis mono_aval'' mono_afilter mono_filter_less'[simplified le_prod_def] fst_conv snd_conv)
+done
+
+lemma mono_step': "S \<sqsubseteq> S' \<Longrightarrow> c \<sqsubseteq> c' \<Longrightarrow> step' S c \<sqsubseteq> step' S' c'"
+apply(induction c c' arbitrary: S S' rule: le_acom.induct)
+apply (auto simp: mono_post mono_bfilter mono_update mono_aval' Let_def le_join_disj
+  split: option.split)
+done
+
+lemma mono_step'2: "mono (step' S)"
+by(simp add: mono_def mono_step'[OF le_refl])
+
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ivl_ITP.thy	Thu Apr 19 17:32:30 2012 +0200
@@ -0,0 +1,285 @@
+(* Author: Tobias Nipkow *)
+
+theory Abs_Int2_ivl_ITP
+imports Abs_Int2_ITP Abs_Int_Tests
+begin
+
+subsection "Interval Analysis"
+
+datatype ivl = I "int option" "int option"
+
+definition "\<gamma>_ivl i = (case i of
+  I (Some l) (Some h) \<Rightarrow> {l..h} |
+  I (Some l) None \<Rightarrow> {l..} |
+  I None (Some h) \<Rightarrow> {..h} |
+  I None None \<Rightarrow> UNIV)"
+
+abbreviation I_Some_Some :: "int \<Rightarrow> int \<Rightarrow> ivl"  ("{_\<dots>_}") where
+"{lo\<dots>hi} == I (Some lo) (Some hi)"
+abbreviation I_Some_None :: "int \<Rightarrow> ivl"  ("{_\<dots>}") where
+"{lo\<dots>} == I (Some lo) None"
+abbreviation I_None_Some :: "int \<Rightarrow> ivl"  ("{\<dots>_}") where
+"{\<dots>hi} == I None (Some hi)"
+abbreviation I_None_None :: "ivl"  ("{\<dots>}") where
+"{\<dots>} == I None None"
+
+definition "num_ivl n = {n\<dots>n}"
+
+fun in_ivl :: "int \<Rightarrow> ivl \<Rightarrow> bool" where
+"in_ivl k (I (Some l) (Some h)) \<longleftrightarrow> l \<le> k \<and> k \<le> h" |
+"in_ivl k (I (Some l) None) \<longleftrightarrow> l \<le> k" |
+"in_ivl k (I None (Some h)) \<longleftrightarrow> k \<le> h" |
+"in_ivl k (I None None) \<longleftrightarrow> True"
+
+instantiation option :: (plus)plus
+begin
+
+fun plus_option where
+"Some x + Some y = Some(x+y)" |
+"_ + _ = None"
+
+instance ..
+
+end
+
+definition empty where "empty = {1\<dots>0}"
+
+fun is_empty where
+"is_empty {l\<dots>h} = (h<l)" |
+"is_empty _ = False"
+
+lemma [simp]: "is_empty(I l h) =
+  (case l of Some l \<Rightarrow> (case h of Some h \<Rightarrow> h<l | None \<Rightarrow> False) | None \<Rightarrow> False)"
+by(auto split:option.split)
+
+lemma [simp]: "is_empty i \<Longrightarrow> \<gamma>_ivl i = {}"
+by(auto simp add: \<gamma>_ivl_def split: ivl.split option.split)
+
+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) \<Rightarrow> I (l1+l2) (h1+h2))"
+
+instantiation ivl :: SL_top
+begin
+
+definition le_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> bool" where
+"le_option pos x y =
+ (case x of (Some i) \<Rightarrow> (case y of Some j \<Rightarrow> i\<le>j | None \<Rightarrow> pos)
+  | None \<Rightarrow> (case y of Some j \<Rightarrow> \<not>pos | None \<Rightarrow> True))"
+
+fun le_aux where
+"le_aux (I l1 h1) (I l2 h2) = (le_option False l2 l1 & le_option True h1 h2)"
+
+definition le_ivl where
+"i1 \<sqsubseteq> i2 =
+ (if is_empty i1 then True else
+  if is_empty i2 then False else le_aux i1 i2)"
+
+definition min_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> int option" where
+"min_option pos o1 o2 = (if le_option pos o1 o2 then o1 else o2)"
+
+definition max_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> int option" where
+"max_option pos o1 o2 = (if le_option pos o1 o2 then o2 else o1)"
+
+definition "i1 \<squnion> i2 =
+ (if is_empty i1 then i2 else if is_empty i2 then i1
+  else case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow>
+          I (min_option False l1 l2) (max_option True h1 h2))"
+
+definition "\<top> = {\<dots>}"
+
+instance
+proof
+  case goal1 thus ?case
+    by(cases x, simp add: le_ivl_def le_option_def split: option.split)
+next
+  case goal2 thus ?case
+    by(cases x, cases y, cases z, auto simp: le_ivl_def le_option_def split: option.splits if_splits)
+next
+  case goal3 thus ?case
+    by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits)
+next
+  case goal4 thus ?case
+    by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits)
+next
+  case goal5 thus ?case
+    by(cases x, cases y, cases z, auto simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits if_splits)
+next
+  case goal6 thus ?case
+    by(cases x, simp add: Top_ivl_def le_ivl_def le_option_def split: option.split)
+qed
+
+end
+
+
+instantiation ivl :: L_top_bot
+begin
+
+definition "i1 \<sqinter> i2 = (if is_empty i1 \<or> is_empty i2 then empty else
+  case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow>
+    I (max_option False l1 l2) (min_option True h1 h2))"
+
+definition "\<bottom> = empty"
+
+instance
+proof
+  case goal1 thus ?case
+    by (simp add:meet_ivl_def empty_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits)
+next
+  case goal2 thus ?case
+    by (simp add: empty_def meet_ivl_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits)
+next
+  case goal3 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 goal4 show ?case by(cases x, simp add: bot_ivl_def empty_def le_ivl_def)
+qed
+
+end
+
+instantiation option :: (minus)minus
+begin
+
+fun minus_option where
+"Some x - Some y = Some(x-y)" |
+"_ - _ = None"
+
+instance ..
+
+end
+
+definition "minus_ivl i1 i2 = (if is_empty i1 | is_empty i2 then empty else
+  case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow> I (l1-h2) (h1-l2))"
+
+lemma gamma_minus_ivl:
+  "n1 : \<gamma>_ivl i1 \<Longrightarrow> n2 : \<gamma>_ivl i2 \<Longrightarrow> n1-n2 : \<gamma>_ivl(minus_ivl i1 i2)"
+by(auto simp add: minus_ivl_def \<gamma>_ivl_def split: ivl.splits option.splits)
+
+definition "filter_plus_ivl i i1 i2 = ((*if is_empty i then empty else*)
+  i1 \<sqinter> minus_ivl i i2, i2 \<sqinter> minus_ivl i i1)"
+
+fun filter_less_ivl :: "bool \<Rightarrow> ivl \<Rightarrow> ivl \<Rightarrow> ivl * ivl" where
+"filter_less_ivl res (I l1 h1) (I l2 h2) =
+  (if is_empty(I l1 h1) \<or> is_empty(I l2 h2) then (empty, empty) else
+   if res
+   then (I l1 (min_option True h1 (h2 - Some 1)),
+         I (max_option False (l1 + Some 1) l2) h2)
+   else (I (max_option False l1 l2) h1, I l2 (min_option True h1 h2)))"
+
+interpretation Val_abs
+where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = plus_ivl
+proof
+  case goal1 thus ?case
+    by(auto simp: \<gamma>_ivl_def le_ivl_def le_option_def split: ivl.split option.split if_splits)
+next
+  case goal2 show ?case by(simp add: \<gamma>_ivl_def Top_ivl_def)
+next
+  case goal3 thus ?case by(simp add: \<gamma>_ivl_def num_ivl_def)
+next
+  case goal4 thus ?case
+    by(auto simp add: \<gamma>_ivl_def plus_ivl_def split: ivl.split option.splits)
+qed
+
+interpretation Val_abs1_gamma
+where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = plus_ivl
+defines aval_ivl is aval'
+proof
+  case goal1 thus ?case
+    by(auto simp add: \<gamma>_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 \<gamma>_ivl_def empty_def)
+qed
+
+lemma mono_minus_ivl:
+  "i1 \<sqsubseteq> i1' \<Longrightarrow> i2 \<sqsubseteq> i2' \<Longrightarrow> minus_ivl i1 i2 \<sqsubseteq> minus_ivl i1' i2'"
+apply(auto simp add: minus_ivl_def empty_def le_ivl_def le_option_def split: ivl.splits)
+  apply(simp split: option.splits)
+ apply(simp split: option.splits)
+apply(simp split: option.splits)
+done
+
+
+interpretation Val_abs1
+where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = plus_ivl
+and test_num' = in_ivl
+and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl
+proof
+  case goal1 thus ?case
+    by (simp add: \<gamma>_ivl_def split: ivl.split option.split)
+next
+  case goal2 thus ?case
+    by(auto simp add: filter_plus_ivl_def)
+      (metis gamma_minus_ivl add_diff_cancel add_commute)+
+next
+  case goal3 thus ?case
+    by(cases a1, cases a2,
+      auto simp: \<gamma>_ivl_def min_option_def max_option_def le_option_def split: if_splits option.splits)
+qed
+
+interpretation Abs_Int1
+where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = plus_ivl
+and test_num' = in_ivl
+and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl
+defines afilter_ivl is afilter
+and bfilter_ivl is bfilter
+and step_ivl is step'
+and AI_ivl is AI
+and aval_ivl' is aval''
+..
+
+
+text{* Monotonicity: *}
+
+interpretation Abs_Int1_mono
+where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = plus_ivl
+and test_num' = in_ivl
+and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl
+proof
+  case goal1 thus ?case
+    by(auto simp: plus_ivl_def le_ivl_def le_option_def empty_def split: if_splits ivl.splits option.splits)
+next
+  case goal2 thus ?case
+    by(auto simp: filter_plus_ivl_def le_prod_def mono_meet mono_minus_ivl)
+next
+  case goal3 thus ?case
+    apply(cases a1, cases b1, cases a2, cases b2, auto simp: le_prod_def)
+    by(auto simp add: empty_def le_ivl_def le_option_def min_option_def max_option_def split: option.splits)
+qed
+
+
+subsubsection "Tests"
+
+value "show_acom_opt (AI_ivl test1_ivl)"
+
+text{* Better than @{text AI_const}: *}
+value "show_acom_opt (AI_ivl test3_const)"
+value "show_acom_opt (AI_ivl test4_const)"
+value "show_acom_opt (AI_ivl test6_const)"
+
+value "show_acom_opt (AI_ivl test2_ivl)"
+value "show_acom (((step_ivl \<top>)^^0) (\<bottom>\<^sub>c test2_ivl))"
+value "show_acom (((step_ivl \<top>)^^1) (\<bottom>\<^sub>c test2_ivl))"
+value "show_acom (((step_ivl \<top>)^^2) (\<bottom>\<^sub>c test2_ivl))"
+
+text{* Fixed point reached in 2 steps. Not so if the start value of x is known: *}
+
+value "show_acom_opt (AI_ivl test3_ivl)"
+value "show_acom (((step_ivl \<top>)^^0) (\<bottom>\<^sub>c test3_ivl))"
+value "show_acom (((step_ivl \<top>)^^1) (\<bottom>\<^sub>c test3_ivl))"
+value "show_acom (((step_ivl \<top>)^^2) (\<bottom>\<^sub>c test3_ivl))"
+value "show_acom (((step_ivl \<top>)^^3) (\<bottom>\<^sub>c test3_ivl))"
+value "show_acom (((step_ivl \<top>)^^4) (\<bottom>\<^sub>c test3_ivl))"
+
+text{* Takes as many iterations as the actual execution. Would diverge if
+loop did not terminate. Worse still, as the following example shows: even if
+the actual execution terminates, the analysis may not. The value of y keeps
+decreasing as the analysis is iterated, no matter how long: *}
+
+value "show_acom (((step_ivl \<top>)^^50) (\<bottom>\<^sub>c test4_ivl))"
+
+text{* Relationships between variables are NOT captured: *}
+value "show_acom_opt (AI_ivl test5_ivl)"
+
+text{* Again, the analysis would not terminate: *}
+value "show_acom (((step_ivl \<top>)^^50) (\<bottom>\<^sub>c test6_ivl))"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy	Thu Apr 19 17:32:30 2012 +0200
@@ -0,0 +1,683 @@
+(* Author: Tobias Nipkow *)
+
+theory Abs_Int3_ITP
+imports Abs_Int2_ivl_ITP
+begin
+
+subsection "Widening and Narrowing"
+
+class WN = SL_top +
+fixes widen :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<nabla>" 65)
+assumes widen1: "x \<sqsubseteq> x \<nabla> y"
+assumes widen2: "y \<sqsubseteq> x \<nabla> y"
+fixes narrow :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<triangle>" 65)
+assumes narrow1: "y \<sqsubseteq> x \<Longrightarrow> y \<sqsubseteq> x \<triangle> y"
+assumes narrow2: "y \<sqsubseteq> x \<Longrightarrow> x \<triangle> y \<sqsubseteq> x"
+
+
+subsubsection "Intervals"
+
+instantiation ivl :: WN
+begin
+
+definition "widen_ivl ivl1 ivl2 =
+  ((*if is_empty ivl1 then ivl2 else
+   if is_empty ivl2 then ivl1 else*)
+     case (ivl1,ivl2) of (I l1 h1, I l2 h2) \<Rightarrow>
+       I (if le_option False l2 l1 \<and> l2 \<noteq> l1 then None else l1)
+         (if le_option True h1 h2 \<and> h1 \<noteq> h2 then None else h1))"
+
+definition "narrow_ivl ivl1 ivl2 =
+  ((*if is_empty ivl1 \<or> is_empty ivl2 then empty else*)
+     case (ivl1,ivl2) of (I l1 h1, I l2 h2) \<Rightarrow>
+       I (if l1 = None then l2 else l1)
+         (if h1 = None then h2 else h1))"
+
+instance
+proof qed
+  (auto simp add: widen_ivl_def narrow_ivl_def le_option_def le_ivl_def empty_def split: ivl.split option.split if_splits)
+
+end
+
+
+subsubsection "Abstract State"
+
+instantiation st :: (WN)WN
+begin
+
+definition "widen_st F1 F2 =
+  FunDom (\<lambda>x. fun F1 x \<nabla> fun F2 x) (inter_list (dom F1) (dom F2))"
+
+definition "narrow_st F1 F2 =
+  FunDom (\<lambda>x. fun F1 x \<triangle> fun F2 x) (inter_list (dom F1) (dom F2))"
+
+instance
+proof
+  case goal1 thus ?case
+    by(simp add: widen_st_def le_st_def lookup_def widen1)
+next
+  case goal2 thus ?case
+    by(simp add: widen_st_def le_st_def lookup_def widen2)
+next
+  case goal3 thus ?case
+    by(auto simp: narrow_st_def le_st_def lookup_def narrow1)
+next
+  case goal4 thus ?case
+    by(auto simp: narrow_st_def le_st_def lookup_def narrow2)
+qed
+
+end
+
+
+subsubsection "Option"
+
+instantiation option :: (WN)WN
+begin
+
+fun widen_option where
+"None \<nabla> x = x" |
+"x \<nabla> None = x" |
+"(Some x) \<nabla> (Some y) = Some(x \<nabla> y)"
+
+fun narrow_option where
+"None \<triangle> x = None" |
+"x \<triangle> None = None" |
+"(Some x) \<triangle> (Some y) = Some(x \<triangle> y)"
+
+instance
+proof
+  case goal1 show ?case
+    by(induct x y rule: widen_option.induct) (simp_all add: widen1)
+next
+  case goal2 show ?case
+    by(induct x y rule: widen_option.induct) (simp_all add: widen2)
+next
+  case goal3 thus ?case
+    by(induct x y rule: narrow_option.induct) (simp_all add: narrow1)
+next
+  case goal4 thus ?case
+    by(induct x y rule: narrow_option.induct) (simp_all add: narrow2)
+qed
+
+end
+
+
+subsubsection "Annotated commands"
+
+fun map2_acom :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a acom \<Rightarrow> 'a acom \<Rightarrow> 'a acom" where
+"map2_acom f (SKIP {a1}) (SKIP {a2}) = (SKIP {f a1 a2})" |
+"map2_acom f (x ::= e {a1}) (x' ::= e' {a2}) = (x ::= e {f a1 a2})" |
+"map2_acom f (c1;c2) (c1';c2') = (map2_acom f c1 c1'; map2_acom f c2 c2')" |
+"map2_acom f (IF b THEN c1 ELSE c2 {a1}) (IF b' THEN c1' ELSE c2' {a2}) =
+  (IF b THEN map2_acom f c1 c1' ELSE map2_acom f c2 c2' {f a1 a2})" |
+"map2_acom f ({a1} WHILE b DO c {a2}) ({a3} WHILE b' DO c' {a4}) =
+  ({f a1 a3} WHILE b DO map2_acom f c c' {f a2 a4})"
+
+abbreviation widen_acom :: "('a::WN)acom \<Rightarrow> 'a acom \<Rightarrow> 'a acom" (infix "\<nabla>\<^sub>c" 65)
+where "widen_acom == map2_acom (op \<nabla>)"
+
+abbreviation narrow_acom :: "('a::WN)acom \<Rightarrow> 'a acom \<Rightarrow> 'a acom" (infix "\<triangle>\<^sub>c" 65)
+where "narrow_acom == map2_acom (op \<triangle>)"
+
+lemma widen1_acom: "strip c = strip c' \<Longrightarrow> c \<sqsubseteq> c \<nabla>\<^sub>c c'"
+by(induct c c' rule: le_acom.induct)(simp_all add: widen1)
+
+lemma widen2_acom: "strip c = strip c' \<Longrightarrow> c' \<sqsubseteq> c \<nabla>\<^sub>c c'"
+by(induct c c' rule: le_acom.induct)(simp_all add: widen2)
+
+lemma narrow1_acom: "y \<sqsubseteq> x \<Longrightarrow> y \<sqsubseteq> x \<triangle>\<^sub>c y"
+by(induct y x rule: le_acom.induct) (simp_all add: narrow1)
+
+lemma narrow2_acom: "y \<sqsubseteq> x \<Longrightarrow> x \<triangle>\<^sub>c y \<sqsubseteq> x"
+by(induct y x rule: le_acom.induct) (simp_all add: narrow2)
+
+
+subsubsection "Post-fixed point computation"
+
+definition iter_widen :: "('a acom \<Rightarrow> 'a acom) \<Rightarrow> 'a acom \<Rightarrow> ('a::WN)acom option"
+where "iter_widen f = while_option (\<lambda>c. \<not> f c \<sqsubseteq> c) (\<lambda>c. c \<nabla>\<^sub>c f c)"
+
+definition iter_narrow :: "('a acom \<Rightarrow> 'a acom) \<Rightarrow> 'a acom \<Rightarrow> 'a::WN acom option"
+where "iter_narrow f = while_option (\<lambda>c. \<not> c \<sqsubseteq> c \<triangle>\<^sub>c f c) (\<lambda>c. c \<triangle>\<^sub>c f c)"
+
+definition pfp_wn ::
+  "(('a::WN)option acom \<Rightarrow> 'a option acom) \<Rightarrow> com \<Rightarrow> 'a option acom option"
+where "pfp_wn f c = (case iter_widen f (\<bottom>\<^sub>c c) of None \<Rightarrow> None
+                     | Some c' \<Rightarrow> iter_narrow f c')"
+
+lemma strip_map2_acom:
+ "strip c1 = strip c2 \<Longrightarrow> strip(map2_acom f c1 c2) = strip c1"
+by(induct f c1 c2 rule: map2_acom.induct) simp_all
+
+lemma iter_widen_pfp: "iter_widen f c = Some c' \<Longrightarrow> f c' \<sqsubseteq> c'"
+by(auto simp add: iter_widen_def dest: while_option_stop)
+
+lemma strip_while: fixes f :: "'a acom \<Rightarrow> 'a acom"
+assumes "\<forall>c. strip (f c) = strip c" and "while_option P f c = Some c'"
+shows "strip c' = strip c"
+using while_option_rule[where P = "\<lambda>c'. strip c' = strip c", OF _ assms(2)]
+by (metis assms(1))
+
+lemma strip_iter_widen: fixes f :: "'a::WN acom \<Rightarrow> 'a acom"
+assumes "\<forall>c. strip (f c) = strip c" and "iter_widen f c = Some c'"
+shows "strip c' = strip c"
+proof-
+  have "\<forall>c. strip(c \<nabla>\<^sub>c f c) = strip c" by (metis assms(1) strip_map2_acom)
+  from strip_while[OF this] assms(2) show ?thesis by(simp add: iter_widen_def)
+qed
+
+lemma iter_narrow_pfp: assumes "mono f" and "f c0 \<sqsubseteq> c0"
+and "iter_narrow f c0 = Some c"
+shows "f c \<sqsubseteq> c \<and> c \<sqsubseteq> c0" (is "?P c")
+proof-
+  { fix c assume "?P c"
+    note 1 = conjunct1[OF this] and 2 = conjunct2[OF this]
+    let ?c' = "c \<triangle>\<^sub>c f c"
+    have "?P ?c'"
+    proof
+      have "f ?c' \<sqsubseteq> f c"  by(rule monoD[OF `mono f` narrow2_acom[OF 1]])
+      also have "\<dots> \<sqsubseteq> ?c'" by(rule narrow1_acom[OF 1])
+      finally show "f ?c' \<sqsubseteq> ?c'" .
+      have "?c' \<sqsubseteq> c" by (rule narrow2_acom[OF 1])
+      also have "c \<sqsubseteq> c0" by(rule 2)
+      finally show "?c' \<sqsubseteq> c0" .
+    qed
+  }
+  with while_option_rule[where P = ?P, OF _ assms(3)[simplified iter_narrow_def]]
+    assms(2) le_refl
+  show ?thesis by blast
+qed
+
+lemma pfp_wn_pfp:
+  "\<lbrakk> mono f;  pfp_wn f c = Some c' \<rbrakk> \<Longrightarrow> f c' \<sqsubseteq> c'"
+unfolding pfp_wn_def
+by (auto dest: iter_widen_pfp iter_narrow_pfp split: option.splits)
+
+lemma strip_pfp_wn:
+  "\<lbrakk> \<forall>c. strip(f c) = strip c; pfp_wn f c = Some c' \<rbrakk> \<Longrightarrow> strip c' = c"
+apply(auto simp add: pfp_wn_def iter_narrow_def split: option.splits)
+by (metis (no_types) strip_map2_acom strip_while strip_bot_acom strip_iter_widen)
+
+locale Abs_Int2 = Abs_Int1_mono
+where \<gamma>=\<gamma> for \<gamma> :: "'av::{WN,L_top_bot} \<Rightarrow> val set"
+begin
+
+definition AI_wn :: "com \<Rightarrow> 'av st option acom option" where
+"AI_wn = pfp_wn (step' \<top>)"
+
+lemma AI_wn_sound: "AI_wn c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
+proof(simp add: CS_def AI_wn_def)
+  assume 1: "pfp_wn (step' \<top>) c = Some c'"
+  from pfp_wn_pfp[OF mono_step'2 1]
+  have 2: "step' \<top> c' \<sqsubseteq> c'" .
+  have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c" by(simp add: strip_pfp_wn[OF _ 1])
+  have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
+  proof(rule lfp_lowerbound[simplified,OF 3])
+    show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
+    proof(rule step_preserves_le[OF _ _])
+      show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
+      show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
+    qed
+  qed
+  from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
+    by (blast intro: mono_gamma_c order_trans)
+qed
+
+end
+
+interpretation Abs_Int2
+where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = plus_ivl
+and test_num' = in_ivl
+and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl
+defines AI_ivl' is AI_wn
+..
+
+
+subsubsection "Tests"
+
+definition "step_up_ivl n = ((\<lambda>c. c \<nabla>\<^sub>c step_ivl \<top> c)^^n)"
+definition "step_down_ivl n = ((\<lambda>c. c \<triangle>\<^sub>c step_ivl \<top> c)^^n)"
+
+text{* For @{const test3_ivl}, @{const AI_ivl} needed as many iterations as
+the loop took to execute. In contrast, @{const AI_ivl'} converges in a
+constant number of steps: *}
+
+value "show_acom (step_up_ivl 1 (\<bottom>\<^sub>c test3_ivl))"
+value "show_acom (step_up_ivl 2 (\<bottom>\<^sub>c test3_ivl))"
+value "show_acom (step_up_ivl 3 (\<bottom>\<^sub>c test3_ivl))"
+value "show_acom (step_up_ivl 4 (\<bottom>\<^sub>c test3_ivl))"
+value "show_acom (step_up_ivl 5 (\<bottom>\<^sub>c test3_ivl))"
+value "show_acom (step_down_ivl 1 (step_up_ivl 5 (\<bottom>\<^sub>c test3_ivl)))"
+value "show_acom (step_down_ivl 2 (step_up_ivl 5 (\<bottom>\<^sub>c test3_ivl)))"
+value "show_acom (step_down_ivl 3 (step_up_ivl 5 (\<bottom>\<^sub>c test3_ivl)))"
+
+text{* Now all the analyses terminate: *}
+
+value "show_acom_opt (AI_ivl' test4_ivl)"
+value "show_acom_opt (AI_ivl' test5_ivl)"
+value "show_acom_opt (AI_ivl' test6_ivl)"
+
+
+subsubsection "Termination: Intervals"
+
+definition m_ivl :: "ivl \<Rightarrow> nat" where
+"m_ivl ivl = (case ivl of I l h \<Rightarrow>
+     (case l of None \<Rightarrow> 0 | Some _ \<Rightarrow> 1) + (case h of None \<Rightarrow> 0 | Some _ \<Rightarrow> 1))"
+
+lemma m_ivl_height: "m_ivl ivl \<le> 2"
+by(simp add: m_ivl_def split: ivl.split option.split)
+
+lemma m_ivl_anti_mono: "(y::ivl) \<sqsubseteq> x \<Longrightarrow> m_ivl x \<le> m_ivl y"
+by(auto simp: m_ivl_def le_option_def le_ivl_def
+        split: ivl.split option.split if_splits)
+
+lemma m_ivl_widen:
+  "~ y \<sqsubseteq> x \<Longrightarrow> m_ivl(x \<nabla> y) < m_ivl x"
+by(auto simp: m_ivl_def widen_ivl_def le_option_def le_ivl_def
+        split: ivl.splits option.splits if_splits)
+
+lemma Top_less_ivl: "\<top> \<sqsubseteq> x \<Longrightarrow> m_ivl x = 0"
+by(auto simp: m_ivl_def le_option_def le_ivl_def empty_def Top_ivl_def
+        split: ivl.split option.split if_splits)
+
+
+definition n_ivl :: "ivl \<Rightarrow> nat" where
+"n_ivl ivl = 2 - m_ivl ivl"
+
+lemma n_ivl_mono: "(x::ivl) \<sqsubseteq> y \<Longrightarrow> n_ivl x \<le> n_ivl y"
+unfolding n_ivl_def by (metis diff_le_mono2 m_ivl_anti_mono)
+
+lemma n_ivl_narrow:
+  "~ x \<sqsubseteq> x \<triangle> y \<Longrightarrow> n_ivl(x \<triangle> y) < n_ivl x"
+by(auto simp: n_ivl_def m_ivl_def narrow_ivl_def le_option_def le_ivl_def
+        split: ivl.splits option.splits if_splits)
+
+
+subsubsection "Termination: Abstract State"
+
+definition "m_st m st = (\<Sum>x\<in>set(dom st). m(fun st x))"
+
+lemma m_st_height: assumes "finite X" and "set (dom S) \<subseteq> X"
+shows "m_st m_ivl S \<le> 2 * card X"
+proof(auto simp: m_st_def)
+  have "(\<Sum>x\<in>set(dom S). m_ivl (fun S x)) \<le> (\<Sum>x\<in>set(dom S). 2)" (is "?L \<le> _")
+    by(rule setsum_mono)(simp add:m_ivl_height)
+  also have "\<dots> \<le> (\<Sum>x\<in>X. 2)"
+    by(rule setsum_mono3[OF assms]) simp
+  also have "\<dots> = 2 * card X" by(simp add: setsum_constant)
+  finally show "?L \<le> \<dots>" .
+qed
+
+lemma m_st_anti_mono:
+  "S1 \<sqsubseteq> S2 \<Longrightarrow> m_st m_ivl S2 \<le> m_st m_ivl S1"
+proof(auto simp: m_st_def le_st_def lookup_def split: if_splits)
+  let ?X = "set(dom S1)" let ?Y = "set(dom S2)"
+  let ?f = "fun S1" let ?g = "fun S2"
+  assume asm: "\<forall>x\<in>?Y. (x \<in> ?X \<longrightarrow> ?f x \<sqsubseteq> ?g x) \<and> (x \<in> ?X \<or> \<top> \<sqsubseteq> ?g x)"
+  hence 1: "\<forall>y\<in>?Y\<inter>?X. m_ivl(?g y) \<le> m_ivl(?f y)" by(simp add: m_ivl_anti_mono)
+  have 0: "\<forall>x\<in>?Y-?X. m_ivl(?g x) = 0" using asm by (auto simp: Top_less_ivl)
+  have "(\<Sum>y\<in>?Y. m_ivl(?g y)) = (\<Sum>y\<in>(?Y-?X) \<union> (?Y\<inter>?X). m_ivl(?g y))"
+    by (metis Un_Diff_Int)
+  also have "\<dots> = (\<Sum>y\<in>?Y-?X. m_ivl(?g y)) + (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?g y))"
+    by(subst setsum_Un_disjoint) auto
+  also have "(\<Sum>y\<in>?Y-?X. m_ivl(?g y)) = 0" using 0 by simp
+  also have "0 + (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?g y)) = (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?g y))" by simp
+  also have "\<dots> \<le> (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?f y))"
+    by(rule setsum_mono)(simp add: 1)
+  also have "\<dots> \<le> (\<Sum>y\<in>?X. m_ivl(?f y))"
+    by(simp add: setsum_mono3[of "?X" "?Y Int ?X", OF _ Int_lower2])
+  finally show "(\<Sum>y\<in>?Y. m_ivl(?g y)) \<le> (\<Sum>x\<in>?X. m_ivl(?f x))"
+    by (metis add_less_cancel_left)
+qed
+
+lemma m_st_widen:
+assumes "\<not> S2 \<sqsubseteq> S1" shows "m_st m_ivl (S1 \<nabla> S2) < m_st m_ivl S1"
+proof-
+  { let ?X = "set(dom S1)" let ?Y = "set(dom S2)"
+    let ?f = "fun S1" let ?g = "fun S2"
+    fix x assume "x \<in> ?X" "\<not> lookup S2 x \<sqsubseteq> ?f x"
+    have "(\<Sum>x\<in>?X\<inter>?Y. m_ivl(?f x \<nabla> ?g x)) < (\<Sum>x\<in>?X. m_ivl(?f x))" (is "?L < ?R")
+    proof cases
+      assume "x : ?Y"
+      have "?L < (\<Sum>x\<in>?X\<inter>?Y. m_ivl(?f x))"
+      proof(rule setsum_strict_mono1, simp)
+        show "\<forall>x\<in>?X\<inter>?Y. m_ivl(?f x \<nabla> ?g x) \<le> m_ivl (?f x)"
+          by (metis m_ivl_anti_mono widen1)
+      next
+        show "\<exists>x\<in>?X\<inter>?Y. m_ivl(?f x \<nabla> ?g x) < m_ivl(?f x)"
+          using `x:?X` `x:?Y` `\<not> lookup S2 x \<sqsubseteq> ?f x`
+          by (metis IntI m_ivl_widen lookup_def)
+      qed
+      also have "\<dots> \<le> ?R" by(simp add: setsum_mono3[OF _ Int_lower1])
+      finally show ?thesis .
+    next
+      assume "x ~: ?Y"
+      have "?L \<le> (\<Sum>x\<in>?X\<inter>?Y. m_ivl(?f x))"
+      proof(rule setsum_mono, simp)
+        fix x assume "x:?X \<and> x:?Y" show "m_ivl(?f x \<nabla> ?g x) \<le> m_ivl (?f x)"
+          by (metis m_ivl_anti_mono widen1)
+      qed
+      also have "\<dots> < m_ivl(?f x) + \<dots>"
+        using m_ivl_widen[OF `\<not> lookup S2 x \<sqsubseteq> ?f x`]
+        by (metis Nat.le_refl add_strict_increasing gr0I not_less0)
+      also have "\<dots> = (\<Sum>y\<in>insert x (?X\<inter>?Y). m_ivl(?f y))"
+        using `x ~: ?Y` by simp
+      also have "\<dots> \<le> (\<Sum>x\<in>?X. m_ivl(?f x))"
+        by(rule setsum_mono3)(insert `x:?X`, auto)
+      finally show ?thesis .
+    qed
+  } with assms show ?thesis
+    by(auto simp: le_st_def widen_st_def m_st_def Int_def)
+qed
+
+definition "n_st m X st = (\<Sum>x\<in>X. m(lookup st x))"
+
+lemma n_st_mono: assumes "set(dom S1) \<subseteq> X" "set(dom S2) \<subseteq> X" "S1 \<sqsubseteq> S2"
+shows "n_st n_ivl X S1 \<le> n_st n_ivl X S2"
+proof-
+  have "(\<Sum>x\<in>X. n_ivl(lookup S1 x)) \<le> (\<Sum>x\<in>X. n_ivl(lookup S2 x))"
+    apply(rule setsum_mono) using assms
+    by(auto simp: le_st_def lookup_def n_ivl_mono split: if_splits)
+  thus ?thesis by(simp add: n_st_def)
+qed
+
+lemma n_st_narrow:
+assumes "finite X" and "set(dom S1) \<subseteq> X" "set(dom S2) \<subseteq> X"
+and "S2 \<sqsubseteq> S1" "\<not> S1 \<sqsubseteq> S1 \<triangle> S2"
+shows "n_st n_ivl X (S1 \<triangle> S2) < n_st n_ivl X S1"
+proof-
+  have 1: "\<forall>x\<in>X. n_ivl (lookup (S1 \<triangle> S2) x) \<le> n_ivl (lookup S1 x)"
+    using assms(2-4)
+    by(auto simp: le_st_def narrow_st_def lookup_def n_ivl_mono narrow2
+            split: if_splits)
+  have 2: "\<exists>x\<in>X. n_ivl (lookup (S1 \<triangle> S2) x) < n_ivl (lookup S1 x)"
+    using assms(2-5)
+    by(auto simp: le_st_def narrow_st_def lookup_def intro: n_ivl_narrow
+            split: if_splits)
+  have "(\<Sum>x\<in>X. n_ivl(lookup (S1 \<triangle> S2) x)) < (\<Sum>x\<in>X. n_ivl(lookup S1 x))"
+    apply(rule setsum_strict_mono1[OF `finite X`]) using 1 2 by blast+
+  thus ?thesis by(simp add: n_st_def)
+qed
+
+
+subsubsection "Termination: Option"
+
+definition "m_o m n opt = (case opt of None \<Rightarrow> n+1 | Some x \<Rightarrow> m x)"
+
+lemma m_o_anti_mono: "finite X \<Longrightarrow> domo S2 \<subseteq> X \<Longrightarrow> S1 \<sqsubseteq> S2 \<Longrightarrow>
+  m_o (m_st m_ivl) (2 * card X) S2 \<le> m_o (m_st m_ivl) (2 * card X) S1"
+apply(induction S1 S2 rule: le_option.induct)
+apply(auto simp: domo_def m_o_def m_st_anti_mono le_SucI m_st_height
+           split: option.splits)
+done
+
+lemma m_o_widen: "\<lbrakk> finite X; domo S2 \<subseteq> X; \<not> S2 \<sqsubseteq> S1 \<rbrakk> \<Longrightarrow>
+  m_o (m_st m_ivl) (2 * card X) (S1 \<nabla> S2) < m_o (m_st m_ivl) (2 * card X) S1"
+by(auto simp: m_o_def domo_def m_st_height less_Suc_eq_le m_st_widen
+        split: option.split)
+
+definition "n_o n opt = (case opt of None \<Rightarrow> 0 | Some x \<Rightarrow> n x + 1)"
+
+lemma n_o_mono: "domo S1 \<subseteq> X \<Longrightarrow> domo S2 \<subseteq> X \<Longrightarrow> S1 \<sqsubseteq> S2 \<Longrightarrow>
+  n_o (n_st n_ivl X) S1 \<le> n_o (n_st n_ivl X) S2"
+apply(induction S1 S2 rule: le_option.induct)
+apply(auto simp: domo_def n_o_def n_st_mono
+           split: option.splits)
+done
+
+lemma n_o_narrow:
+  "\<lbrakk> finite X; domo S1 \<subseteq> X; domo S2 \<subseteq> X; S2 \<sqsubseteq> S1; \<not> S1 \<sqsubseteq> S1 \<triangle> S2 \<rbrakk>
+  \<Longrightarrow> n_o (n_st n_ivl X) (S1 \<triangle> S2) < n_o (n_st n_ivl X) S1"
+apply(induction S1 S2 rule: narrow_option.induct)
+apply(auto simp: n_o_def domo_def n_st_narrow)
+done
+
+lemma domo_widen_subset: "domo (S1 \<nabla> S2) \<subseteq> domo S1 \<union> domo S2"
+apply(induction S1 S2 rule: widen_option.induct)
+apply (auto simp: domo_def widen_st_def)
+done
+
+lemma domo_narrow_subset: "domo (S1 \<triangle> S2) \<subseteq> domo S1 \<union> domo S2"
+apply(induction S1 S2 rule: narrow_option.induct)
+apply (auto simp: domo_def narrow_st_def)
+done
+
+subsubsection "Termination: Commands"
+
+lemma strip_widen_acom[simp]:
+  "strip c' = strip (c::'a::WN acom) \<Longrightarrow>  strip (c \<nabla>\<^sub>c c') = strip c"
+by(induction "widen::'a\<Rightarrow>'a\<Rightarrow>'a" c c' rule: map2_acom.induct) simp_all
+
+lemma strip_narrow_acom[simp]:
+  "strip c' = strip (c::'a::WN acom) \<Longrightarrow>  strip (c \<triangle>\<^sub>c c') = strip c"
+by(induction "narrow::'a\<Rightarrow>'a\<Rightarrow>'a" c c' rule: map2_acom.induct) simp_all
+
+lemma annos_widen_acom[simp]: "strip c1 = strip (c2::'a::WN acom) \<Longrightarrow>
+  annos(c1 \<nabla>\<^sub>c c2) = map (%(x,y).x\<nabla>y) (zip (annos c1) (annos(c2::'a::WN acom)))"
+by(induction "widen::'a\<Rightarrow>'a\<Rightarrow>'a" c1 c2 rule: map2_acom.induct)
+  (simp_all add: size_annos_same2)
+
+lemma annos_narrow_acom[simp]: "strip c1 = strip (c2::'a::WN acom) \<Longrightarrow>
+  annos(c1 \<triangle>\<^sub>c c2) = map (%(x,y).x\<triangle>y) (zip (annos c1) (annos(c2::'a::WN acom)))"
+by(induction "narrow::'a\<Rightarrow>'a\<Rightarrow>'a" c1 c2 rule: map2_acom.induct)
+  (simp_all add: size_annos_same2)
+
+lemma widen_acom_Com[simp]: "strip c2 = strip c1 \<Longrightarrow>
+  c1 : Com X \<Longrightarrow> c2 : Com X \<Longrightarrow> (c1 \<nabla>\<^sub>c c2) : Com X"
+apply(auto simp add: Com_def)
+apply(rename_tac S S' x)
+apply(erule in_set_zipE)
+apply(auto simp: domo_def split: option.splits)
+apply(case_tac S)
+apply(case_tac S')
+apply simp
+apply fastforce
+apply(case_tac S')
+apply fastforce
+apply (fastforce simp: widen_st_def)
+done
+
+lemma narrow_acom_Com[simp]: "strip c2 = strip c1 \<Longrightarrow>
+  c1 : Com X \<Longrightarrow> c2 : Com X \<Longrightarrow> (c1 \<triangle>\<^sub>c c2) : Com X"
+apply(auto simp add: Com_def)
+apply(rename_tac S S' x)
+apply(erule in_set_zipE)
+apply(auto simp: domo_def split: option.splits)
+apply(case_tac S)
+apply(case_tac S')
+apply simp
+apply fastforce
+apply(case_tac S')
+apply fastforce
+apply (fastforce simp: narrow_st_def)
+done
+
+definition "m_c m c = (let as = annos c in \<Sum>i=0..<size as. m(as!i))"
+
+lemma measure_m_c: "finite X \<Longrightarrow> {(c, c \<nabla>\<^sub>c c') |c c'\<Colon>ivl st option acom.
+     strip c' = strip c \<and> c : Com X \<and> c' : Com X \<and> \<not> c' \<sqsubseteq> c}\<inverse>
+    \<subseteq> measure(m_c(m_o (m_st m_ivl) (2*card(X))))"
+apply(auto simp: m_c_def Let_def Com_def)
+apply(subgoal_tac "length(annos c') = length(annos c)")
+prefer 2 apply (simp add: size_annos_same2)
+apply (auto)
+apply(rule setsum_strict_mono1)
+apply simp
+apply (clarsimp)
+apply(erule m_o_anti_mono)
+apply(rule subset_trans[OF domo_widen_subset])
+apply fastforce
+apply(rule widen1)
+apply(auto simp: le_iff_le_annos listrel_iff_nth)
+apply(rule_tac x=n in bexI)
+prefer 2 apply simp
+apply(erule m_o_widen)
+apply (simp)+
+done
+
+lemma measure_n_c: "finite X \<Longrightarrow> {(c, c \<triangle>\<^sub>c c') |c c'.
+  strip c = strip c' \<and> c \<in> Com X \<and> c' \<in> Com X \<and> c' \<sqsubseteq> c \<and> \<not> c \<sqsubseteq> c \<triangle>\<^sub>c c'}\<inverse>
+  \<subseteq> measure(m_c(n_o (n_st n_ivl X)))"
+apply(auto simp: m_c_def Let_def Com_def)
+apply(subgoal_tac "length(annos c') = length(annos c)")
+prefer 2 apply (simp add: size_annos_same2)
+apply (auto)
+apply(rule setsum_strict_mono1)
+apply simp
+apply (clarsimp)
+apply(rule n_o_mono)
+using domo_narrow_subset apply fastforce
+apply fastforce
+apply(rule narrow2)
+apply(fastforce simp: le_iff_le_annos listrel_iff_nth)
+apply(auto simp: le_iff_le_annos listrel_iff_nth strip_narrow_acom)
+apply(rule_tac x=n in bexI)
+prefer 2 apply simp
+apply(erule n_o_narrow)
+apply (simp)+
+done
+
+
+subsubsection "Termination: Post-Fixed Point Iterations"
+
+lemma iter_widen_termination:
+fixes c0 :: "'a::WN acom"
+assumes P_f: "\<And>c. P c \<Longrightarrow> P(f c)"
+assumes P_widen: "\<And>c c'. P c \<Longrightarrow> P c' \<Longrightarrow> P(c \<nabla>\<^sub>c c')"
+and "wf({(c::'a acom,c \<nabla>\<^sub>c c')|c c'. P c \<and> P c' \<and> ~ c' \<sqsubseteq> c}^-1)"
+and "P c0" and "c0 \<sqsubseteq> f c0" shows "EX c. iter_widen f c0 = Some c"
+proof(simp add: iter_widen_def, rule wf_while_option_Some[where P = "P"])
+  show "wf {(cc', c). (P c \<and> \<not> f c \<sqsubseteq> c) \<and> cc' = c \<nabla>\<^sub>c f c}"
+    apply(rule wf_subset[OF assms(3)]) by(blast intro: P_f)
+next
+  show "P c0" by(rule `P c0`)
+next
+  fix c assume "P c" thus "P (c \<nabla>\<^sub>c f c)" by(simp add: P_f P_widen)
+qed
+
+lemma iter_narrow_termination:
+assumes P_f: "\<And>c. P c \<Longrightarrow> P(c \<triangle>\<^sub>c f c)"
+and wf: "wf({(c, c \<triangle>\<^sub>c f c)|c c'. P c \<and> ~ c \<sqsubseteq> c \<triangle>\<^sub>c f c}^-1)"
+and "P c0" shows "EX c. iter_narrow f c0 = Some c"
+proof(simp add: iter_narrow_def, rule wf_while_option_Some[where P = "P"])
+  show "wf {(c', c). (P c \<and> \<not> c \<sqsubseteq> c \<triangle>\<^sub>c f c) \<and> c' = c \<triangle>\<^sub>c f c}"
+    apply(rule wf_subset[OF wf]) by(blast intro: P_f)
+next
+  show "P c0" by(rule `P c0`)
+next
+  fix c assume "P c" thus "P (c \<triangle>\<^sub>c f c)" by(simp add: P_f)
+qed
+
+lemma iter_winden_step_ivl_termination:
+  "EX c. iter_widen (step_ivl \<top>) (\<bottom>\<^sub>c c0) = Some c"
+apply(rule iter_widen_termination[where
+  P = "%c. strip c = c0 \<and> c : Com(vars c0)"])
+apply (simp_all add: step'_Com bot_acom)
+apply(rule wf_subset)
+apply(rule wf_measure)
+apply(rule subset_trans)
+prefer 2
+apply(rule measure_m_c[where X = "vars c0", OF finite_cvars])
+apply blast
+done
+
+lemma iter_narrow_step_ivl_termination:
+  "c0 \<in> Com (vars(strip c0)) \<Longrightarrow> step_ivl \<top> c0 \<sqsubseteq> c0 \<Longrightarrow>
+  EX c. iter_narrow (step_ivl \<top>) c0 = Some c"
+apply(rule iter_narrow_termination[where
+  P = "%c. strip c = strip c0 \<and> c : Com(vars(strip c0)) \<and> step_ivl \<top> c \<sqsubseteq> c"])
+apply (simp_all add: step'_Com)
+apply(clarify)
+apply(frule narrow2_acom, drule mono_step'[OF le_refl], erule le_trans[OF _ narrow1_acom])
+apply assumption
+apply(rule wf_subset)
+apply(rule wf_measure)
+apply(rule subset_trans)
+prefer 2
+apply(rule measure_n_c[where X = "vars(strip c0)", OF finite_cvars])
+apply auto
+by (metis bot_least domo_Top order_refl step'_Com strip_step')
+
+(* FIXME: simplify type system: Combine Com(X) and vars <= X?? *)
+lemma while_Com:
+fixes c :: "'a st option acom"
+assumes "while_option P f c = Some c'"
+and "!!c. strip(f c) = strip c"
+and "\<forall>c::'a st option acom. c : Com(X) \<longrightarrow> vars(strip c) \<subseteq> X \<longrightarrow> f c : Com(X)"
+and "c : Com(X)" and "vars(strip c) \<subseteq> X" shows "c' : Com(X)"
+using while_option_rule[where P = "\<lambda>c'. c' : Com(X) \<and> vars(strip c') \<subseteq> X", OF _ assms(1)]
+by(simp add: assms(2-))
+
+lemma iter_widen_Com: fixes f :: "'a::WN st option acom \<Rightarrow> 'a st option acom"
+assumes "iter_widen f c = Some c'"
+and "\<forall>c. c : Com(X) \<longrightarrow> vars(strip c) \<subseteq> X \<longrightarrow> f c : Com(X)"
+and "!!c. strip(f c) = strip c"
+and "c : Com(X)" and "vars (strip c) \<subseteq> X" shows "c' : Com(X)"
+proof-
+  have "\<forall>c. c : Com(X) \<longrightarrow> vars(strip c) \<subseteq> X \<longrightarrow> c \<nabla>\<^sub>c f c : Com(X)"
+    by (metis (full_types) widen_acom_Com assms(2,3))
+  from while_Com[OF assms(1)[simplified iter_widen_def] _ this assms(4,5)]
+  show ?thesis using assms(3) by(simp)
+qed
+
+
+context Abs_Int2
+begin
+
+lemma iter_widen_step'_Com:
+  "iter_widen (step' \<top>) c = Some c' \<Longrightarrow> vars(strip c) \<subseteq> X \<Longrightarrow> c : Com(X)
+   \<Longrightarrow> c' : Com(X)"
+apply(subgoal_tac "strip c'= strip c")
+prefer 2 apply (metis strip_iter_widen strip_step')
+apply(drule iter_widen_Com)
+prefer 3 apply assumption
+prefer 3 apply assumption
+apply (auto simp: step'_Com)
+done
+
+end
+
+theorem AI_ivl'_termination:
+  "EX c'. AI_ivl' c = Some c'"
+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 (metis bot_acom_Com iter_widen_step'_Com[OF _ subset_refl] strip_iter_widen strip_step')
+apply(erule iter_widen_pfp)
+done
+
+end
+
+
+(* interesting(?) relic
+lemma widen_assoc:
+  "~ (y::ivl) \<sqsubseteq> x \<Longrightarrow> ~ z \<sqsubseteq> x \<nabla> y \<Longrightarrow> ((x::ivl) \<nabla> y) \<nabla> z = x \<nabla> (y \<nabla> z)"
+apply(cases x)
+apply(cases y)
+apply(cases z)
+apply(rename_tac x1 x2 y1 y2 z1 z2)
+apply(simp add: le_ivl_def)
+apply(case_tac x1)
+apply(case_tac x2)
+apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits)
+apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits)
+apply(case_tac x2)
+apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits)
+apply(case_tac y1)
+apply(case_tac y2)
+apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits)
+apply(case_tac z1)
+apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1]
+apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1]
+apply(case_tac y2)
+apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1]
+apply(case_tac z1)
+apply(auto simp add:le_option_def widen_ivl_def split: if_splits ivl.splits option.splits)[1]
+apply(case_tac z2)
+apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits)[1]
+apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits)[1]
+done
+
+lemma widen_step_trans:
+  "~ (y::ivl) \<sqsubseteq> x \<Longrightarrow> ~ z \<sqsubseteq> x \<nabla> y \<Longrightarrow> EX u. (x \<nabla> y) \<nabla> z = x \<nabla> u \<and> ~ u \<sqsubseteq> x"
+by (metis widen_assoc preord_class.le_trans widen1)
+*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_State_ITP.thy	Thu Apr 19 17:32:30 2012 +0200
@@ -0,0 +1,98 @@
+(* Author: Tobias Nipkow *)
+
+theory Abs_State_ITP
+imports Abs_Int0_ITP
+  "~~/src/HOL/Library/Char_ord" "~~/src/HOL/Library/List_lexord"
+  (* Library import merely to allow string lists to be sorted for output *)
+begin
+
+subsection "Abstract State with Computable Ordering"
+
+text{* A concrete type of state with computable @{text"\<sqsubseteq>"}: *}
+
+datatype 'a st = FunDom "vname \<Rightarrow> 'a" "vname list"
+
+fun "fun" where "fun (FunDom f xs) = f"
+fun dom where "dom (FunDom f xs) = xs"
+
+definition [simp]: "inter_list xs ys = [x\<leftarrow>xs. x \<in> set ys]"
+
+definition "show_st S = [(x,fun S x). x \<leftarrow> sort(dom S)]"
+
+definition "show_acom = map_acom (Option.map show_st)"
+definition "show_acom_opt = Option.map show_acom"
+
+definition "lookup F x = (if x : set(dom F) then fun F x else \<top>)"
+
+definition "update F x y =
+  FunDom ((fun F)(x:=y)) (if x \<in> set(dom F) then dom F else x # dom F)"
+
+lemma lookup_update: "lookup (update S x y) = (lookup S)(x:=y)"
+by(rule ext)(auto simp: lookup_def update_def)
+
+definition "\<gamma>_st \<gamma> F = {f. \<forall>x. f x \<in> \<gamma>(lookup F x)}"
+
+instantiation st :: (SL_top) SL_top
+begin
+
+definition "le_st F G = (ALL x : set(dom G). lookup F x \<sqsubseteq> fun G x)"
+
+definition
+"join_st F G =
+ FunDom (\<lambda>x. fun F x \<squnion> fun G x) (inter_list (dom F) (dom G))"
+
+definition "\<top> = FunDom (\<lambda>x. \<top>) []"
+
+instance
+proof
+  case goal2 thus ?case
+    apply(auto simp: le_st_def)
+    by (metis lookup_def preord_class.le_trans top)
+qed (auto simp: le_st_def lookup_def join_st_def Top_st_def)
+
+end
+
+lemma mono_lookup: "F \<sqsubseteq> F' \<Longrightarrow> lookup F x \<sqsubseteq> lookup F' x"
+by(auto simp add: lookup_def le_st_def)
+
+lemma mono_update: "a \<sqsubseteq> a' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> update S x a \<sqsubseteq> update S' x a'"
+by(auto simp add: le_st_def lookup_def update_def)
+
+locale Gamma = Val_abs where \<gamma>=\<gamma> for \<gamma> :: "'av::SL_top \<Rightarrow> val set"
+begin
+
+abbreviation \<gamma>\<^isub>f :: "'av st \<Rightarrow> state set"
+where "\<gamma>\<^isub>f == \<gamma>_st \<gamma>"
+
+abbreviation \<gamma>\<^isub>o :: "'av st option \<Rightarrow> state set"
+where "\<gamma>\<^isub>o == \<gamma>_option \<gamma>\<^isub>f"
+
+abbreviation \<gamma>\<^isub>c :: "'av st option acom \<Rightarrow> state set acom"
+where "\<gamma>\<^isub>c == map_acom \<gamma>\<^isub>o"
+
+lemma gamma_f_Top[simp]: "\<gamma>\<^isub>f Top = UNIV"
+by(auto simp: Top_st_def \<gamma>_st_def lookup_def)
+
+lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o Top = UNIV"
+by (simp add: Top_option_def)
+
+(* FIXME (maybe also le \<rightarrow> sqle?) *)
+
+lemma mono_gamma_f: "f \<sqsubseteq> g \<Longrightarrow> \<gamma>\<^isub>f f \<subseteq> \<gamma>\<^isub>f g"
+apply(simp add:\<gamma>_st_def subset_iff lookup_def le_st_def split: if_splits)
+by (metis UNIV_I mono_gamma gamma_Top subsetD)
+
+lemma mono_gamma_o:
+  "sa \<sqsubseteq> sa' \<Longrightarrow> \<gamma>\<^isub>o sa \<subseteq> \<gamma>\<^isub>o sa'"
+by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f)
+
+lemma mono_gamma_c: "ca \<sqsubseteq> ca' \<Longrightarrow> \<gamma>\<^isub>c ca \<le> \<gamma>\<^isub>c ca'"
+by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o)
+
+lemma in_gamma_option_iff:
+  "x : \<gamma>_option r u \<longleftrightarrow> (\<exists>u'. u = Some u' \<and> x : r u')"
+by (cases u) auto
+
+end
+
+end
--- a/src/HOL/IMP/Abs_State.thy	Thu Apr 19 12:28:10 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,98 +0,0 @@
-(* Author: Tobias Nipkow *)
-
-theory Abs_State
-imports Abs_Int0_fun
-  "~~/src/HOL/Library/Char_ord" "~~/src/HOL/Library/List_lexord"
-  (* Library import merely to allow string lists to be sorted for output *)
-begin
-
-subsection "Abstract State with Computable Ordering"
-
-text{* A concrete type of state with computable @{text"\<sqsubseteq>"}: *}
-
-datatype 'a st = FunDom "vname \<Rightarrow> 'a" "vname list"
-
-fun "fun" where "fun (FunDom f xs) = f"
-fun dom where "dom (FunDom f xs) = xs"
-
-definition [simp]: "inter_list xs ys = [x\<leftarrow>xs. x \<in> set ys]"
-
-definition "show_st S = [(x,fun S x). x \<leftarrow> sort(dom S)]"
-
-definition "show_acom = map_acom (Option.map show_st)"
-definition "show_acom_opt = Option.map show_acom"
-
-definition "lookup F x = (if x : set(dom F) then fun F x else \<top>)"
-
-definition "update F x y =
-  FunDom ((fun F)(x:=y)) (if x \<in> set(dom F) then dom F else x # dom F)"
-
-lemma lookup_update: "lookup (update S x y) = (lookup S)(x:=y)"
-by(rule ext)(auto simp: lookup_def update_def)
-
-definition "\<gamma>_st \<gamma> F = {f. \<forall>x. f x \<in> \<gamma>(lookup F x)}"
-
-instantiation st :: (SL_top) SL_top
-begin
-
-definition "le_st F G = (ALL x : set(dom G). lookup F x \<sqsubseteq> fun G x)"
-
-definition
-"join_st F G =
- FunDom (\<lambda>x. fun F x \<squnion> fun G x) (inter_list (dom F) (dom G))"
-
-definition "\<top> = FunDom (\<lambda>x. \<top>) []"
-
-instance
-proof
-  case goal2 thus ?case
-    apply(auto simp: le_st_def)
-    by (metis lookup_def preord_class.le_trans top)
-qed (auto simp: le_st_def lookup_def join_st_def Top_st_def)
-
-end
-
-lemma mono_lookup: "F \<sqsubseteq> F' \<Longrightarrow> lookup F x \<sqsubseteq> lookup F' x"
-by(auto simp add: lookup_def le_st_def)
-
-lemma mono_update: "a \<sqsubseteq> a' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> update S x a \<sqsubseteq> update S' x a'"
-by(auto simp add: le_st_def lookup_def update_def)
-
-locale Gamma = Val_abs where \<gamma>=\<gamma> for \<gamma> :: "'av::SL_top \<Rightarrow> val set"
-begin
-
-abbreviation \<gamma>\<^isub>f :: "'av st \<Rightarrow> state set"
-where "\<gamma>\<^isub>f == \<gamma>_st \<gamma>"
-
-abbreviation \<gamma>\<^isub>o :: "'av st option \<Rightarrow> state set"
-where "\<gamma>\<^isub>o == \<gamma>_option \<gamma>\<^isub>f"
-
-abbreviation \<gamma>\<^isub>c :: "'av st option acom \<Rightarrow> state set acom"
-where "\<gamma>\<^isub>c == map_acom \<gamma>\<^isub>o"
-
-lemma gamma_f_Top[simp]: "\<gamma>\<^isub>f Top = UNIV"
-by(auto simp: Top_st_def \<gamma>_st_def lookup_def)
-
-lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o Top = UNIV"
-by (simp add: Top_option_def)
-
-(* FIXME (maybe also le \<rightarrow> sqle?) *)
-
-lemma mono_gamma_f: "f \<sqsubseteq> g \<Longrightarrow> \<gamma>\<^isub>f f \<subseteq> \<gamma>\<^isub>f g"
-apply(simp add:\<gamma>_st_def subset_iff lookup_def le_st_def split: if_splits)
-by (metis UNIV_I mono_gamma gamma_Top subsetD)
-
-lemma mono_gamma_o:
-  "sa \<sqsubseteq> sa' \<Longrightarrow> \<gamma>\<^isub>o sa \<subseteq> \<gamma>\<^isub>o sa'"
-by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f)
-
-lemma mono_gamma_c: "ca \<sqsubseteq> ca' \<Longrightarrow> \<gamma>\<^isub>c ca \<le> \<gamma>\<^isub>c ca'"
-by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o)
-
-lemma in_gamma_option_iff:
-  "x : \<gamma>_option r u \<longleftrightarrow> (\<exists>u'. u = Some u' \<and> x : r u')"
-by (cases u) auto
-
-end
-
-end
--- a/src/HOL/IMP/ROOT.ML	Thu Apr 19 12:28:10 2012 +0200
+++ b/src/HOL/IMP/ROOT.ML	Thu Apr 19 17:32:30 2012 +0200
@@ -1,5 +1,6 @@
 no_document use_thys
   ["~~/src/HOL/ex/Interpretation_with_Defs",
+   "~~/src/HOL/Library/While_Combinator",
    "~~/src/HOL/Library/Char_ord", "~~/src/HOL/Library/List_lexord"
   ];
 
@@ -21,10 +22,10 @@
  "HoareT",
  "Collecting1",
  "Collecting_list",
- "Abs_Int0",
- "Abs_Int0_parity",
- "Abs_Int0_const",
- "Abs_Int2",
+ "Abs_Int_Tests",
+ "Abs_Int_ITP/Abs_Int1_parity_ITP",
+ "Abs_Int_ITP/Abs_Int1_const_ITP",
+ "Abs_Int_ITP/Abs_Int3_ITP",
  "Abs_Int_Den/Abs_Int_den2",
  "Procs_Dyn_Vars_Dyn",
  "Procs_Stat_Vars_Dyn",
--- a/src/HOL/IMP/document/root.tex	Thu Apr 19 12:28:10 2012 +0200
+++ b/src/HOL/IMP/document/root.tex	Thu Apr 19 17:32:30 2012 +0200
@@ -54,6 +54,7 @@
 \author{TN \& GK}
 \maketitle
 
+\setcounter{tocdepth}{2}
 \tableofcontents
 \newpage
 
--- a/src/HOL/IsaMakefile	Thu Apr 19 12:28:10 2012 +0200
+++ b/src/HOL/IsaMakefile	Thu Apr 19 17:32:30 2012 +0200
@@ -525,9 +525,12 @@
 HOL-IMP: HOL $(OUT)/HOL-IMP
 
 $(OUT)/HOL-IMP: $(OUT)/HOL \
-  IMP/ACom.thy IMP/Abs_Int0_fun.thy IMP/Abs_State.thy IMP/Abs_Int0.thy \
-  IMP/Abs_Int0_const.thy IMP/Abs_Int1.thy IMP/Abs_Int1_ivl.thy \
-  IMP/Abs_Int2.thy IMP/Abs_Int_Den/Abs_Int_den0_fun.thy \
+  IMP/ACom.thy \
+  IMP/Abs_Int_ITP/Abs_Int0_ITP.thy IMP/Abs_Int_ITP/Abs_State_ITP.thy \
+  IMP/Abs_Int_ITP/Abs_Int1_ITP.thy IMP/Abs_Int_ITP/Abs_Int1_const_ITP.thy \
+  IMP/Abs_Int_ITP/Abs_Int1_parity_ITP.thy IMP/Abs_Int_ITP/Abs_Int2_ITP.thy \
+  IMP/Abs_Int_ITP/Abs_Int2_ivl_ITP.thy IMP/Abs_Int_ITP/Abs_Int3_ITP.thy \
+  IMP/Abs_Int_Den/Abs_Int_den0_fun.thy \
   IMP/Abs_Int_Den/Abs_State_den.thy  IMP/Abs_Int_Den/Abs_Int_den0.thy \
   IMP/Abs_Int_Den/Abs_Int_den0_const.thy IMP/Abs_Int_Den/Abs_Int_den1.thy \
   IMP/Abs_Int_Den/Abs_Int_den1_ivl.thy IMP/Abs_Int_Den/Abs_Int_den2.thy \