tune simpset for Complete_Lattices
authornoschinl
Tue, 13 Sep 2011 16:21:48 +0200
changeset 44918 6a80fbc4e72c
parent 44917 8df4c332cb1c
child 44919 482f1807976e
tune simpset for Complete_Lattices
src/HOL/Big_Operators.thy
src/HOL/Complete_Lattices.thy
src/HOL/Induct/Sexp.thy
src/HOL/Lattices.thy
src/HOL/Library/Extended_Real.thy
src/HOL/Library/Kleene_Algebra.thy
src/HOL/Main.thy
src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy
src/HOL/Probability/Caratheodory.thy
src/HOL/Probability/Radon_Nikodym.thy
src/HOL/UNITY/ProgressSets.thy
--- a/src/HOL/Big_Operators.thy	Tue Sep 13 13:17:52 2011 +0200
+++ b/src/HOL/Big_Operators.thy	Tue Sep 13 16:21:48 2011 +0200
@@ -1433,11 +1433,10 @@
 proof -
   interpret ab_semigroup_idem_mult inf
     by (rule ab_semigroup_idem_mult_inf)
-  from `A \<noteq> {}` obtain b B where "A = insert b B" by auto
+  from `A \<noteq> {}` obtain b B where "A = {b} \<union> B" by auto
   moreover with `finite A` have "finite B" by simp
-  ultimately show ?thesis  
-  by (simp add: Inf_fin_def fold1_eq_fold_idem inf_Inf_fold_inf [symmetric])
-    (simp add: Inf_fold_inf)
+  ultimately show ?thesis
+    by (simp add: Inf_fin_def fold1_eq_fold_idem inf_Inf_fold_inf [symmetric])
 qed
 
 lemma Sup_fin_Sup:
@@ -1446,11 +1445,10 @@
 proof -
   interpret ab_semigroup_idem_mult sup
     by (rule ab_semigroup_idem_mult_sup)
-  from `A \<noteq> {}` obtain b B where "A = insert b B" by auto
+  from `A \<noteq> {}` obtain b B where "A = {b} \<union> B" by auto
   moreover with `finite A` have "finite B" by simp
   ultimately show ?thesis  
   by (simp add: Sup_fin_def fold1_eq_fold_idem sup_Sup_fold_sup [symmetric])
-    (simp add: Sup_fold_sup)
 qed
 
 end
--- a/src/HOL/Complete_Lattices.thy	Tue Sep 13 13:17:52 2011 +0200
+++ b/src/HOL/Complete_Lattices.thy	Tue Sep 13 16:21:48 2011 +0200
@@ -126,16 +126,16 @@
 lemma SUP_upper2: "i \<in> A \<Longrightarrow> u \<sqsubseteq> f i \<Longrightarrow> u \<sqsubseteq> (\<Squnion>i\<in>A. f i)"
   using SUP_upper [of i A f] by auto
 
-lemma le_Inf_iff (*[simp]*): "b \<sqsubseteq> \<Sqinter>A \<longleftrightarrow> (\<forall>a\<in>A. b \<sqsubseteq> a)"
+lemma le_Inf_iff: "b \<sqsubseteq> \<Sqinter>A \<longleftrightarrow> (\<forall>a\<in>A. b \<sqsubseteq> a)"
   by (auto intro: Inf_greatest dest: Inf_lower)
 
-lemma le_INF_iff (*[simp]*): "u \<sqsubseteq> (\<Sqinter>i\<in>A. f i) \<longleftrightarrow> (\<forall>i\<in>A. u \<sqsubseteq> f i)"
+lemma le_INF_iff: "u \<sqsubseteq> (\<Sqinter>i\<in>A. f i) \<longleftrightarrow> (\<forall>i\<in>A. u \<sqsubseteq> f i)"
   by (auto simp add: INF_def le_Inf_iff)
 
-lemma Sup_le_iff (*[simp]*): "\<Squnion>A \<sqsubseteq> b \<longleftrightarrow> (\<forall>a\<in>A. a \<sqsubseteq> b)"
+lemma Sup_le_iff: "\<Squnion>A \<sqsubseteq> b \<longleftrightarrow> (\<forall>a\<in>A. a \<sqsubseteq> b)"
   by (auto intro: Sup_least dest: Sup_upper)
 
