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