-lemma SUP_le_iff (*[simp]*): "(\<Squnion>i\<in>A. f i) \<sqsubseteq> u \<longleftrightarrow> (\<forall>i\<in>A. f i \<sqsubseteq> u)"
+lemma SUP_le_iff: "(\<Squnion>i\<in>A. f i) \<sqsubseteq> u \<longleftrightarrow> (\<forall>i\<in>A. f i \<sqsubseteq> u)"
   by (auto simp add: SUP_def Sup_le_iff)
 
 lemma Inf_empty [simp]:
@@ -160,22 +160,22 @@
   "\<Squnion>UNIV = \<top>"
   by (auto intro!: antisym Sup_upper)
 
-lemma Inf_insert (*[simp]*): "\<Sqinter>insert a A = a \<sqinter> \<Sqinter>A"
+lemma Inf_insert [simp]: "\<Sqinter>insert a A = a \<sqinter> \<Sqinter>A"
   by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower)
 
 lemma INF_insert: "(\<Sqinter>x\<in>insert a A. f x) = f a \<sqinter> INFI A f"
   by (simp add: INF_def Inf_insert)
 
-lemma Sup_insert (*[simp]*): "\<Squnion>insert a A = a \<squnion> \<Squnion>A"
+lemma Sup_insert [simp]: "\<Squnion>insert a A = a \<squnion> \<Squnion>A"
   by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper)
 
 lemma SUP_insert: "(\<Squnion>x\<in>insert a A. f x) = f a \<squnion> SUPR A f"
   by (simp add: SUP_def Sup_insert)
 
-lemma INF_image (*[simp]*): "(\<Sqinter>x\<in>f`A. g x) = (\<Sqinter>x\<in>A. g (f x))"
+lemma INF_image [simp]: "(\<Sqinter>x\<in>f`A. g x) = (\<Sqinter>x\<in>A. g (f x))"
   by (simp add: INF_def image_image)
 
-lemma SUP_image (*[simp]*): "(\<Squnion>x\<in>f`A. g x) = (\<Squnion>x\<in>A. g (f x))"
+lemma SUP_image [simp]: "(\<Squnion>x\<in>f`A. g x) = (\<Squnion>x\<in>A. g (f x))"
   by (simp add: SUP_def image_image)
 
 lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<sqsubseteq> a}"
@@ -210,7 +210,7 @@
 
 lemma INF_mono:
   "(\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<sqsubseteq> g m) \<Longrightarrow> (\<Sqinter>n\<in>A. f n) \<sqsubseteq> (\<Sqinter>n\<in>B. g n)"
-  by (force intro!: Inf_mono simp: INF_def)
+  unfolding INF_def by (rule Inf_mono) fast
 
 lemma Sup_mono:
   assumes "\<And>a. a \<in> A \<Longrightarrow> \<exists>b\<in>B. a \<sqsubseteq> b"
@@ -224,7 +224,7 @@
 
 lemma SUP_mono:
   "(\<And>n. n \<in> A \<Longrightarrow> \<exists>m\<in>B. f n \<sqsubseteq> g m) \<Longrightarrow> (\<Squnion>n\<in>A. f n) \<sqsubseteq> (\<Squnion>n\<in>B. g n)"
-  by (force intro!: Sup_mono simp: SUP_def)
+  unfolding SUP_def by (rule Sup_mono) fast
 
 lemma INF_superset_mono:
   "B \<subseteq> A \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<sqsubseteq> g x) \<Longrightarrow> (\<Sqinter>x\<in>A. f x) \<sqsubseteq> (\<Sqinter>x\<in>B. g x)"
@@ -278,11 +278,14 @@
 lemma INF_inf_distrib: "(\<Sqinter>a\<in>A. f a) \<sqinter> (\<Sqinter>a\<in>A. g a) = (\<Sqinter>a\<in>A. f a \<sqinter> g a)"
   by (rule antisym) (rule INF_greatest, auto intro: le_infI1 le_infI2 INF_lower INF_mono)
 
-lemma SUP_sup_distrib: "(\<Squnion>a\<in>A. f a) \<squnion> (\<Squnion>a\<in>A. g a) = (\<Squnion>a\<in>A. f a \<squnion> g a)"
-  by (rule antisym) (auto intro: le_supI1 le_supI2 SUP_upper SUP_mono,
-    rule SUP_least, auto intro: le_supI1 le_supI2 SUP_upper SUP_mono)
+lemma SUP_sup_distrib: "(\<Squnion>a\<in>A. f a) \<squnion> (\<Squnion>a\<in>A. g a) = (\<Squnion>a\<in>A. f a \<squnion> g a)" (is "?L = ?R")
+proof (rule antisym)
+  show "?L \<le> ?R" by (auto intro: le_supI1 le_supI2 SUP_upper SUP_mono)
+next
+  show "?R \<le> ?L" by (rule SUP_least) (auto intro: le_supI1 le_supI2 SUP_upper)
+qed
 
-lemma Inf_top_conv (*[simp]*) [no_atp]:
+lemma Inf_top_conv [simp, no_atp]:
   "\<Sqinter>A = \<top> \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
   "\<top> = \<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
 proof -
@@ -304,12 +307,12 @@
   then show "\<top> = \<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)" by auto
 qed
 
-lemma INF_top_conv (*[simp]*):
+lemma INF_top_conv [simp]:
  "(\<Sqinter>x\<in>A. B x) = \<top> \<longleftrightarrow> (\<forall>x\<in>A. B x = \<top>)"
  "\<top> = (\<Sqinter>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. B x = \<top>)"
   by (auto simp add: INF_def Inf_top_conv)
 
-lemma Sup_bot_conv (*[simp]*) [no_atp]:
+lemma Sup_bot_conv [simp, no_atp]:
   "\<Squnion>A = \<bottom> \<longleftrightarrow> (\<forall>x\<in>A. x = \<bottom>)" (is ?P)
   "\<bottom> = \<Squnion>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<bottom>)" (is ?Q)
 proof -
@@ -318,7 +321,7 @@
   from dual.Inf_top_conv show ?P and ?Q by simp_all
 qed
 
-lemma SUP_bot_conv (*[simp]*):
+lemma SUP_bot_conv [simp]:
  "(\<Squnion>x\<in>A. B x) = \<bottom> \<longleftrightarrow> (\<forall>x\<in>A. B x = \<bottom>)"
  "\<bottom> = (\<Squnion>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. B x = \<bottom>)"
   by (auto simp add: SUP_def Sup_bot_conv)
@@ -329,10 +332,10 @@
 lemma SUP_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Squnion>i\<in>A. f) = f"
   by (auto intro: antisym SUP_upper SUP_least)
 
-lemma INF_top (*[simp]*): "(\<Sqinter>x\<in>A. \<top>) = \<top>"
+lemma INF_top [simp]: "(\<Sqinter>x\<in>A. \<top>) = \<top>"
   by (cases "A = {}") (simp_all add: INF_empty)
 
-lemma SUP_bot (*[simp]*): "(\<Squnion>x\<in>A. \<bottom>) = \<bottom>"
+lemma SUP_bot [simp]: "(\<Squnion>x\<in>A. \<bottom>) = \<bottom>"
   by (cases "A = {}") (simp_all add: SUP_empty)
 
 lemma INF_commute: "(\<Sqinter>i\<in>A. \<Sqinter>j\<in>B. f i j) = (\<Sqinter>j\<in>B. \<Sqinter>i\<in>A. f i j)"
@@ -492,23 +495,23 @@
   "class.complete_linorder Sup Inf sup (op \<ge>) (op >) inf \<top> \<bottom>"
   by (rule class.complete_linorder.intro, rule dual_complete_lattice, rule dual_linorder)
 
-lemma Inf_less_iff (*[simp]*):
+lemma Inf_less_iff:
   "\<Sqinter>S \<sqsubset> a \<longleftrightarrow> (\<exists>x\<in>S. x \<sqsubset> a)"
   unfolding not_le [symmetric] le_Inf_iff by auto
 
-lemma INF_less_iff (*[simp]*):
+lemma INF_less_iff:
   "(\<Sqinter>i\<in>A. f i) \<sqsubset> a \<longleftrightarrow> (\<exists>x\<in>A. f x \<sqsubset> a)"
   unfolding INF_def Inf_less_iff by auto
 
-lemma less_Sup_iff (*[simp]*):
+lemma less_Sup_iff:
   "a \<sqsubset> \<Squnion>S \<longleftrightarrow> (\<exists>x\<in>S. a \<sqsubset> x)"
   unfolding not_le [symmetric] Sup_le_iff by auto
 
-lemma less_SUP_iff (*[simp]*):
+lemma less_SUP_iff:
   "a \<sqsubset> (\<Squnion>i\<in>A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a \<sqsubset> f x)"
   unfolding SUP_def less_Sup_iff by auto
 
-lemma Sup_eq_top_iff (*[simp]*):
+lemma Sup_eq_top_iff [simp]:
   "\<Squnion>A = \<top> \<longleftrightarrow> (\<forall>x<\<top>. \<exists>i\<in>A. x < i)"
 proof
   assume *: "\<Squnion>A = \<top>"
@@ -530,11 +533,11 @@
   qed
 qed
 
-lemma SUP_eq_top_iff (*[simp]*):
+lemma SUP_eq_top_iff [simp]:
   "(\<Squnion>i\<in>A. f i) = \<top> \<longleftrightarrow> (\<forall>x<\<top>. \<exists>i\<in>A. x < f i)"
   unfolding SUP_def Sup_eq_top_iff by auto
 
-lemma Inf_eq_bot_iff (*[simp]*):
+lemma Inf_eq_bot_iff [simp]:
   "\<Sqinter>A = \<bottom> \<longleftrightarrow> (\<forall>x>\<bottom>. \<exists>i\<in>A. i < x)"
 proof -
   interpret dual: complete_linorder Sup Inf sup "op \<ge>" "op >" inf \<top> \<bottom>
@@ -542,7 +545,7 @@
   from dual.Sup_eq_top_iff show ?thesis .
 qed
 
-lemma INF_eq_bot_iff (*[simp]*):
+lemma INF_eq_bot_iff [simp]:
   "(\<Sqinter>i\<in>A. f i) = \<bottom> \<longleftrightarrow> (\<forall>x>\<bottom>. \<exists>i\<in>A. f i < x)"
   unfolding INF_def Inf_eq_bot_iff by auto
 
--- a/src/HOL/Induct/Sexp.thy	Tue Sep 13 13:17:52 2011 +0200
+++ b/src/HOL/Induct/Sexp.thy	Tue Sep 13 16:21:48 2011 +0200
@@ -73,7 +73,7 @@
 (** Introduction rules for 'pred_sexp' **)
 
 lemma pred_sexp_subset_Sigma: "pred_sexp <= sexp <*> sexp"
-by (simp add: pred_sexp_def, blast)
+  by (simp add: pred_sexp_def) blast
 
 (* (a,b) \<in> pred_sexp^+ ==> a \<in> sexp *)
 lemmas trancl_pred_sexpD1 = 
--- a/src/HOL/Lattices.thy	Tue Sep 13 13:17:52 2011 +0200
+++ b/src/HOL/Lattices.thy	Tue Sep 13 16:21:48 2011 +0200
@@ -180,10 +180,10 @@
 lemma inf_left_commute: "x \<sqinter> (y \<sqinter> z) = y \<sqinter> (x \<sqinter> z)"
   by (fact inf.left_commute)
 
-lemma inf_idem (*[simp]*): "x \<sqinter> x = x"
+lemma inf_idem [simp]: "x \<sqinter> x = x"
   by (fact inf.idem)
 
-lemma inf_left_idem (*[simp]*): "x \<sqinter> (x \<sqinter> y) = x \<sqinter> y"
+lemma inf_left_idem [simp]: "x \<sqinter> (x \<sqinter> y) = x \<sqinter> y"
   by (fact inf.left_idem)
 
 lemma inf_absorb1: "x \<sqsubseteq> y \<Longrightarrow> x \<sqinter> y = x"
@@ -219,10 +219,10 @@
 lemma sup_left_commute: "x \<squnion> (y \<squnion> z) = y \<squnion> (x \<squnion> z)"
   by (fact sup.left_commute)
 
-lemma sup_idem (*[simp]*): "x \<squnion> x = x"
+lemma sup_idem [simp]: "x \<squnion> x = x"
   by (fact sup.idem)
 
-lemma sup_left_idem (*[simp]*): "x \<squnion> (x \<squnion> y) = x \<squnion> y"
+lemma sup_left_idem [simp]: "x \<squnion> (x \<squnion> y) = x \<squnion> y"
   by (fact sup.left_idem)
 
 lemma sup_absorb1: "y \<sqsubseteq> x \<Longrightarrow> x \<squnion> y = x"
@@ -243,10 +243,10 @@
   by (rule class.lattice.intro, rule dual_semilattice, rule class.semilattice_sup.intro, rule dual_order)
     (unfold_locales, auto)
 
-lemma inf_sup_absorb (*[simp]*): "x \<sqinter> (x \<squnion> y) = x"
+lemma inf_sup_absorb [simp]: "x \<sqinter> (x \<squnion> y) = x"
   by (blast intro: antisym inf_le1 inf_greatest sup_ge1)
 
-lemma sup_inf_absorb (*[simp]*): "x \<squnion> (x \<sqinter> y) = x"
+lemma sup_inf_absorb [simp]: "x \<squnion> (x \<sqinter> y) = x"
   by (blast intro: antisym sup_ge1 sup_least inf_le1)
 
 lemmas inf_sup_aci = inf_aci sup_aci
@@ -267,8 +267,9 @@
 assumes D: "!!x y z. x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)"
 shows "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
 proof-
-  have "x \<squnion> (y \<sqinter> z) = (x \<squnion> (x \<sqinter> z)) \<squnion> (y \<sqinter> z)" by(simp add:sup_inf_absorb)
-  also have "\<dots> = x \<squnion> (z \<sqinter> (x \<squnion> y))" by(simp add:D inf_commute sup_assoc)
+  have "x \<squnion> (y \<sqinter> z) = (x \<squnion> (x \<sqinter> z)) \<squnion> (y \<sqinter> z)" by simp
+  also have "\<dots> = x \<squnion> (z \<sqinter> (x \<squnion> y))"
+    by (simp add: D inf_commute sup_assoc del: sup_inf_absorb)
   also have "\<dots> = ((x \<squnion> y) \<sqinter> x) \<squnion> ((x \<squnion> y) \<sqinter> z)"
     by(simp add:inf_sup_absorb inf_commute)
   also have "\<dots> = (x \<squnion> y) \<sqinter> (x \<squnion> z)" by(simp add:D)
@@ -279,8 +280,9 @@
 assumes D: "!!x y z. x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
 shows "x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)"
 proof-
-  have "x \<sqinter> (y \<squnion> z) = (x \<sqinter> (x \<squnion> z)) \<sqinter> (y \<squnion> z)" by(simp add:inf_sup_absorb)
-  also have "\<dots> = x \<sqinter> (z \<squnion> (x \<sqinter> y))" by(simp add:D sup_commute inf_assoc)
+  have "x \<sqinter> (y \<squnion> z) = (x \<sqinter> (x \<squnion> z)) \<sqinter> (y \<squnion> z)" by simp
+  also have "\<dots> = x \<sqinter> (z \<squnion> (x \<sqinter> y))"
+    by (simp add: D sup_commute inf_assoc del: inf_sup_absorb)
   also have "\<dots> = ((x \<sqinter> y) \<squnion> x) \<sqinter> ((x \<sqinter> y) \<squnion> z)"
     by(simp add:sup_inf_absorb sup_commute)
   also have "\<dots> = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" by(simp add:D)
@@ -439,11 +441,11 @@
   by (rule class.boolean_algebra.intro, rule dual_bounded_lattice, rule dual_distrib_lattice)
     (unfold_locales, auto simp add: inf_compl_bot sup_compl_top diff_eq)
 
-lemma compl_inf_bot (*[simp]*):
+lemma compl_inf_bot [simp]:
   "- x \<sqinter> x = \<bottom>"
   by (simp add: inf_commute inf_compl_bot)
 
-lemma compl_sup_top (*[simp]*):
+lemma compl_sup_top [simp]:
   "- x \<squnion> x = \<top>"
   by (simp add: sup_commute sup_compl_top)
 
@@ -525,7 +527,7 @@
   then show "- y \<sqsubseteq> - x" by (simp only: le_iff_inf)
 qed
 
-lemma compl_le_compl_iff (*[simp]*):
+lemma compl_le_compl_iff [simp]:
   "- x \<sqsubseteq> - y \<longleftrightarrow> y \<sqsubseteq> x"
   by (auto dest: compl_mono)
 
--- a/src/HOL/Library/Extended_Real.thy	Tue Sep 13 13:17:52 2011 +0200
+++ b/src/HOL/Library/Extended_Real.thy	Tue Sep 13 16:21:48 2011 +0200
@@ -1506,7 +1506,8 @@
       proof cases
         assume "\<forall>i. f i = 0"
         moreover then have "range f = {0}" by auto
-        ultimately show "c * SUPR UNIV f \<le> y" using * by (auto simp: SUPR_def)
+        ultimately show "c * SUPR UNIV f \<le> y" using *
+          by (auto simp: SUPR_def min_max.sup_absorb1)
       next
         assume "\<not> (\<forall>i. f i = 0)"
         then obtain i where "f i \<noteq> 0" by auto
@@ -1568,7 +1569,8 @@
         hence "0 < r" using `0 < e` by auto
         then obtain n ::nat where *: "1 / real n < r" "0 < n"
           using ex_inverse_of_nat_less by (auto simp: real_eq_of_nat inverse_eq_divide)
-        have "Sup A \<le> f n + 1 / ereal (real n)" using f[THEN spec, of n] by auto
+        have "Sup A \<le> f n + 1 / ereal (real n)" using f[THEN spec, of n]
+          by auto
         also have "1 / ereal (real n) \<le> e" using real * by (auto simp: one_ereal_def )
         with bound have "f n + 1 / ereal (real n) \<le> y + e" by (rule add_mono) simp
         finally show "Sup A \<le> y + e" .
@@ -1625,7 +1627,7 @@
   then show "Sup ((\<lambda>x. a + x) ` A) \<le> a + Sup A" .
   show "a + Sup A \<le> Sup ((\<lambda>x. a + x) ` A)"
   proof (cases a)
-    case PInf with `A \<noteq> {}` show ?thesis by (auto simp: image_constant)
+    case PInf with `A \<noteq> {}` show ?thesis by (auto simp: image_constant min_max.sup_absorb1)
   next
     case (real r)
     then have **: "op + (- a) ` op + a ` A = A"
--- a/src/HOL/Library/Kleene_Algebra.thy	Tue Sep 13 13:17:52 2011 +0200
+++ b/src/HOL/Library/Kleene_Algebra.thy	Tue Sep 13 16:21:48 2011 +0200
@@ -377,19 +377,18 @@
   have [simp]: "1 \<le> star a"
     unfolding star_cont[of 1 a 1, simplified] 
     by (subst power_0[symmetric]) (rule le_SUPI [OF UNIV_I])
-  
-  show "1 + a * star a \<le> star a"
-    apply (rule plus_leI, simp)
-    apply (simp add:star_cont[of a a 1, simplified])
-    apply (simp add:star_cont[of 1 a 1, simplified])
-    apply (subst power_Suc[symmetric])
-    by (intro SUP_leI le_SUPI UNIV_I)
+
+  have "a * star a \<le> star a"
+    using star_cont[of a a 1] star_cont[of 1 a 1]
+    by (auto simp add: power_Suc[symmetric] simp del: power_Suc
+      intro: SUP_leI le_SUPI)
 
-  show "1 + star a * a \<le> star a" 
-    apply (rule plus_leI, simp)
-    apply (simp add:star_cont[of 1 a a, simplified])
-    apply (simp add:star_cont[of 1 a 1, simplified])
-    by (auto intro: SUP_leI le_SUPI simp add: power_Suc[symmetric] power_commutes simp del: power_Suc)
+  then show "1 + a * star a \<le> star a"
+    by simp
+
+  then show "1 + star a * a \<le> star a"
+    using star_cont[of a a 1] star_cont[of 1 a a]
+    by (simp add: power_commutes)
 
   show "a * x \<le> x \<Longrightarrow> star a * x \<le> x"
   proof -
--- a/src/HOL/Main.thy	Tue Sep 13 13:17:52 2011 +0200
+++ b/src/HOL/Main.thy	Tue Sep 13 16:21:48 2011 +0200
@@ -11,17 +11,4 @@
 
 text {* See further \cite{Nipkow-et-al:2002:tutorial} *}
 
-text {* Compatibility layer -- to be dropped *}
-
-lemma Inf_bool_def:
-  "Inf A \<longleftrightarrow> (\<forall>x\<in>A. x)"
-  by (auto intro: bool_induct)
-
-lemma Sup_bool_def:
-  "Sup A \<longleftrightarrow> (\<exists>x\<in>A. x)"
-  by auto
-
-declare Complete_Lattices.Inf_bool_def [simp del]
-declare Complete_Lattices.Sup_bool_def [simp del]
-
 end
--- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Tue Sep 13 13:17:52 2011 +0200
+++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Tue Sep 13 16:21:48 2011 +0200
@@ -101,7 +101,7 @@
     then show False using MInf by auto
   next
     case PInf then have "S={\<infinity>}" by (metis Inf_eq_PInfty assms(2))
-    then show False using `Inf S ~: S` by simp
+    then show False using `Inf S ~: S` by (simp add: top_ereal_def)
   next
     case (real r) then have fin: "\<bar>Inf S\<bar> \<noteq> \<infinity>" by simp
     from ereal_open_cont_interval[OF a this] guess e . note e = this
@@ -143,7 +143,8 @@
     from ereal_open_cont_interval[OF assms(1) * fin] guess e . note e = this
     then obtain b where b_def: "Inf S-e<b & b<Inf S"
       using fin ereal_between[of "Inf S" e] ereal_dense[of "Inf S-e"] by auto
-    hence "b: {Inf S-e <..< Inf S+e}" using e fin ereal_between[of "Inf S" e] by auto
+    hence "b: {Inf S-e <..< Inf S+e}" using e fin ereal_between[of "Inf S" e]
+      by auto
     hence "b:S" using e by auto
     hence False using b_def by (metis complete_lattice_class.Inf_lower leD)
   } ultimately show False by auto
@@ -247,7 +248,7 @@
     show "eventually (\<lambda>x. a * X x \<in> S) net"
       by (rule eventually_mono[OF _ *]) auto
   qed
-qed (auto intro: tendsto_const)
+qed auto
 
 lemma ereal_lim_uminus:
   fixes X :: "'a \<Rightarrow> ereal" shows "((\<lambda>i. - X i) ---> -L) net \<longleftrightarrow> (X ---> L) net"
@@ -306,7 +307,7 @@
     assume S: "S = {Inf S<..}"
     then have "Inf S < l" using `l \<in> S` by auto
     then have "eventually (\<lambda>x. Inf S < f x) net" using ev by auto
-    then show "eventually (\<lambda>x. f x \<in> S) net"  by (subst S) auto
+    then show "eventually (\<lambda>x. f x \<in> S) net" by (subst S) auto
   qed auto
 next
   fix l y assume S: "\<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> l \<in> S \<longrightarrow> eventually  (\<lambda>x. f x \<in> S) net" "y < l"
--- a/src/HOL/Probability/Caratheodory.thy	Tue Sep 13 13:17:52 2011 +0200
+++ b/src/HOL/Probability/Caratheodory.thy	Tue Sep 13 16:21:48 2011 +0200
@@ -613,7 +613,7 @@
   assumes posf: "positive M f"
   shows "increasing \<lparr> space = space M, sets = Pow (space M) \<rparr>
                     (\<lambda>x. Inf (measure_set M f x))"
-apply (auto simp add: increasing_def)
+apply (clarsimp simp add: increasing_def)
 apply (rule complete_lattice_class.Inf_greatest)
 apply (rule complete_lattice_class.Inf_lower)
 apply (clarsimp simp add: measure_set_def, rule_tac x=A in exI, blast)
--- a/src/HOL/Probability/Radon_Nikodym.thy	Tue Sep 13 13:17:52 2011 +0200
+++ b/src/HOL/Probability/Radon_Nikodym.thy	Tue Sep 13 16:21:48 2011 +0200
@@ -411,7 +411,9 @@
   also have "\<dots> = ?y"
   proof (rule antisym)
     show "(SUP i. integral\<^isup>P M (?g i)) \<le> ?y"
-      using g_in_G by (auto intro!: exI Sup_mono simp: SUPR_def)
+      using g_in_G
+      using [[simp_trace]]
+      by (auto intro!: exI Sup_mono simp: SUPR_def)
     show "?y \<le> (SUP i. integral\<^isup>P M (?g i))" unfolding y_eq
       by (auto intro!: SUP_mono positive_integral_mono Max_ge)
   qed
--- a/src/HOL/UNITY/ProgressSets.thy	Tue Sep 13 13:17:52 2011 +0200
+++ b/src/HOL/UNITY/ProgressSets.thy	Tue Sep 13 16:21:48 2011 +0200
@@ -77,7 +77,7 @@
 by (simp add: cl_def, blast)
 
 lemma subset_cl: "r \<subseteq> cl L r"
-by (simp add: cl_def, blast)
+by (simp add: cl_def le_Inf_iff)
 
 text{*A reformulation of @{thm subset_cl}*}
 lemma clI: "x \<in> r ==> x \<in> cl L r"