merged
authorwenzelm
Mon, 25 Jul 2016 11:30:31 +0200
changeset 63547 00521f181510
parent 63538 d7b5e2a222c2 (current diff)
parent 63546 5f097087fa1e (diff)
child 63548 6c2c16fef8f1
merged
src/HOL/Library/Multiset.thy
src/HOL/Number_Theory/Factorial_Ring.thy
--- a/src/HOL/Archimedean_Field.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Archimedean_Field.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -395,9 +395,9 @@
       by blast
     then have l: "l = - r"
       by simp
-    moreover with \<open>l \<noteq> 0\<close> False have "r > 0"
+    with \<open>l \<noteq> 0\<close> False have "r > 0"
       by simp
-    ultimately show ?thesis
+    with l show ?thesis
       using pos_mod_bound [of r]
       by (auto simp add: zmod_zminus2_eq_if less_le field_simps intro: floor_unique)
   qed
--- a/src/HOL/Cardinals/Ordinal_Arithmetic.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Cardinals/Ordinal_Arithmetic.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -310,11 +310,11 @@
   have minim[simp]: "minim r' (Field r') \<in> Field r'"
     using wo_rel.minim_inField[unfolded wo_rel_def, OF WELL' _ r'] by auto
   { fix b
-    assume "(b, minim r' (Field r')) \<in> r'"
-    moreover hence "b \<in> Field r'" unfolding Field_def by auto
+    assume b: "(b, minim r' (Field r')) \<in> r'"
+    hence "b \<in> Field r'" unfolding Field_def by auto
     hence "(minim r' (Field r'), b) \<in> r'"
       using wo_rel.minim_least[unfolded wo_rel_def, OF WELL' subset_refl] r' by auto
-    ultimately have "b = minim r' (Field r')"
+    with b have "b = minim r' (Field r')"
       by (metis WELL' antisym_def linear_order_on_def partial_order_on_def well_order_on_def)
   } note * = this
   have 1: "Well_order (r *o r')" using assms by (auto simp add: oprod_Well_order)
@@ -772,9 +772,9 @@
             have "G \<subseteq> fin_support r.zero (Field s)"
             unfolding FinFunc_def fin_support_def proof safe
               fix g assume "g \<in> G"
-              with GF obtain f where "f \<in> F" "g = f(z := r.zero)" by auto
-              moreover with SUPPF have "finite (SUPP f)" by blast
-              ultimately show "finite (SUPP g)"
+              with GF obtain f where f: "f \<in> F" "g = f(z := r.zero)" by auto
+              with SUPPF have "finite (SUPP f)" by blast
+              with f show "finite (SUPP g)"
                 by (elim finite_subset[rotated]) (auto simp: support_def)
             qed
             moreover from F GF zField have "G \<subseteq> Func (Field s) (Field r)"
@@ -831,13 +831,12 @@
                   case True
                   with f have "f(z := r.zero) \<in> G" unfolding G_def by blast
                   with g0(2) f0z have "(f0(z := r.zero), f(z := r.zero)) \<in> oexp" by auto
-                  hence "(f0(z := r.zero, z := x0), f(z := r.zero, z := x0)) \<in> oexp"
+                  hence oexp: "(f0(z := r.zero, z := x0), f(z := r.zero, z := x0)) \<in> oexp"
                     by (elim fun_upd_same_oexp[OF _ _ zField x0Field]) simp
-                  moreover
                   with f F(1) x0min True
                     have "(f(z := x0), f) \<in> oexp" unfolding G_def r.isMinim_def
                     by (intro fun_upd_smaller_oexp[OF _ zField x0Field]) auto
-                  ultimately show ?thesis using transD[OF oexp_trans, of f0 "f(z := x0)" f]
+                  with oexp show ?thesis using transD[OF oexp_trans, of f0 "f(z := x0)" f]
                     unfolding f0_def by auto
                 next
                   case False note notG = this
@@ -1336,10 +1335,9 @@
       hence max_Field: "t.max_fun_diff h (F g) \<in> {a \<in> Field t. h a \<noteq> F g a}"
         by (rule rt.max_fun_diff_in[OF _ gh(2,3)])
       { assume *: "t.max_fun_diff h (F g) \<notin> f ` Field s"
-        with max_Field have "F g (t.max_fun_diff h (F g)) = r.zero" unfolding F_def by auto
-        moreover
+        with max_Field have **: "F g (t.max_fun_diff h (F g)) = r.zero" unfolding F_def by auto
         with * gh(4) have "h (t.max_fun_diff h (F g)) = r.zero" unfolding Let_def by auto
-        ultimately have False using max_Field gh(2,3) unfolding FinFunc_def Func_def by auto
+        with ** have False using max_Field gh(2,3) unfolding FinFunc_def Func_def by auto
       }
       hence max_f_Field: "t.max_fun_diff h (F g) \<in> f ` Field s" by blast
       { fix z assume z: "z \<in> Field t - f ` Field s"
@@ -1433,7 +1431,6 @@
           with f inj have neq: "?f h \<noteq> ?f g"
             unfolding fun_eq_iff inj_on_def rt.oexp_def map_option_case FinFunc_def Func_def Let_def
             by simp metis
-          moreover
           with hg have "t.max_fun_diff (?f h) (?f g) = t.max_fun_diff h g" unfolding rt.oexp_def
             using rt.max_fun_diff[OF \<open>h \<noteq> g\<close>] rt.max_fun_diff_in[OF \<open>h \<noteq> g\<close>]
             by (subst t.max_fun_diff_def, intro t.maxim_equality)
@@ -1441,7 +1438,7 @@
           with Field_fg Field_fh hg fz f_underS compat neq have "(?f h, ?f g) \<in> st.oexp"
              using rt.max_fun_diff[OF \<open>h \<noteq> g\<close>] rt.max_fun_diff_in[OF \<open>h \<noteq> g\<close>] unfolding st.Field_oexp
              unfolding rt.oexp_def st.oexp_def Let_def compat_def by auto
-          ultimately show "?f h \<in> underS (s ^o t) (?f g)" unfolding underS_def by auto
+          with neq show "?f h \<in> underS (s ^o t) (?f g)" unfolding underS_def by auto
         qed
         ultimately have "?f g \<in> Field (s ^o t) \<and> ?f ` underS (r ^o t) g \<subseteq> underS (s ^o t) (?f g)"
           by blast
@@ -1609,12 +1606,12 @@
     have **: "(\<Union>g \<in> fg ` Field t. rs.SUPP g) =
               (\<Union>g \<in> fg ` Field t - {rs.const}. rs.SUPP g)"
       unfolding support_def by auto
-    from * have "\<forall>g \<in> fg ` Field t. finite (rs.SUPP g)" "finite (rst.SUPP fg)"
+    from * have ***: "\<forall>g \<in> fg ` Field t. finite (rs.SUPP g)" "finite (rst.SUPP fg)"
       unfolding rs.Field_oexp FinFunc_def Func_def fin_support_def Option.these_def by force+
-    moreover hence "finite (fg ` Field t - {rs.const})" using *
+    hence "finite (fg ` Field t - {rs.const})" using *
       unfolding support_def rs.zero_oexp[OF Field] FinFunc_def Func_def
       by (elim finite_surj[of _ _ fg]) (fastforce simp: image_iff Option.these_def)
-    ultimately have "finite ((\<Union>g \<in> fg ` Field t. rs.SUPP g) \<times> rst.SUPP fg)"
+    with *** have "finite ((\<Union>g \<in> fg ` Field t. rs.SUPP g) \<times> rst.SUPP fg)"
       by (subst **) (auto intro!: finite_cartesian_product)
     with * show "?g \<in> FinFunc r (s *o t)"
       unfolding Field_oprod rs.Field_oexp FinFunc_def Func_def fin_support_def Option.these_def
@@ -1680,8 +1677,9 @@
           ordIso_transitive[OF oone_ordIso_oexp[OF ordIso_symmetric[OF oone_ordIso] t] oone_ordIso])
   next
      case False
-     moreover hence "s *o t \<noteq> {}" unfolding oprod_def Field_def by fastforce
-     ultimately show ?thesis using \<open>r = {}\<close> ozero_ordIso
+     hence "s *o t \<noteq> {}" unfolding oprod_def Field_def by fastforce
+     with False show ?thesis
+       using \<open>r = {}\<close> ozero_ordIso
        by (auto simp add: s t oprod_Well_order ozero_def)
   qed
 next
--- a/src/HOL/Cardinals/Wellorder_Constructions.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Cardinals/Wellorder_Constructions.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -606,13 +606,13 @@
 
 lemma not_isSucc_zero: "\<not> isSucc zero"
 proof
-  assume "isSucc zero"
-  moreover
+  assume *: "isSucc zero"
   then obtain i where "aboveS i \<noteq> {}" and 1: "minim (Field r) = succ i"
   unfolding isSucc_def zero_def by auto
   hence "succ i \<in> Field r" by auto
-  ultimately show False by (metis REFL isSucc_def minim_least refl_on_domain
-    subset_refl succ_in succ_not_in zero_def)
+  with * show False
+    by (metis REFL isSucc_def minim_least refl_on_domain
+        subset_refl succ_in succ_not_in zero_def)
 qed
 
 lemma isLim_zero[simp]: "isLim zero"
--- a/src/HOL/Conditionally_Complete_Lattices.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Conditionally_Complete_Lattices.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -505,19 +505,23 @@
 
 instance
 proof
-  fix x :: nat and X :: "nat set"
-  { assume "x \<in> X" "bdd_below X" then show "Inf X \<le> x"
-      by (simp add: Inf_nat_def Least_le) }
-  { assume "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> x \<le> y" then show "x \<le> Inf X"
-      unfolding Inf_nat_def ex_in_conv[symmetric] by (rule LeastI2_ex) }
-  { assume "x \<in> X" "bdd_above X" then show "x \<le> Sup X"
-      by (simp add: Sup_nat_def bdd_above_nat) }
-  { assume "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> y \<le> x"
-    moreover then have "bdd_above X"
+  fix x :: nat
+  fix X :: "nat set"
+  show "Inf X \<le> x" if "x \<in> X" "bdd_below X"
+    using that by (simp add: Inf_nat_def Least_le)
+  show "x \<le> Inf X" if "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> x \<le> y"
+    using that unfolding Inf_nat_def ex_in_conv[symmetric] by (rule LeastI2_ex)
+  show "x \<le> Sup X" if "x \<in> X" "bdd_above X"
+    using that by (simp add: Sup_nat_def bdd_above_nat)
+  show "Sup X \<le> x" if "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> y \<le> x"
+  proof -
+    from that have "bdd_above X"
       by (auto simp: bdd_above_def)
-    ultimately show "Sup X \<le> x"
-      by (simp add: Sup_nat_def bdd_above_nat) }
+    with that show ?thesis 
+      by (simp add: Sup_nat_def bdd_above_nat)
+  qed
 qed
+
 end
 
 instantiation int :: conditionally_complete_linorder
--- a/src/HOL/Corec_Examples/LFilter.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Corec_Examples/LFilter.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -30,7 +30,6 @@
   from this(1,2) obtain a where "P (lhd ((ltl ^^ a) xs))"
     by (atomize_elim, induct x xs rule: llist.set_induct)
        (auto simp: funpow_Suc_right simp del: funpow.simps(2) intro: exI[of _ 0] exI[of _ "Suc i" for i])
-  moreover
   with \<open>\<not> P (lhd xs)\<close>
     have "(LEAST n. P (lhd ((ltl ^^ n) xs))) = Suc (LEAST n. P (lhd ((ltl ^^ Suc n) xs)))"
     by (intro Least_Suc) auto
--- a/src/HOL/Corec_Examples/Tests/Misc_Mono.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Corec_Examples/Tests/Misc_Mono.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -382,7 +382,6 @@
   from this(1,2) obtain a where "P (head ((tail ^^ a) xs))"
     by (atomize_elim, induct xs x rule: llist_in.induct) (auto simp: funpow_Suc_right
       simp del: funpow.simps(2) intro: exI[of _ 0] exI[of _ "Suc i" for i])
-  moreover
   with \<open>\<not> P (head xs)\<close>
     have "(LEAST n. P (head ((tail ^^ n) xs))) = Suc (LEAST n. P (head ((tail ^^ Suc n) xs)))"
     by (intro Least_Suc) auto
--- a/src/HOL/Filter.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Filter.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -489,9 +489,9 @@
   assumes **: "\<And>i P. i \<in> I \<Longrightarrow> \<forall>\<^sub>F x in F i. P x \<Longrightarrow> \<forall>\<^sub>F x in F' i. T P x"
   shows "\<forall>\<^sub>F x in \<Sqinter>i\<in>I. F' i. T P x"
 proof -
-  from * obtain X where "finite X" "X \<subseteq> I" "\<forall>\<^sub>F x in \<Sqinter>i\<in>X. F i. P x"
+  from * obtain X where X: "finite X" "X \<subseteq> I" "\<forall>\<^sub>F x in \<Sqinter>i\<in>X. F i. P x"
     unfolding eventually_INF[of _ _ I] by auto
-  moreover then have "eventually (T P) (INFIMUM X F')"
+  then have "eventually (T P) (INFIMUM X F')"
     apply (induction X arbitrary: P)
     apply (auto simp: eventually_inf T2)
     subgoal for x S P Q R
@@ -501,7 +501,7 @@
       apply (auto intro: T1) []
       done
     done
-  ultimately show "\<forall>\<^sub>F x in \<Sqinter>i\<in>I. F' i. T P x"
+  with X show "\<forall>\<^sub>F x in \<Sqinter>i\<in>I. F' i. T P x"
     by (subst eventually_INF) auto
 qed
 
@@ -798,9 +798,10 @@
   shows "(\<forall>\<^sub>F (x, y) in A \<times>\<^sub>F B. P x) \<longleftrightarrow> (\<forall>\<^sub>F x in A. P x)"
   unfolding eventually_prod_filter
 proof safe
-  fix R Q assume "\<forall>\<^sub>F x in A. R x" "\<forall>\<^sub>F x in B. Q x" "\<forall>x y. R x \<longrightarrow> Q y \<longrightarrow> P x"
-  moreover with \<open>B \<noteq> bot\<close> obtain y where "Q y" by (auto dest: eventually_happens)
-  ultimately show "eventually P A"
+  fix R Q
+  assume *: "\<forall>\<^sub>F x in A. R x" "\<forall>\<^sub>F x in B. Q x" "\<forall>x y. R x \<longrightarrow> Q y \<longrightarrow> P x"
+  with \<open>B \<noteq> bot\<close> obtain y where "Q y" by (auto dest: eventually_happens)
+  with * show "eventually P A"
     by (force elim: eventually_mono)
 next
   assume "eventually P A"
@@ -813,9 +814,10 @@
   shows "(\<forall>\<^sub>F (x, y) in A \<times>\<^sub>F B. P y) \<longleftrightarrow> (\<forall>\<^sub>F y in B. P y)"
   unfolding eventually_prod_filter
 proof safe
-  fix R Q assume "\<forall>\<^sub>F x in A. R x" "\<forall>\<^sub>F x in B. Q x" "\<forall>x y. R x \<longrightarrow> Q y \<longrightarrow> P y"
-  moreover with \<open>A \<noteq> bot\<close> obtain x where "R x" by (auto dest: eventually_happens)
-  ultimately show "eventually P B"
+  fix R Q
+  assume *: "\<forall>\<^sub>F x in A. R x" "\<forall>\<^sub>F x in B. Q x" "\<forall>x y. R x \<longrightarrow> Q y \<longrightarrow> P y"
+  with \<open>A \<noteq> bot\<close> obtain x where "R x" by (auto dest: eventually_happens)
+  with * show "eventually P B"
     by (force elim: eventually_mono)
 next
   assume "eventually P B"
@@ -827,35 +829,40 @@
   fixes F :: "'a \<Rightarrow> 'b filter"
   assumes *: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> \<exists>k\<in>I. F k \<le> F i \<sqinter> F j"
   shows "(INF i:I. F i) = bot \<longleftrightarrow> (\<exists>i\<in>I. F i = bot)"
-proof cases
-  assume "\<exists>i\<in>I. F i = bot"
-  moreover then have "(INF i:I. F i) \<le> bot"
+proof (cases "\<exists>i\<in>I. F i = bot")
+  case True
+  then have "(INF i:I. F i) \<le> bot"
     by (auto intro: INF_lower2)
-  ultimately show ?thesis
+  with True show ?thesis
     by (auto simp: bot_unique)
 next
-  assume **: "\<not> (\<exists>i\<in>I. F i = bot)"
+  case False
   moreover have "(INF i:I. F i) \<noteq> bot"
-  proof cases
-    assume "I \<noteq> {}"
+  proof (cases "I = {}")
+    case True
+    then show ?thesis
+      by (auto simp add: filter_eq_iff)
+  next
+    case False': False
     show ?thesis
     proof (rule INF_filter_not_bot)
-      fix J assume "finite J" "J \<subseteq> I"
+      fix J
+      assume "finite J" "J \<subseteq> I"
       then have "\<exists>k\<in>I. F k \<le> (\<Sqinter>i\<in>J. F i)"
-      proof (induction J)
-        case empty then show ?case
+      proof (induct J)
+        case empty
+        then show ?case
           using \<open>I \<noteq> {}\<close> by auto
       next
         case (insert i J)
-        moreover then obtain k where "k \<in> I" "F k \<le> (\<Sqinter>i\<in>J. F i)" by auto
-        moreover note *[of i k]
-        ultimately show ?case
+        then obtain k where "k \<in> I" "F k \<le> (\<Sqinter>i\<in>J. F i)" by auto
+        with insert *[of i k] show ?case
           by auto
       qed
-      with ** show "(\<Sqinter>i\<in>J. F i) \<noteq> \<bottom>"
+      with False show "(\<Sqinter>i\<in>J. F i) \<noteq> \<bottom>"
         by (auto simp: bot_unique)
     qed
-  qed (auto simp add: filter_eq_iff)
+  qed
   ultimately show ?thesis
     by auto
 qed
@@ -883,9 +890,9 @@
   shows "A \<times>\<^sub>F B \<le> C \<times>\<^sub>F D \<longleftrightarrow> A \<le> C \<and> B \<le> D"
 proof safe
   assume *: "A \<times>\<^sub>F B \<le> C \<times>\<^sub>F D"
-  moreover with assms have "A \<times>\<^sub>F B \<noteq> bot"
+  with assms have "A \<times>\<^sub>F B \<noteq> bot"
     by (auto simp: bot_unique prod_filter_eq_bot)
-  ultimately have "C \<times>\<^sub>F D \<noteq> bot"
+  with * have "C \<times>\<^sub>F D \<noteq> bot"
     by (auto simp: bot_unique)
   then have nCD: "C \<noteq> bot" "D \<noteq> bot"
     by (auto simp: prod_filter_eq_bot)
--- a/src/HOL/Hilbert_Choice.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Hilbert_Choice.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -317,14 +317,17 @@
 proof -
   define Sseq where "Sseq = rec_nat S (\<lambda>n T. T - {SOME e. e \<in> T})"
   define pick where "pick n = (SOME e. e \<in> Sseq n)" for n
-  { fix n have "Sseq n \<subseteq> S" "\<not> finite (Sseq n)" by (induct n) (auto simp add: Sseq_def inf) }
-  moreover then have *: "\<And>n. pick n \<in> Sseq n"
+  have *: "Sseq n \<subseteq> S" "\<not> finite (Sseq n)" for n
+    by (induct n) (auto simp add: Sseq_def inf)
+  then have **: "\<And>n. pick n \<in> Sseq n"
     unfolding pick_def by (subst (asm) finite.simps) (auto simp add: ex_in_conv intro: someI_ex)
-  ultimately have "range pick \<subseteq> S" by auto
+  with * have "range pick \<subseteq> S" by auto
   moreover
-  { fix n m
-    have "pick n \<notin> Sseq (n + Suc m)" by (induct m) (auto simp add: Sseq_def pick_def)
-    with * have "pick n \<noteq> pick (n + Suc m)" by auto
+  {
+    fix n m
+    have "pick n \<notin> Sseq (n + Suc m)"
+      by (induct m) (auto simp add: Sseq_def pick_def)
+    with ** have "pick n \<noteq> pick (n + Suc m)" by auto
   }
   then have "inj pick" by (intro linorder_injI) (auto simp add: less_iff_Suc_add)
   ultimately show ?thesis by blast
--- a/src/HOL/IMP/Compiler2.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/IMP/Compiler2.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -366,9 +366,8 @@
     rest: "P @ P' \<turnstile> (i', s'', stk'') \<rightarrow>^k (n, s', stk')"
     by auto
   from step `size P \<le> i`
-  have "P' \<turnstile> (i - size P, s, stk) \<rightarrow> (i' - size P, s'', stk'')" 
+  have *: "P' \<turnstile> (i - size P, s, stk) \<rightarrow> (i' - size P, s'', stk'')" 
     by (rule exec1_drop_left)
-  moreover
   then have "i' - size P \<in> succs P' 0"
     by (fastforce dest!: succs_iexec1 simp: exec1_def simp del: iexec.simps)
   with `exits P' \<subseteq> {0..}`
@@ -376,8 +375,7 @@
   from rest this `exits P' \<subseteq> {0..}`     
   have "P' \<turnstile> (i' - size P, s'', stk'') \<rightarrow>^k (n - size P, s', stk')"
     by (rule Suc.IH)
-  ultimately
-  show ?case by auto
+  with * show ?case by auto
 qed
 
 lemmas exec_n_drop_Cons = 
--- a/src/HOL/IMP/Def_Init_Small.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/IMP/Def_Init_Small.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -58,10 +58,9 @@
   "(c,s) \<rightarrow> (c',s') \<Longrightarrow> D (dom s) c A \<Longrightarrow> EX A'. D (dom s') c' A' & A <= A'"
 proof (induction arbitrary: A rule: small_step_induct)
   case (While b c s)
-  then obtain A' where "vars b \<subseteq> dom s" "A = dom s" "D (dom s) c A'" by blast
-  moreover
+  then obtain A' where A': "vars b \<subseteq> dom s" "A = dom s" "D (dom s) c A'" by blast
   then obtain A'' where "D A' c A''" by (metis D_incr D_mono)
-  ultimately have "D (dom s) (IF b THEN c;; WHILE b DO c ELSE SKIP) (dom s)"
+  with A' have "D (dom s) (IF b THEN c;; WHILE b DO c ELSE SKIP) (dom s)"
     by (metis D.If[OF `vars b \<subseteq> dom s` D.Seq[OF `D (dom s) c A'` D.While[OF _ `D A' c A''`]] D.Skip] D_incr Int_absorb1 subset_trans)
   thus ?case by (metis D_incr `A = dom s`)
 next
--- a/src/HOL/IMP/Sec_TypingT.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/IMP/Sec_TypingT.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -107,9 +107,9 @@
     have "s = s' (\<le> l)" by auto
     moreover
     from termi_if_non0[OF 1 0, of t] obtain t' where
-      "(IF b THEN c1 ELSE c2,t) \<Rightarrow> t'" ..
+      t': "(IF b THEN c1 ELSE c2,t) \<Rightarrow> t'" ..
     moreover
-    from confinement[OF this 1] `\<not> sec b \<le> l`
+    from confinement[OF t' 1] `\<not> sec b \<le> l`
     have "t = t' (\<le> l)" by auto
     ultimately
     show ?case using `s = t (\<le> l)` by auto
@@ -134,9 +134,9 @@
     have "s = s' (\<le> l)" by auto
     moreover
     from termi_if_non0[OF 1 0, of t] obtain t' where
-      "(IF b THEN c1 ELSE c2,t) \<Rightarrow> t'" ..
+      t': "(IF b THEN c1 ELSE c2,t) \<Rightarrow> t'" ..
     moreover
-    from confinement[OF this 1] `\<not> sec b \<le> l`
+    from confinement[OF t' 1] `\<not> sec b \<le> l`
     have "t = t' (\<le> l)" by auto
     ultimately
     show ?case using `s = t (\<le> l)` by auto
--- a/src/HOL/Inductive.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Inductive.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -270,10 +270,10 @@
   show "g (lfp (\<lambda>x. f (g x))) \<le> lfp (\<lambda>x. g (f x))"
   proof (rule lfp_greatest)
     fix u
-    assume "g (f u) \<le> u"
-    moreover then have "g (lfp (\<lambda>x. f (g x))) \<le> g (f u)"
+    assume u: "g (f u) \<le> u"
+    then have "g (lfp (\<lambda>x. f (g x))) \<le> g (f u)"
       by (intro assms[THEN monoD] lfp_lowerbound)
-    ultimately show "g (lfp (\<lambda>x. f (g x))) \<le> u"
+    with u show "g (lfp (\<lambda>x. f (g x))) \<le> u"
       by auto
   qed
 qed
@@ -307,10 +307,11 @@
     by (rule gfp_upperbound) (simp add: gfp_unfold[OF *, symmetric])
   show "gfp (\<lambda>x. g (f x)) \<le> g (gfp (\<lambda>x. f (g x)))"
   proof (rule gfp_least)
-    fix u assume "u \<le> g (f u)"
-    moreover then have "g (f u) \<le> g (gfp (\<lambda>x. f (g x)))"
+    fix u
+    assume u: "u \<le> g (f u)"
+    then have "g (f u) \<le> g (gfp (\<lambda>x. f (g x)))"
       by (intro assms[THEN monoD] gfp_upperbound)
-    ultimately show "u \<le> g (gfp (\<lambda>x. f (g x)))"
+    with u show "u \<le> g (gfp (\<lambda>x. f (g x)))"
       by auto
   qed
 qed
--- a/src/HOL/Int.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Int.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -633,13 +633,13 @@
   case equal with assms(1) show P by simp
 next
   case greater
-  then have "nat k > 0" by simp
-  moreover from this have "k = int (nat k)" by auto
+  then have *: "nat k > 0" by simp
+  moreover from * have "k = int (nat k)" by auto
   ultimately show P using assms(2) by blast
 next
   case less
-  then have "nat (- k) > 0" by simp
-  moreover from this have "k = - int (nat (- k))" by auto
+  then have *: "nat (- k) > 0" by simp
+  moreover from * have "k = - int (nat (- k))" by auto
   ultimately show P using assms(3) by blast
 qed
 
--- a/src/HOL/Library/Bourbaki_Witt_Fixpoint.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Library/Bourbaki_Witt_Fixpoint.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -7,7 +7,9 @@
 
 section \<open>The Bourbaki-Witt tower construction for transfinite iteration\<close>
 
-theory Bourbaki_Witt_Fixpoint imports Main begin
+theory Bourbaki_Witt_Fixpoint
+  imports Main
+begin
 
 lemma ChainsI [intro?]:
   "(\<And>a b. \<lbrakk> a \<in> Y; b \<in> Y \<rbrakk> \<Longrightarrow> (a, b) \<in> r \<or> (b, a) \<in> r) \<Longrightarrow> Y \<in> Chains r"
@@ -131,10 +133,10 @@
         with Sup.IH[OF z] \<open>y = a\<close> Sup.hyps(3)[OF z]
         show ?thesis by(auto dest: iterates_above_ge intro: a)
       next
-        assume "y \<in> iterates_above (f a)"
-        moreover with increasing[OF a] have "y \<in> Field leq"
+        assume *: "y \<in> iterates_above (f a)"
+        with increasing[OF a] have "y \<in> Field leq"
           by(auto dest!: iterates_above_Field intro: FieldI2)
-        ultimately show ?thesis using y by(auto)
+        with * show ?thesis using y by auto
       qed
     qed
     thus ?thesis by simp
--- a/src/HOL/Library/Continuum_Not_Denumerable.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Library/Continuum_Not_Denumerable.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -153,11 +153,11 @@
   assumes "a < b" and "countable A"
   shows "\<exists>x\<in>{a<..<b}. x \<notin> A"
 proof -
-  from \<open>countable A\<close> have "countable (A \<inter> {a<..<b})"
+  from \<open>countable A\<close> have *: "countable (A \<inter> {a<..<b})"
     by auto
-  moreover with \<open>a < b\<close> have "\<not> countable {a<..<b}"
+  with \<open>a < b\<close> have "\<not> countable {a<..<b}"
     by (simp add: uncountable_open_interval)
-  ultimately have "A \<inter> {a<..<b} \<noteq> {a<..<b}"
+  with * have "A \<inter> {a<..<b} \<noteq> {a<..<b}"
     by auto
   then have "A \<inter> {a<..<b} \<subset> {a<..<b}"
     by (intro psubsetI) auto
--- a/src/HOL/Library/Discrete.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Library/Discrete.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -28,11 +28,13 @@
   assume "n > 0"
   show "P n"
   proof (cases "n = 1")
-    case True with one show ?thesis by simp
+    case True
+    with one show ?thesis by simp
   next
-    case False with \<open>n > 0\<close> have "n \<ge> 2" by auto
-    moreover with * have "P (n div 2)" .
-    ultimately show ?thesis by (rule double)
+    case False
+    with \<open>n > 0\<close> have "n \<ge> 2" by auto
+    with * have "P (n div 2)" .
+    with \<open>n \<ge> 2\<close> show ?thesis by (rule double)
   qed
 qed
   
--- a/src/HOL/Library/Extended_Nonnegative_Real.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Library/Extended_Nonnegative_Real.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -1067,11 +1067,11 @@
   fix x y :: ereal assume xy: "0 \<le> x" "0 \<le> y" "x < y"
   moreover
   from ereal_dense3[OF \<open>x < y\<close>]
-  obtain r where "x < ereal (real_of_rat r)" "ereal (real_of_rat r) < y"
+  obtain r where r: "x < ereal (real_of_rat r)" "ereal (real_of_rat r) < y"
     by auto
-  moreover then have "0 \<le> r"
+  then have "0 \<le> r"
     using le_less_trans[OF \<open>0 \<le> x\<close> \<open>x < ereal (real_of_rat r)\<close>] by auto
-  ultimately show "\<exists>r. x < (sup 0 \<circ> ereal) (real_of_rat r) \<and> (sup 0 \<circ> ereal) (real_of_rat r) < y"
+  with r show "\<exists>r. x < (sup 0 \<circ> ereal) (real_of_rat r) \<and> (sup 0 \<circ> ereal) (real_of_rat r) < y"
     by (intro exI[of _ r]) (auto simp: max_absorb2)
 qed
 
@@ -1172,11 +1172,11 @@
       from ennreal_Ex_less_of_nat[OF r] guess n .. note n = this
       have "\<not> (X \<subseteq> enat ` {.. n})"
         using \<open>infinite X\<close> by (auto dest: finite_subset)
-      then obtain x where "x \<in> X" "x \<notin> enat ` {..n}"
+      then obtain x where x: "x \<in> X" "x \<notin> enat ` {..n}"
         by blast
-      moreover then have "of_nat n \<le> x"
+      then have "of_nat n \<le> x"
         by (cases x) (auto simp: of_nat_eq_enat)
-      ultimately show ?thesis
+      with x show ?thesis
         by (auto intro!: bexI[of _ x] less_le_trans[OF n])
     qed
     then have "(SUP x : X. ennreal_of_enat x) = top"
@@ -1217,8 +1217,8 @@
     using zero_neq_one by (intro exI)
   show "\<And>x y::ennreal. x < y \<Longrightarrow> \<exists>z>x. z < y"
   proof transfer
-    fix x y :: ereal assume "0 \<le> x" "x < y"
-    moreover from dense[OF this(2)] guess z ..
+    fix x y :: ereal assume "0 \<le> x" and *: "x < y"
+    moreover from dense[OF *] guess z ..
     ultimately show "\<exists>z\<in>Collect (op \<le> 0). x < z \<and> z < y"
       by (intro bexI[of _ z]) auto
   qed
--- a/src/HOL/Library/Extended_Real.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Library/Extended_Real.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -2036,14 +2036,15 @@
 lemma SUP_ereal_add_left:
   assumes "I \<noteq> {}" "c \<noteq> -\<infinity>"
   shows "(SUP i:I. f i + c :: ereal) = (SUP i:I. f i) + c"
-proof cases
-  assume "(SUP i:I. f i) = - \<infinity>"
-  moreover then have "\<And>i. i \<in> I \<Longrightarrow> f i = -\<infinity>"
+proof (cases "(SUP i:I. f i) = - \<infinity>")
+  case True
+  then have "\<And>i. i \<in> I \<Longrightarrow> f i = -\<infinity>"
     unfolding Sup_eq_MInfty by auto
-  ultimately show ?thesis
+  with True show ?thesis
     by (cases c) (auto simp: \<open>I \<noteq> {}\<close>)
 next
-  assume "(SUP i:I. f i) \<noteq> - \<infinity>" then show ?thesis
+  case False
+  then show ?thesis
     by (subst continuous_at_Sup_mono[where f="\<lambda>x. x + c"])
        (auto simp: continuous_at_imp_continuous_at_within continuous_at mono_def ereal_add_mono \<open>I \<noteq> {}\<close> \<open>c \<noteq> -\<infinity>\<close>)
 qed
@@ -2158,14 +2159,15 @@
   assumes "I \<noteq> {}"
   assumes f: "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i" and c: "0 \<le> c"
   shows "(SUP i:I. c * f i) = c * (SUP i:I. f i)"
-proof cases
-  assume "(SUP i: I. f i) = 0"
-  moreover then have "\<And>i. i \<in> I \<Longrightarrow> f i = 0"
+proof (cases "(SUP i: I. f i) = 0")
+  case True
+  then have "\<And>i. i \<in> I \<Longrightarrow> f i = 0"
     by (metis SUP_upper f antisym)
-  ultimately show ?thesis
+  with True show ?thesis
     by simp
 next
-  assume "(SUP i:I. f i) \<noteq> 0" then show ?thesis
+  case False
+  then show ?thesis
     by (subst continuous_at_Sup_mono[where f="\<lambda>x. c * x"])
        (auto simp: mono_def continuous_at continuous_at_imp_continuous_at_within \<open>I \<noteq> {}\<close>
              intro!: ereal_mult_left_mono c)
--- a/src/HOL/Library/Formal_Power_Series.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Library/Formal_Power_Series.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -1512,7 +1512,7 @@
     assume "Lcm A \<noteq> 0"
     from unbounded obtain f where f: "f \<in> A" "subdegree (Lcm A) < subdegree f"
       unfolding bdd_above_def by (auto simp: not_le)
-    moreover from this and \<open>Lcm A \<noteq> 0\<close> have "subdegree f \<le> subdegree (Lcm A)"
+    moreover from f and \<open>Lcm A \<noteq> 0\<close> have "subdegree f \<le> subdegree (Lcm A)"
       by (intro dvd_imp_subdegree_le dvd_Lcm) simp_all
     ultimately show False by simp
   qed
--- a/src/HOL/Library/Function_Growth.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Library/Function_Growth.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -36,8 +36,8 @@
   shows "a ^ (m - n) = (a ^ m) div (a ^ n)"
 proof -
   define q where "q = m - n"
-  moreover with assms have "m = q + n" by (simp add: q_def)
-  ultimately show ?thesis using \<open>a \<noteq> 0\<close> by (simp add: power_add)
+  with assms have "m = q + n" by (simp add: q_def)
+  with q_def show ?thesis using \<open>a \<noteq> 0\<close> by (simp add: power_add)
 qed
 
 
--- a/src/HOL/Library/More_List.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Library/More_List.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -317,10 +317,10 @@
   then show ?Q
   proof (rule nth_equalityI [rule_format])
     fix n
-    assume "n < length ?xs"
-    moreover with len have "n < length ?ys"
+    assume n: "n < length ?xs"
+    with len have "n < length ?ys"
       by simp
-    ultimately have xs: "nth_default dflt ?xs n = ?xs ! n"
+    with n have xs: "nth_default dflt ?xs n = ?xs ! n"
       and ys: "nth_default dflt ?ys n = ?ys ! n"
       by (simp_all only: nth_default_nth)
     with eq show "?xs ! n = ?ys ! n"
--- a/src/HOL/Library/Multiset.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Library/Multiset.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -2429,11 +2429,11 @@
         moreover have "(a, b) \<in> (r \<inter> D \<times> D)\<inverse>"
           using \<open>(b, a) \<in> r\<close> using \<open>b \<in># B\<close> and \<open>a \<in># B\<close>
           by (auto simp: D_def)
-        ultimately obtain x where "x \<in># X" and "(a, x) \<in> r"
+        ultimately obtain x where x: "x \<in># X" "(a, x) \<in> r"
           using "1.IH" by blast
-        moreover then have "(b, x) \<in> r"
+        then have "(b, x) \<in> r"
           using \<open>trans r\<close> and \<open>(b, a) \<in> r\<close> by (auto dest: transD)
-        ultimately show ?thesis by blast
+        with x show ?thesis by blast
       qed blast
     qed }
   note B_less = this
@@ -2512,9 +2512,9 @@
   obtains a M0 K where "M = M0 + {#a#}" "N = M0 + K"
     "a \<notin># K" "\<And>b. b \<in># K \<Longrightarrow> b < a"
 proof -
-  from assms obtain a M0 K where "M = M0 + {#a#}" "N = M0 + K"
-    "b \<in># K \<Longrightarrow> b < a" for b by (blast elim: mult1E)
-  moreover from this(3) [of a] have "a \<notin># K" by auto
+  from assms obtain a M0 K where "M = M0 + {#a#}" "N = M0 + K" and
+    *: "b \<in># K \<Longrightarrow> b < a" for b by (blast elim: mult1E)
+  moreover from * [of a] have "a \<notin># K" by auto
   ultimately show thesis by (auto intro: that)
 qed
 
--- a/src/HOL/Library/Order_Continuity.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Library/Order_Continuity.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -72,10 +72,11 @@
   shows "sup_continuous (\<lambda>x. f (g x))"
   unfolding sup_continuous_def
 proof safe
-  fix M :: "nat \<Rightarrow> 'c" assume "mono M"
-  moreover then have "mono (\<lambda>i. g (M i))"
+  fix M :: "nat \<Rightarrow> 'c"
+  assume M: "mono M"
+  then have "mono (\<lambda>i. g (M i))"
     using sup_continuous_mono[OF g] by (auto simp: mono_def)
-  ultimately show "f (g (SUPREMUM UNIV M)) = (SUP i. f (g (M i)))"
+  with M show "f (g (SUPREMUM UNIV M)) = (SUP i. f (g (M i)))"
     by (auto simp: sup_continuous_def g[THEN sup_continuousD] f[THEN sup_continuousD])
 qed
 
@@ -269,10 +270,11 @@
   shows "inf_continuous (\<lambda>x. f (g x))"
   unfolding inf_continuous_def
 proof safe
-  fix M :: "nat \<Rightarrow> 'c" assume "antimono M"
-  moreover then have "antimono (\<lambda>i. g (M i))"
+  fix M :: "nat \<Rightarrow> 'c"
+  assume M: "antimono M"
+  then have "antimono (\<lambda>i. g (M i))"
     using inf_continuous_mono[OF g] by (auto simp: mono_def antimono_def)
-  ultimately show "f (g (INFIMUM UNIV M)) = (INF i. f (g (M i)))"
+  with M show "f (g (INFIMUM UNIV M)) = (INF i. f (g (M i)))"
     by (auto simp: inf_continuous_def g[THEN inf_continuousD] f[THEN inf_continuousD])
 qed
 
--- a/src/HOL/Library/Perm.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Library/Perm.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -78,15 +78,15 @@
   assume "card (affected f) = 1"
   then obtain a where *: "affected f = {a}"
     by (rule card_1_singletonE)
-  then have "f \<langle>$\<rangle> a \<noteq> a"
+  then have **: "f \<langle>$\<rangle> a \<noteq> a"
     by (simp add: in_affected [symmetric])
-  moreover with * have "f \<langle>$\<rangle> a \<notin> affected f"
+  with * have "f \<langle>$\<rangle> a \<notin> affected f"
     by simp
   then have "f \<langle>$\<rangle> (f \<langle>$\<rangle> a) = f \<langle>$\<rangle> a"
     by (simp add: in_affected)
   then have "inv (apply f) (f \<langle>$\<rangle> (f \<langle>$\<rangle> a)) = inv (apply f) (f \<langle>$\<rangle> a)"
     by simp
-  ultimately show False by simp
+  with ** show False by simp
 qed
 
 
@@ -203,15 +203,18 @@
     using assms by auto
   then show "(g * f) \<langle>$\<rangle> a = (f * g) \<langle>$\<rangle> a"
   proof cases
-    case 1 moreover with * have "f \<langle>$\<rangle> a \<notin> affected g"
+    case 1
+    with * have "f \<langle>$\<rangle> a \<notin> affected g"
       by auto
-    ultimately show ?thesis by (simp add: in_affected apply_times)
+    with 1 show ?thesis by (simp add: in_affected apply_times)
   next
-    case 2 moreover with * have "g \<langle>$\<rangle> a \<notin> affected f"
+    case 2
+    with * have "g \<langle>$\<rangle> a \<notin> affected f"
       by auto
-    ultimately show ?thesis by (simp add: in_affected apply_times)
+    with 2 show ?thesis by (simp add: in_affected apply_times)
   next
-    case 3 then show ?thesis by (simp add: in_affected apply_times)
+    case 3
+    then show ?thesis by (simp add: in_affected apply_times)
   qed
 qed
 
@@ -593,9 +596,9 @@
   define m where "m = order f a - n mod order f a - 1"
   moreover have "order f a - n mod order f a > 0"
     by simp
-  ultimately have "order f a - n mod order f a = Suc m"
+  ultimately have *: "order f a - n mod order f a = Suc m"
     by arith
-  moreover from this have m2: "order f a - Suc n mod order f a = (if m = 0 then order f a else m)"
+  moreover from * have m2: "order f a - Suc n mod order f a = (if m = 0 then order f a else m)"
     by (auto simp add: mod_Suc)
   ultimately show ?case
     using Suc
--- a/src/HOL/Library/Permutations.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Library/Permutations.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -1269,9 +1269,9 @@
     where x'y': "map_of xs x = Some x'" "map_of xs y = Some y'" 
     by (cases "map_of xs x"; cases "map_of xs y")
        (simp_all add: map_of_eq_None_iff)
-  moreover from this x'y' have "(x,x') \<in> set xs" "(y,y') \<in> set xs"
+  moreover from x'y' have *: "(x,x') \<in> set xs" "(y,y') \<in> set xs"
     by (force dest: map_of_SomeD)+
-  moreover from this eq x'y' have "x' = y'" by simp
+  moreover from * eq x'y' have "x' = y'" by simp
   ultimately show "x = y" using assms
     by (force simp: distinct_map dest: inj_onD[of _ _ "(x,x')" "(y,y')"])
 qed
--- a/src/HOL/Library/Polynomial_FPS.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Library/Polynomial_FPS.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -125,10 +125,10 @@
     
   from assms have "\<not>monom 1 (Suc n) dvd p"
     by (auto simp: monom_1_dvd_iff simp del: power_Suc)
-  then obtain k where "k \<le> n" "fps_of_poly p $ k \<noteq> 0" 
+  then obtain k where k: "k \<le> n" "fps_of_poly p $ k \<noteq> 0" 
     by (auto simp: monom_1_dvd_iff' less_Suc_eq_le)
-  moreover from this and zero[of k] have "k = n" by linarith
-  ultimately show "fps_of_poly p $ n \<noteq> 0" by simp
+  with zero[of k] have "k = n" by linarith
+  with k show "fps_of_poly p $ n \<noteq> 0" by simp
 qed
 
 lemma fps_of_poly_dvd:
--- a/src/HOL/Limits.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Limits.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -8,13 +8,13 @@
 section \<open>Limits on Real Vector Spaces\<close>
 
 theory Limits
-imports Real_Vector_Spaces
+  imports Real_Vector_Spaces
 begin
 
 subsection \<open>Filter going to infinity norm\<close>
 
-definition at_infinity :: "'a::real_normed_vector filter" where
-  "at_infinity = (INF r. principal {x. r \<le> norm x})"
+definition at_infinity :: "'a::real_normed_vector filter"
+  where "at_infinity = (INF r. principal {x. r \<le> norm x})"
 
 lemma eventually_at_infinity: "eventually P at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. b \<le> norm x \<longrightarrow> P x)"
   unfolding at_infinity_def
@@ -22,21 +22,24 @@
      (auto simp: subset_eq eventually_principal intro!: exI[of _ "max a b" for a b])
 
 corollary eventually_at_infinity_pos:
-   "eventually p at_infinity \<longleftrightarrow> (\<exists>b. 0 < b \<and> (\<forall>x. norm x \<ge> b \<longrightarrow> p x))"
-apply (simp add: eventually_at_infinity, auto)
-apply (case_tac "b \<le> 0")
-using norm_ge_zero order_trans zero_less_one apply blast
-apply (force simp:)
-done
-
-lemma at_infinity_eq_at_top_bot:
-  "(at_infinity :: real filter) = sup at_top at_bot"
+  "eventually p at_infinity \<longleftrightarrow> (\<exists>b. 0 < b \<and> (\<forall>x. norm x \<ge> b \<longrightarrow> p x))"
+  apply (simp add: eventually_at_infinity)
+  apply auto
+  apply (case_tac "b \<le> 0")
+  using norm_ge_zero order_trans zero_less_one apply blast
+  apply force
+  done
+
+lemma at_infinity_eq_at_top_bot: "(at_infinity :: real filter) = sup at_top at_bot"
   apply (simp add: filter_eq_iff eventually_sup eventually_at_infinity
-                   eventually_at_top_linorder eventually_at_bot_linorder)
+      eventually_at_top_linorder eventually_at_bot_linorder)
   apply safe
-  apply (rule_tac x="b" in exI, simp)
-  apply (rule_tac x="- b" in exI, simp)
-  apply (rule_tac x="max (- Na) N" in exI, auto simp: abs_real_def)
+    apply (rule_tac x="b" in exI)
+    apply simp
+   apply (rule_tac x="- b" in exI)
+   apply simp
+  apply (rule_tac x="max (- Na) N" in exI)
+  apply (auto simp: abs_real_def)
   done
 
 lemma at_top_le_at_infinity: "at_top \<le> (at_infinity :: real filter)"
@@ -45,23 +48,21 @@
 lemma at_bot_le_at_infinity: "at_bot \<le> (at_infinity :: real filter)"
   unfolding at_infinity_eq_at_top_bot by simp
 
-lemma filterlim_at_top_imp_at_infinity:
-  fixes f :: "_ \<Rightarrow> real"
-  shows "filterlim f at_top F \<Longrightarrow> filterlim f at_infinity F"
+lemma filterlim_at_top_imp_at_infinity: "filterlim f at_top F \<Longrightarrow> filterlim f at_infinity F"
+  for f :: "_ \<Rightarrow> real"
   by (rule filterlim_mono[OF _ at_top_le_at_infinity order_refl])
 
-lemma lim_infinity_imp_sequentially:
-  "(f \<longlongrightarrow> l) at_infinity \<Longrightarrow> ((\<lambda>n. f(n)) \<longlongrightarrow> l) sequentially"
-by (simp add: filterlim_at_top_imp_at_infinity filterlim_compose filterlim_real_sequentially)
+lemma lim_infinity_imp_sequentially: "(f \<longlongrightarrow> l) at_infinity \<Longrightarrow> ((\<lambda>n. f(n)) \<longlongrightarrow> l) sequentially"
+  by (simp add: filterlim_at_top_imp_at_infinity filterlim_compose filterlim_real_sequentially)
 
 
 subsubsection \<open>Boundedness\<close>
 
-definition Bfun :: "('a \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a filter \<Rightarrow> bool" where
-  Bfun_metric_def: "Bfun f F = (\<exists>y. \<exists>K>0. eventually (\<lambda>x. dist (f x) y \<le> K) F)"
-
-abbreviation Bseq :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where
-  "Bseq X \<equiv> Bfun X sequentially"
+definition Bfun :: "('a \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a filter \<Rightarrow> bool"
+  where Bfun_metric_def: "Bfun f F = (\<exists>y. \<exists>K>0. eventually (\<lambda>x. dist (f x) y \<le> K) F)"
+
+abbreviation Bseq :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool"
+  where "Bseq X \<equiv> Bfun X sequentially"
 
 lemma Bseq_conv_Bfun: "Bseq X \<longleftrightarrow> Bfun X sequentially" ..
 
@@ -71,11 +72,11 @@
 lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X"
   unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg)
 
-lemma Bfun_def:
-  "Bfun f F \<longleftrightarrow> (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) F)"
+lemma Bfun_def: "Bfun f F \<longleftrightarrow> (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) F)"
   unfolding Bfun_metric_def norm_conv_dist
 proof safe
-  fix y K assume K: "0 < K" and *: "eventually (\<lambda>x. dist (f x) y \<le> K) F"
+  fix y K
+  assume K: "0 < K" and *: "eventually (\<lambda>x. dist (f x) y \<le> K) F"
   moreover have "eventually (\<lambda>x. dist (f x) 0 \<le> dist (f x) y + dist 0 y) F"
     by (intro always_eventually) (metis dist_commute dist_triangle)
   with * have "eventually (\<lambda>x. dist (f x) 0 \<le> K + dist 0 y) F"
@@ -85,19 +86,19 @@
 qed (force simp del: norm_conv_dist [symmetric])
 
 lemma BfunI:
-  assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) F" shows "Bfun f F"
-unfolding Bfun_def
+  assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) F"
+  shows "Bfun f F"
+  unfolding Bfun_def
 proof (intro exI conjI allI)
   show "0 < max K 1" by simp
-next
   show "eventually (\<lambda>x. norm (f x) \<le> max K 1) F"
-    using K by (rule eventually_mono, simp)
+    using K by (rule eventually_mono) simp
 qed
 
 lemma BfunE:
   assumes "Bfun f F"
   obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) F"
-using assms unfolding Bfun_def by blast
+  using assms unfolding Bfun_def by blast
 
 lemma Cauchy_Bseq: "Cauchy X \<Longrightarrow> Bseq X"
   unfolding Cauchy_def Bfun_metric_def eventually_sequentially
@@ -124,57 +125,66 @@
 lemma Bseq_def: "Bseq X \<longleftrightarrow> (\<exists>K>0. \<forall>n. norm (X n) \<le> K)"
   unfolding Bfun_def eventually_sequentially
 proof safe
-  fix N K assume "0 < K" "\<forall>n\<ge>N. norm (X n) \<le> K"
+  fix N K
+  assume "0 < K" "\<forall>n\<ge>N. norm (X n) \<le> K"
   then show "\<exists>K>0. \<forall>n. norm (X n) \<le> K"
     by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] max.strict_coboundedI2)
        (auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj)
 qed auto
 
-lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
-unfolding Bseq_def by auto
-
-lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)"
-by (simp add: Bseq_def)
-
-lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
-by (auto simp add: Bseq_def)
-
-lemma Bseq_bdd_above: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_above (range X)"
+lemma BseqE: "Bseq X \<Longrightarrow> (\<And>K. 0 < K \<Longrightarrow> \<forall>n. norm (X n) \<le> K \<Longrightarrow> Q) \<Longrightarrow> Q"
+  unfolding Bseq_def by auto
+
+lemma BseqD: "Bseq X \<Longrightarrow> \<exists>K. 0 < K \<and> (\<forall>n. norm (X n) \<le> K)"
+  by (simp add: Bseq_def)
+
+lemma BseqI: "0 < K \<Longrightarrow> \<forall>n. norm (X n) \<le> K \<Longrightarrow> Bseq X"
+  by (auto simp add: Bseq_def)
+
+lemma Bseq_bdd_above: "Bseq X \<Longrightarrow> bdd_above (range X)"
+  for X :: "nat \<Rightarrow> real"
 proof (elim BseqE, intro bdd_aboveI2)
-  fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "X n \<le> K"
+  fix K n
+  assume "0 < K" "\<forall>n. norm (X n) \<le> K"
+  then show "X n \<le> K"
     by (auto elim!: allE[of _ n])
 qed
 
-lemma Bseq_bdd_above':
-  "Bseq (X::nat \<Rightarrow> 'a :: real_normed_vector) \<Longrightarrow> bdd_above (range (\<lambda>n. norm (X n)))"
+lemma Bseq_bdd_above': "Bseq X \<Longrightarrow> bdd_above (range (\<lambda>n. norm (X n)))"
+  for X :: "nat \<Rightarrow> 'a :: real_normed_vector"
 proof (elim BseqE, intro bdd_aboveI2)
-  fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "norm (X n) \<le> K"
+  fix K n
+  assume "0 < K" "\<forall>n. norm (X n) \<le> K"
+  then show "norm (X n) \<le> K"
     by (auto elim!: allE[of _ n])
 qed
 
-lemma Bseq_bdd_below: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_below (range X)"
+lemma Bseq_bdd_below: "Bseq X \<Longrightarrow> bdd_below (range X)"
+  for X :: "nat \<Rightarrow> real"
 proof (elim BseqE, intro bdd_belowI2)
-  fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "- K \<le> X n"
+  fix K n
+  assume "0 < K" "\<forall>n. norm (X n) \<le> K"
+  then show "- K \<le> X n"
     by (auto elim!: allE[of _ n])
 qed
 
 lemma Bseq_eventually_mono:
   assumes "eventually (\<lambda>n. norm (f n) \<le> norm (g n)) sequentially" "Bseq g"
-  shows   "Bseq f"
+  shows "Bseq f"
 proof -
   from assms(1) obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> norm (f n) \<le> norm (g n)"
     by (auto simp: eventually_at_top_linorder)
-  moreover from assms(2) obtain K where K: "\<And>n. norm (g n) \<le> K" by (blast elim!: BseqE)
+  moreover from assms(2) obtain K where K: "\<And>n. norm (g n) \<le> K"
+    by (blast elim!: BseqE)
   ultimately have "norm (f n) \<le> max K (Max {norm (f n) |n. n < N})" for n
     apply (cases "n < N")
-    apply (rule max.coboundedI2, rule Max.coboundedI, auto) []
-    apply (rule max.coboundedI1, force intro: order.trans[OF N K])
+    subgoal by (rule max.coboundedI2, rule Max.coboundedI) auto
+    subgoal by (rule max.coboundedI1) (force intro: order.trans[OF N K])
     done
-  thus ?thesis by (blast intro: BseqI')
+  then show ?thesis by (blast intro: BseqI')
 qed
 
-lemma lemma_NBseq_def:
-  "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
+lemma lemma_NBseq_def: "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) \<longleftrightarrow> (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
 proof safe
   fix K :: real
   from reals_Archimedean2 obtain n :: nat where "K < real n" ..
@@ -188,47 +198,50 @@
     using of_nat_0_less_iff by blast
 qed
 
-text\<open>alternative definition for Bseq\<close>
-lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
-apply (simp add: Bseq_def)
-apply (simp (no_asm) add: lemma_NBseq_def)
-done
-
-lemma lemma_NBseq_def2:
-     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
-apply (subst lemma_NBseq_def, auto)
-apply (rule_tac x = "Suc N" in exI)
-apply (rule_tac [2] x = N in exI)
-apply (auto simp add: of_nat_Suc)
- prefer 2 apply (blast intro: order_less_imp_le)
-apply (drule_tac x = n in spec, simp)
-done
-
-(* yet another definition for Bseq *)
-lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
-by (simp add: Bseq_def lemma_NBseq_def2)
-
-subsubsection\<open>A Few More Equivalence Theorems for Boundedness\<close>
-
-text\<open>alternative formulation for boundedness\<close>
-lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
-apply (unfold Bseq_def, safe)
-apply (rule_tac [2] x = "k + norm x" in exI)
-apply (rule_tac x = K in exI, simp)
-apply (rule exI [where x = 0], auto)
-apply (erule order_less_le_trans, simp)
-apply (drule_tac x=n in spec)
-apply (drule order_trans [OF norm_triangle_ineq2])
-apply simp
-done
-
-text\<open>alternative formulation for boundedness\<close>
-lemma Bseq_iff3:
-  "Bseq X \<longleftrightarrow> (\<exists>k>0. \<exists>N. \<forall>n. norm (X n + - X N) \<le> k)" (is "?P \<longleftrightarrow> ?Q")
+text \<open>Alternative definition for \<open>Bseq\<close>.\<close>
+lemma Bseq_iff: "Bseq X \<longleftrightarrow> (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
+  by (simp add: Bseq_def) (simp add: lemma_NBseq_def)
+
+lemma lemma_NBseq_def2: "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
+  apply (subst lemma_NBseq_def)
+  apply auto
+   apply (rule_tac x = "Suc N" in exI)
+   apply (rule_tac [2] x = N in exI)
+   apply auto
+   prefer 2 apply (blast intro: order_less_imp_le)
+  apply (drule_tac x = n in spec)
+  apply simp
+  done
+
+text \<open>Yet another definition for Bseq.\<close>
+lemma Bseq_iff1a: "Bseq X \<longleftrightarrow> (\<exists>N. \<forall>n. norm (X n) < real (Suc N))"
+  by (simp add: Bseq_def lemma_NBseq_def2)
+
+subsubsection \<open>A Few More Equivalence Theorems for Boundedness\<close>
+
+text \<open>Alternative formulation for boundedness.\<close>
+lemma Bseq_iff2: "Bseq X \<longleftrightarrow> (\<exists>k > 0. \<exists>x. \<forall>n. norm (X n + - x) \<le> k)"
+  apply (unfold Bseq_def)
+  apply safe
+   apply (rule_tac [2] x = "k + norm x" in exI)
+   apply (rule_tac x = K in exI)
+   apply simp
+   apply (rule exI [where x = 0])
+   apply auto
+   apply (erule order_less_le_trans)
+   apply simp
+  apply (drule_tac x=n in spec)
+  apply (drule order_trans [OF norm_triangle_ineq2])
+  apply simp
+  done
+
+text \<open>Alternative formulation for boundedness.\<close>
+lemma Bseq_iff3: "Bseq X \<longleftrightarrow> (\<exists>k>0. \<exists>N. \<forall>n. norm (X n + - X N) \<le> k)"
+  (is "?P \<longleftrightarrow> ?Q")
 proof
   assume ?P
-  then obtain K
-    where *: "0 < K" and **: "\<And>n. norm (X n) \<le> K" by (auto simp add: Bseq_def)
+  then obtain K where *: "0 < K" and **: "\<And>n. norm (X n) \<le> K"
+    by (auto simp add: Bseq_def)
   from * have "0 < K + norm (X 0)" by (rule order_less_le_trans) simp
   from ** have "\<forall>n. norm (X n - X 0) \<le> K + norm (X 0)"
     by (auto intro: order_trans norm_triangle_ineq4)
@@ -236,129 +249,150 @@
     by simp
   with \<open>0 < K + norm (X 0)\<close> show ?Q by blast
 next
-  assume ?Q then show ?P by (auto simp add: Bseq_iff2)
+  assume ?Q
+  then show ?P by (auto simp add: Bseq_iff2)
 qed
 
-lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
-apply (simp add: Bseq_def)
-apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
-apply (drule_tac x = n in spec, arith)
-done
-
-
-subsubsection\<open>Upper Bounds and Lubs of Bounded Sequences\<close>
-
-lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X"
+lemma BseqI2: "\<forall>n. k \<le> f n \<and> f n \<le> K \<Longrightarrow> Bseq f"
+  for k K :: real
+  apply (simp add: Bseq_def)
+  apply (rule_tac x = "(\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI)
+  apply auto
+  apply (drule_tac x = n in spec)
+  apply arith
+  done
+
+
+subsubsection \<open>Upper Bounds and Lubs of Bounded Sequences\<close>
+
+lemma Bseq_minus_iff: "Bseq (\<lambda>n. - (X n) :: 'a::real_normed_vector) \<longleftrightarrow> Bseq X"
   by (simp add: Bseq_def)
 
 lemma Bseq_add:
-  assumes "Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector)"
-  shows   "Bseq (\<lambda>x. f x + c)"
+  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
+  assumes "Bseq f"
+  shows "Bseq (\<lambda>x. f x + c)"
 proof -
-  from assms obtain K where K: "\<And>x. norm (f x) \<le> K" unfolding Bseq_def by blast
+  from assms obtain K where K: "\<And>x. norm (f x) \<le> K"
+    unfolding Bseq_def by blast
   {
     fix x :: nat
     have "norm (f x + c) \<le> norm (f x) + norm c" by (rule norm_triangle_ineq)
     also have "norm (f x) \<le> K" by (rule K)
     finally have "norm (f x + c) \<le> K + norm c" by simp
   }
-  thus ?thesis by (rule BseqI')
+  then show ?thesis by (rule BseqI')
 qed
 
-lemma Bseq_add_iff: "Bseq (\<lambda>x. f x + c) \<longleftrightarrow> Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector)"
+lemma Bseq_add_iff: "Bseq (\<lambda>x. f x + c) \<longleftrightarrow> Bseq f"
+  for f :: "nat \<Rightarrow> 'a::real_normed_vector"
   using Bseq_add[of f c] Bseq_add[of "\<lambda>x. f x + c" "-c"] by auto
 
 lemma Bseq_mult:
-  assumes "Bseq (f :: nat \<Rightarrow> 'a :: real_normed_field)"
-  assumes "Bseq (g :: nat \<Rightarrow> 'a :: real_normed_field)"
-  shows   "Bseq (\<lambda>x. f x * g x)"
+  fixes f g :: "nat \<Rightarrow> 'a::real_normed_field"
+  assumes "Bseq f" and "Bseq g"
+  shows "Bseq (\<lambda>x. f x * g x)"
 proof -
-  from assms obtain K1 K2 where K: "\<And>x. norm (f x) \<le> K1" "K1 > 0" "\<And>x. norm (g x) \<le> K2" "K2 > 0"
+  from assms obtain K1 K2 where K: "norm (f x) \<le> K1" "K1 > 0" "norm (g x) \<le> K2" "K2 > 0"
+    for x
     unfolding Bseq_def by blast
-  hence "\<And>x. norm (f x * g x) \<le> K1 * K2" by (auto simp: norm_mult intro!: mult_mono)
-  thus ?thesis by (rule BseqI')
+  then have "norm (f x * g x) \<le> K1 * K2" for x
+    by (auto simp: norm_mult intro!: mult_mono)
+  then show ?thesis by (rule BseqI')
 qed
 
 lemma Bfun_const [simp]: "Bfun (\<lambda>_. c) F"
   unfolding Bfun_metric_def by (auto intro!: exI[of _ c] exI[of _ "1::real"])
 
-lemma Bseq_cmult_iff: "(c :: 'a :: real_normed_field) \<noteq> 0 \<Longrightarrow> Bseq (\<lambda>x. c * f x) \<longleftrightarrow> Bseq f"
+lemma Bseq_cmult_iff:
+  fixes c :: "'a::real_normed_field"
+  assumes "c \<noteq> 0"
+  shows "Bseq (\<lambda>x. c * f x) \<longleftrightarrow> Bseq f"
 proof
-  assume "c \<noteq> 0" "Bseq (\<lambda>x. c * f x)"
-  find_theorems "Bfun (\<lambda>_. ?c) _"
-  from Bfun_const this(2) have "Bseq (\<lambda>x. inverse c * (c * f x))" by (rule Bseq_mult)
-  with \<open>c \<noteq> 0\<close> show "Bseq f" by (simp add: divide_simps)
+  assume "Bseq (\<lambda>x. c * f x)"
+  with Bfun_const have "Bseq (\<lambda>x. inverse c * (c * f x))"
+    by (rule Bseq_mult)
+  with \<open>c \<noteq> 0\<close> show "Bseq f"
+    by (simp add: divide_simps)
 qed (intro Bseq_mult Bfun_const)
 
-lemma Bseq_subseq: "Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector) \<Longrightarrow> Bseq (\<lambda>x. f (g x))"
+lemma Bseq_subseq: "Bseq f \<Longrightarrow> Bseq (\<lambda>x. f (g x))"
+  for f :: "nat \<Rightarrow> 'a::real_normed_vector"
   unfolding Bseq_def by auto
 
-lemma Bseq_Suc_iff: "Bseq (\<lambda>n. f (Suc n)) \<longleftrightarrow> Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector)"
+lemma Bseq_Suc_iff: "Bseq (\<lambda>n. f (Suc n)) \<longleftrightarrow> Bseq f"
+  for f :: "nat \<Rightarrow> 'a::real_normed_vector"
   using Bseq_offset[of f 1] by (auto intro: Bseq_subseq)
 
 lemma increasing_Bseq_subseq_iff:
-  assumes "\<And>x y. x \<le> y \<Longrightarrow> norm (f x :: 'a :: real_normed_vector) \<le> norm (f y)" "subseq g"
-  shows   "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f"
+  assumes "\<And>x y. x \<le> y \<Longrightarrow> norm (f x :: 'a::real_normed_vector) \<le> norm (f y)" "subseq g"
+  shows "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f"
 proof
   assume "Bseq (\<lambda>x. f (g x))"
-  then obtain K where K: "\<And>x. norm (f (g x)) \<le> K" unfolding Bseq_def by auto
+  then obtain K where K: "\<And>x. norm (f (g x)) \<le> K"
+    unfolding Bseq_def by auto
   {
     fix x :: nat
     from filterlim_subseq[OF assms(2)] obtain y where "g y \<ge> x"
       by (auto simp: filterlim_at_top eventually_at_top_linorder)
-    hence "norm (f x) \<le> norm (f (g y))" using assms(1) by blast
+    then have "norm (f x) \<le> norm (f (g y))"
+      using assms(1) by blast
     also have "norm (f (g y)) \<le> K" by (rule K)
     finally have "norm (f x) \<le> K" .
   }
-  thus "Bseq f" by (rule BseqI')
-qed (insert Bseq_subseq[of f g], simp_all)
+  then show "Bseq f" by (rule BseqI')
+qed (use Bseq_subseq[of f g] in simp_all)
 
 lemma nonneg_incseq_Bseq_subseq_iff:
-  assumes "\<And>x. f x \<ge> 0" "incseq (f :: nat \<Rightarrow> real)" "subseq g"
-  shows   "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f"
+  fixes f :: "nat \<Rightarrow> real"
+    and g :: "nat \<Rightarrow> nat"
+  assumes "\<And>x. f x \<ge> 0" "incseq f" "subseq g"
+  shows "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f"
   using assms by (intro increasing_Bseq_subseq_iff) (auto simp: incseq_def)
 
-lemma Bseq_eq_bounded: "range f \<subseteq> {a .. b::real} \<Longrightarrow> Bseq f"
+lemma Bseq_eq_bounded: "range f \<subseteq> {a..b} \<Longrightarrow> Bseq f"
+  for a b :: real
   apply (simp add: subset_eq)
   apply (rule BseqI'[where K="max (norm a) (norm b)"])
   apply (erule_tac x=n in allE)
   apply auto
   done
 
-lemma incseq_bounded: "incseq X \<Longrightarrow> \<forall>i. X i \<le> (B::real) \<Longrightarrow> Bseq X"
+lemma incseq_bounded: "incseq X \<Longrightarrow> \<forall>i. X i \<le> B \<Longrightarrow> Bseq X"
+  for B :: real
   by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def)
 
-lemma decseq_bounded: "decseq X \<Longrightarrow> \<forall>i. (B::real) \<le> X i \<Longrightarrow> Bseq X"
+lemma decseq_bounded: "decseq X \<Longrightarrow> \<forall>i. B \<le> X i \<Longrightarrow> Bseq X"
+  for B :: real
   by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def)
 
+
 subsection \<open>Bounded Monotonic Sequences\<close>
 
-subsubsection\<open>A Bounded and Monotonic Sequence Converges\<close>
+subsubsection \<open>A Bounded and Monotonic Sequence Converges\<close>
 
 (* TODO: delete *)
 (* FIXME: one use in NSA/HSEQ.thy *)
-lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X \<longlonglongrightarrow> L)"
+lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n \<longrightarrow> X n = X m \<Longrightarrow> \<exists>L. X \<longlonglongrightarrow> L"
   apply (rule_tac x="X m" in exI)
   apply (rule filterlim_cong[THEN iffD2, OF refl refl _ tendsto_const])
   unfolding eventually_sequentially
   apply blast
   done
 
+
 subsection \<open>Convergence to Zero\<close>
 
 definition Zfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
   where "Zfun f F = (\<forall>r>0. eventually (\<lambda>x. norm (f x) < r) F)"
 
-lemma ZfunI:
-  "(\<And>r. 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F) \<Longrightarrow> Zfun f F"
-  unfolding Zfun_def by simp
-
-lemma ZfunD:
-  "\<lbrakk>Zfun f F; 0 < r\<rbrakk> \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F"
-  unfolding Zfun_def by simp
-
-lemma Zfun_ssubst:
-  "eventually (\<lambda>x. f x = g x) F \<Longrightarrow> Zfun g F \<Longrightarrow> Zfun f F"
+lemma ZfunI: "(\<And>r. 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F) \<Longrightarrow> Zfun f F"
+  by (simp add: Zfun_def)
+
+lemma ZfunD: "Zfun f F \<Longrightarrow> 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F"
+  by (simp add: Zfun_def)
+
+lemma Zfun_ssubst: "eventually (\<lambda>x. f x = g x) F \<Longrightarrow> Zfun g F \<Longrightarrow> Zfun f F"
   unfolding Zfun_def by (auto elim!: eventually_rev_mp)
 
 lemma Zfun_zero: "Zfun (\<lambda>x. 0) F"
@@ -369,28 +403,29 @@
 
 lemma Zfun_imp_Zfun:
   assumes f: "Zfun f F"
-  assumes g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F"
+    and g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F"
   shows "Zfun (\<lambda>x. g x) F"
-proof (cases)
-  assume K: "0 < K"
+proof (cases "0 < K")
+  case K: True
   show ?thesis
   proof (rule ZfunI)
-    fix r::real assume "0 < r"
-    hence "0 < r / K" using K by simp
+    fix r :: real
+    assume "0 < r"
+    then have "0 < r / K" using K by simp
     then have "eventually (\<lambda>x. norm (f x) < r / K) F"
       using ZfunD [OF f] by blast
     with g show "eventually (\<lambda>x. norm (g x) < r) F"
     proof eventually_elim
       case (elim x)
-      hence "norm (f x) * K < r"
+      then have "norm (f x) * K < r"
         by (simp add: pos_less_divide_eq K)
-      thus ?case
+      then show ?case
         by (simp add: order_le_less_trans [OF elim(1)])
     qed
   qed
 next
-  assume "\<not> 0 < K"
-  hence K: "K \<le> 0" by (simp only: not_less)
+  case False
+  then have K: "K \<le> 0" by (simp only: not_less)
   show ?thesis
   proof (rule ZfunI)
     fix r :: real
@@ -406,15 +441,17 @@
   qed
 qed
 
-lemma Zfun_le: "\<lbrakk>Zfun g F; \<forall>x. norm (f x) \<le> norm (g x)\<rbrakk> \<Longrightarrow> Zfun f F"
-  by (erule_tac K="1" in Zfun_imp_Zfun, simp)
+lemma Zfun_le: "Zfun g F \<Longrightarrow> \<forall>x. norm (f x) \<le> norm (g x) \<Longrightarrow> Zfun f F"
+  by (erule Zfun_imp_Zfun [where K = 1]) simp
 
 lemma Zfun_add:
-  assumes f: "Zfun f F" and g: "Zfun g F"
+  assumes f: "Zfun f F"
+    and g: "Zfun g F"
   shows "Zfun (\<lambda>x. f x + g x) F"
 proof (rule ZfunI)
-  fix r::real assume "0 < r"
-  hence r: "0 < r / 2" by simp
+  fix r :: real
+  assume "0 < r"
+  then have r: "0 < r / 2" by simp
   have "eventually (\<lambda>x. norm (f x) < r/2) F"
     using f r by (rule ZfunD)
   moreover
@@ -436,14 +473,14 @@
 lemma Zfun_minus: "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. - f x) F"
   unfolding Zfun_def by simp
 
-lemma Zfun_diff: "\<lbrakk>Zfun f F; Zfun g F\<rbrakk> \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F"
+lemma Zfun_diff: "Zfun f F \<Longrightarrow> Zfun g F \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F"
   using Zfun_add [of f F "\<lambda>x. - g x"] by (simp add: Zfun_minus)
 
 lemma (in bounded_linear) Zfun:
   assumes g: "Zfun g F"
   shows "Zfun (\<lambda>x. f (g x)) F"
 proof -
-  obtain K where "\<And>x. norm (f x) \<le> norm x * K"
+  obtain K where "norm (f x) \<le> norm x * K" for x
     using bounded by blast
   then have "eventually (\<lambda>x. norm (f (g x)) \<le> norm (g x) * K) F"
     by simp
@@ -453,12 +490,13 @@
 
 lemma (in bounded_bilinear) Zfun:
   assumes f: "Zfun f F"
-  assumes g: "Zfun g F"
+    and g: "Zfun g F"
   shows "Zfun (\<lambda>x. f x ** g x) F"
 proof (rule ZfunI)
-  fix r::real assume r: "0 < r"
+  fix r :: real
+  assume r: "0 < r"
   obtain K where K: "0 < K"
-    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
+    and norm_le: "norm (x ** y) \<le> norm x * norm y * K" for x y
     using pos_bounded by blast
   from K have K': "0 < inverse K"
     by (rule positive_imp_inverse_positive)
@@ -481,12 +519,10 @@
   qed
 qed
 
-lemma (in bounded_bilinear) Zfun_left:
-  "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. f x ** a) F"
+lemma (in bounded_bilinear) Zfun_left: "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. f x ** a) F"
   by (rule bounded_linear_left [THEN bounded_linear.Zfun])
 
-lemma (in bounded_bilinear) Zfun_right:
-  "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. a ** f x) F"
+lemma (in bounded_bilinear) Zfun_right: "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. a ** f x) F"
   by (rule bounded_linear_right [THEN bounded_linear.Zfun])
 
 lemmas Zfun_mult = bounded_bilinear.Zfun [OF bounded_bilinear_mult]
@@ -496,19 +532,22 @@
 lemma tendsto_Zfun_iff: "(f \<longlongrightarrow> a) F = Zfun (\<lambda>x. f x - a) F"
   by (simp only: tendsto_iff Zfun_def dist_norm)
 
-lemma tendsto_0_le: "\<lbrakk>(f \<longlongrightarrow> 0) F; eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F\<rbrakk>
-                     \<Longrightarrow> (g \<longlongrightarrow> 0) F"
+lemma tendsto_0_le:
+  "(f \<longlongrightarrow> 0) F \<Longrightarrow> eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F \<Longrightarrow> (g \<longlongrightarrow> 0) F"
   by (simp add: Zfun_imp_Zfun tendsto_Zfun_iff)
 
+
 subsubsection \<open>Distance and norms\<close>
 
 lemma tendsto_dist [tendsto_intros]:
-  fixes l m :: "'a :: metric_space"
-  assumes f: "(f \<longlongrightarrow> l) F" and g: "(g \<longlongrightarrow> m) F"
+  fixes l m :: "'a::metric_space"
+  assumes f: "(f \<longlongrightarrow> l) F"
+    and g: "(g \<longlongrightarrow> m) F"
   shows "((\<lambda>x. dist (f x) (g x)) \<longlongrightarrow> dist l m) F"
 proof (rule tendstoI)
-  fix e :: real assume "0 < e"
-  hence e2: "0 < e/2" by simp
+  fix e :: real
+  assume "0 < e"
+  then have e2: "0 < e/2" by simp
   from tendstoD [OF f e2] tendstoD [OF g e2]
   show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) F"
   proof (eventually_elim)
@@ -516,9 +555,9 @@
     then show "dist (dist (f x) (g x)) (dist l m) < e"
       unfolding dist_real_def
       using dist_triangle2 [of "f x" "g x" "l"]
-      using dist_triangle2 [of "g x" "l" "m"]
-      using dist_triangle3 [of "l" "m" "f x"]
-      using dist_triangle [of "f x" "m" "g x"]
+        and dist_triangle2 [of "g x" "l" "m"]
+        and dist_triangle3 [of "l" "m" "f x"]
+        and dist_triangle [of "f x" "m" "g x"]
       by arith
   qed
 qed
@@ -533,33 +572,28 @@
   shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. dist (f x) (g x))"
   unfolding continuous_on_def by (auto intro: tendsto_dist)
 
-lemma tendsto_norm [tendsto_intros]:
-  "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. norm (f x)) \<longlongrightarrow> norm a) F"
+lemma tendsto_norm [tendsto_intros]: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. norm (f x)) \<longlongrightarrow> norm a) F"
   unfolding norm_conv_dist by (intro tendsto_intros)
 
-lemma continuous_norm [continuous_intros]:
-  "continuous F f \<Longrightarrow> continuous F (\<lambda>x. norm (f x))"
+lemma continuous_norm [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. norm (f x))"
   unfolding continuous_def by (rule tendsto_norm)
 
 lemma continuous_on_norm [continuous_intros]:
   "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. norm (f x))"
   unfolding continuous_on_def by (auto intro: tendsto_norm)
 
-lemma tendsto_norm_zero:
-  "(f \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) F"
-  by (drule tendsto_norm, simp)
-
-lemma tendsto_norm_zero_cancel:
-  "((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> 0) F"
+lemma tendsto_norm_zero: "(f \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) F"
+  by (drule tendsto_norm) simp
+
+lemma tendsto_norm_zero_cancel: "((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> 0) F"
   unfolding tendsto_iff dist_norm by simp
 
-lemma tendsto_norm_zero_iff:
-  "((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) F \<longleftrightarrow> (f \<longlongrightarrow> 0) F"
+lemma tendsto_norm_zero_iff: "((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) F \<longleftrightarrow> (f \<longlongrightarrow> 0) F"
   unfolding tendsto_iff dist_norm by simp
 
-lemma tendsto_rabs [tendsto_intros]:
-  "(f \<longlongrightarrow> (l::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> \<bar>l\<bar>) F"
-  by (fold real_norm_def, rule tendsto_norm)
+lemma tendsto_rabs [tendsto_intros]: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> \<bar>l\<bar>) F"
+  for l :: real
+  by (fold real_norm_def) (rule tendsto_norm)
 
 lemma continuous_rabs [continuous_intros]:
   "continuous F f \<Longrightarrow> continuous F (\<lambda>x. \<bar>f x :: real\<bar>)"
@@ -569,17 +603,15 @@
   "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. \<bar>f x :: real\<bar>)"
   unfolding real_norm_def[symmetric] by (rule continuous_on_norm)
 
-lemma tendsto_rabs_zero:
-  "(f \<longlongrightarrow> (0::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> 0) F"
-  by (fold real_norm_def, rule tendsto_norm_zero)
-
-lemma tendsto_rabs_zero_cancel:
-  "((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> (0::real)) F \<Longrightarrow> (f \<longlongrightarrow> 0) F"
-  by (fold real_norm_def, rule tendsto_norm_zero_cancel)
-
-lemma tendsto_rabs_zero_iff:
-  "((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> (0::real)) F \<longleftrightarrow> (f \<longlongrightarrow> 0) F"
-  by (fold real_norm_def, rule tendsto_norm_zero_iff)
+lemma tendsto_rabs_zero: "(f \<longlongrightarrow> (0::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> 0) F"
+  by (fold real_norm_def) (rule tendsto_norm_zero)
+
+lemma tendsto_rabs_zero_cancel: "((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> (0::real)) F \<Longrightarrow> (f \<longlongrightarrow> 0) F"
+  by (fold real_norm_def) (rule tendsto_norm_zero_cancel)
+
+lemma tendsto_rabs_zero_iff: "((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> (0::real)) F \<longleftrightarrow> (f \<longlongrightarrow> 0) F"
+  by (fold real_norm_def) (rule tendsto_norm_zero_iff)
+
 
 subsection \<open>Topological Monoid\<close>
 
@@ -606,17 +638,22 @@
 
 lemma tendsto_add_zero:
   fixes f g :: "_ \<Rightarrow> 'b::topological_monoid_add"
-  shows "\<lbrakk>(f \<longlongrightarrow> 0) F; (g \<longlongrightarrow> 0) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) \<longlongrightarrow> 0) F"
-  by (drule (1) tendsto_add, simp)
+  shows "(f \<longlongrightarrow> 0) F \<Longrightarrow> (g \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f x + g x) \<longlongrightarrow> 0) F"
+  by (drule (1) tendsto_add) simp
 
 lemma tendsto_setsum [tendsto_intros]:
   fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::topological_comm_monoid_add"
   assumes "\<And>i. i \<in> I \<Longrightarrow> (f i \<longlongrightarrow> a i) F"
   shows "((\<lambda>x. \<Sum>i\<in>I. f i x) \<longlongrightarrow> (\<Sum>i\<in>I. a i)) F"
 proof (cases "finite I")
-  assume "finite I" thus ?thesis using assms
-    by (induct, simp, simp add: tendsto_add)
-qed simp
+  case True
+  then show ?thesis
+    using assms by induct (simp_all add: tendsto_add)
+next
+  case False
+  then show ?thesis
+    by simp
+qed
 
 lemma continuous_setsum [continuous_intros]:
   fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::topological_comm_monoid_add"
@@ -629,10 +666,13 @@
   unfolding continuous_on_def by (auto intro: tendsto_setsum)
 
 instance nat :: topological_comm_monoid_add
-  proof qed (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal)
+  by standard
+    (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal)
 
 instance int :: topological_comm_monoid_add
-  proof qed (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal)
+  by standard
+    (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal)
+
 
 subsubsection \<open>Topological group\<close>
 
@@ -640,7 +680,7 @@
   assumes tendsto_uminus_nhds: "(uminus \<longlongrightarrow> - a) (nhds a)"
 begin
 
-lemma tendsto_minus [tendsto_intros]: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. - f x) \<longlongrightarrow> -a) F"
+lemma tendsto_minus [tendsto_intros]: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. - f x) \<longlongrightarrow> - a) F"
   by (rule filterlim_compose[OF tendsto_uminus_nhds])
 
 end
@@ -649,29 +689,26 @@
 
 instance topological_ab_group_add < topological_comm_monoid_add ..
 
-lemma continuous_minus [continuous_intros]:
-  fixes f :: "'a::t2_space \<Rightarrow> 'b::topological_group_add"
-  shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. - f x)"
+lemma continuous_minus [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. - f x)"
+  for f :: "'a::t2_space \<Rightarrow> 'b::topological_group_add"
   unfolding continuous_def by (rule tendsto_minus)
 
-lemma continuous_on_minus [continuous_intros]:
-  fixes f :: "_ \<Rightarrow> 'b::topological_group_add"
-  shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. - f x)"
+lemma continuous_on_minus [continuous_intros]: "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. - f x)"
+  for f :: "_ \<Rightarrow> 'b::topological_group_add"
   unfolding continuous_on_def by (auto intro: tendsto_minus)
 
-lemma tendsto_minus_cancel:
-  fixes a :: "'a::topological_group_add"
-  shows "((\<lambda>x. - f x) \<longlongrightarrow> - a) F \<Longrightarrow> (f \<longlongrightarrow> a) F"
-  by (drule tendsto_minus, simp)
+lemma tendsto_minus_cancel: "((\<lambda>x. - f x) \<longlongrightarrow> - a) F \<Longrightarrow> (f \<longlongrightarrow> a) F"
+  for a :: "'a::topological_group_add"
+  by (drule tendsto_minus) simp
 
 lemma tendsto_minus_cancel_left:
-    "(f \<longlongrightarrow> - (y::_::topological_group_add)) F \<longleftrightarrow> ((\<lambda>x. - f x) \<longlongrightarrow> y) F"
+  "(f \<longlongrightarrow> - (y::_::topological_group_add)) F \<longleftrightarrow> ((\<lambda>x. - f x) \<longlongrightarrow> y) F"
   using tendsto_minus_cancel[of f "- y" F]  tendsto_minus[of f "- y" F]
   by auto
 
 lemma tendsto_diff [tendsto_intros]:
   fixes a b :: "'a::topological_group_add"
-  shows "\<lbrakk>(f \<longlongrightarrow> a) F; (g \<longlongrightarrow> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x - g x) \<longlongrightarrow> a - b) F"
+  shows "(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> ((\<lambda>x. f x - g x) \<longlongrightarrow> a - b) F"
   using tendsto_add [of f a F "\<lambda>x. - g x" "- b"] by (simp add: tendsto_minus)
 
 lemma continuous_diff [continuous_intros]:
@@ -689,7 +726,8 @@
 
 instance real_normed_vector < topological_ab_group_add
 proof
-  fix a b :: 'a show "((\<lambda>x. fst x + snd x) \<longlongrightarrow> a + b) (nhds a \<times>\<^sub>F nhds b)"
+  fix a b :: 'a
+  show "((\<lambda>x. fst x + snd x) \<longlongrightarrow> a + b) (nhds a \<times>\<^sub>F nhds b)"
     unfolding tendsto_Zfun_iff add_diff_add
     using tendsto_fst[OF filterlim_ident, of "(a,b)"] tendsto_snd[OF filterlim_ident, of "(a,b)"]
     by (intro Zfun_add)
@@ -702,32 +740,28 @@
 
 lemmas real_tendsto_sandwich = tendsto_sandwich[where 'b=real]
 
+
 subsubsection \<open>Linear operators and multiplication\<close>
 
-lemma linear_times:
-  fixes c::"'a::real_algebra" shows "linear (\<lambda>x. c * x)"
+lemma linear_times: "linear (\<lambda>x. c * x)"
+  for c :: "'a::real_algebra"
   by (auto simp: linearI distrib_left)
 
-lemma (in bounded_linear) tendsto:
-  "(g \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f (g x)) \<longlongrightarrow> f a) F"
+lemma (in bounded_linear) tendsto: "(g \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f (g x)) \<longlongrightarrow> f a) F"
   by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun)
 
-lemma (in bounded_linear) continuous:
-  "continuous F g \<Longrightarrow> continuous F (\<lambda>x. f (g x))"
+lemma (in bounded_linear) continuous: "continuous F g \<Longrightarrow> continuous F (\<lambda>x. f (g x))"
   using tendsto[of g _ F] by (auto simp: continuous_def)
 
-lemma (in bounded_linear) continuous_on:
-  "continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f (g x))"
+lemma (in bounded_linear) continuous_on: "continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f (g x))"
   using tendsto[of g] by (auto simp: continuous_on_def)
 
-lemma (in bounded_linear) tendsto_zero:
-  "(g \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f (g x)) \<longlongrightarrow> 0) F"
-  by (drule tendsto, simp only: zero)
+lemma (in bounded_linear) tendsto_zero: "(g \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f (g x)) \<longlongrightarrow> 0) F"
+  by (drule tendsto) (simp only: zero)
 
 lemma (in bounded_bilinear) tendsto:
-  "\<lbrakk>(f \<longlongrightarrow> a) F; (g \<longlongrightarrow> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x ** g x) \<longlongrightarrow> a ** b) F"
-  by (simp only: tendsto_Zfun_iff prod_diff_prod
-                 Zfun_add Zfun Zfun_left Zfun_right)
+  "(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> ((\<lambda>x. f x ** g x) \<longlongrightarrow> a ** b) F"
+  by (simp only: tendsto_Zfun_iff prod_diff_prod Zfun_add Zfun Zfun_left Zfun_right)
 
 lemma (in bounded_bilinear) continuous:
   "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x ** g x)"
@@ -739,7 +773,7 @@
 
 lemma (in bounded_bilinear) tendsto_zero:
   assumes f: "(f \<longlongrightarrow> 0) F"
-  assumes g: "(g \<longlongrightarrow> 0) F"
+    and g: "(g \<longlongrightarrow> 0) F"
   shows "((\<lambda>x. f x ** g x) \<longlongrightarrow> 0) F"
   using tendsto [OF f g] by (simp add: zero_left)
 
@@ -760,15 +794,13 @@
 lemmas tendsto_mult [tendsto_intros] =
   bounded_bilinear.tendsto [OF bounded_bilinear_mult]
 
-lemma tendsto_mult_left:
-  fixes c::"'a::real_normed_algebra"
-  shows "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. c * (f x)) \<longlongrightarrow> c * l) F"
-by (rule tendsto_mult [OF tendsto_const])
-
-lemma tendsto_mult_right:
-  fixes c::"'a::real_normed_algebra"
-  shows "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. (f x) * c) \<longlongrightarrow> l * c) F"
-by (rule tendsto_mult [OF _ tendsto_const])
+lemma tendsto_mult_left: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. c * (f x)) \<longlongrightarrow> c * l) F"
+  for c :: "'a::real_normed_algebra"
+  by (rule tendsto_mult [OF tendsto_const])
+
+lemma tendsto_mult_right: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. (f x) * c) \<longlongrightarrow> l * c) F"
+  for c :: "'a::real_normed_algebra"
+  by (rule tendsto_mult [OF _ tendsto_const])
 
 lemmas continuous_of_real [continuous_intros] =
   bounded_linear.continuous [OF bounded_linear_of_real]
@@ -797,14 +829,12 @@
 lemmas tendsto_mult_right_zero =
   bounded_bilinear.tendsto_right_zero [OF bounded_bilinear_mult]
 
-lemma tendsto_power [tendsto_intros]:
-  fixes f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra}"
-  shows "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f x ^ n) \<longlongrightarrow> a ^ n) F"
+lemma tendsto_power [tendsto_intros]: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f x ^ n) \<longlongrightarrow> a ^ n) F"
+  for f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra}"
   by (induct n) (simp_all add: tendsto_mult)
 
-lemma continuous_power [continuous_intros]:
-  fixes f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
-  shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. (f x)^n)"
+lemma continuous_power [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. (f x)^n)"
+  for f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
   unfolding continuous_def by (rule tendsto_power)
 
 lemma continuous_on_power [continuous_intros]:
@@ -817,9 +847,13 @@
   assumes "\<And>i. i \<in> S \<Longrightarrow> (f i \<longlongrightarrow> L i) F"
   shows "((\<lambda>x. \<Prod>i\<in>S. f i x) \<longlongrightarrow> (\<Prod>i\<in>S. L i)) F"
 proof (cases "finite S")
-  assume "finite S" thus ?thesis using assms
-    by (induct, simp, simp add: tendsto_mult)
-qed simp
+  case True
+  then show ?thesis using assms
+    by induct (simp_all add: tendsto_mult)
+next
+  case False
+  then show ?thesis by simp
+qed
 
 lemma continuous_setprod [continuous_intros]:
   fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
@@ -832,20 +866,20 @@
   unfolding continuous_on_def by (auto intro: tendsto_setprod)
 
 lemma tendsto_of_real_iff:
-  "((\<lambda>x. of_real (f x) :: 'a :: real_normed_div_algebra) \<longlongrightarrow> of_real c) F \<longleftrightarrow> (f \<longlongrightarrow> c) F"
+  "((\<lambda>x. of_real (f x) :: 'a::real_normed_div_algebra) \<longlongrightarrow> of_real c) F \<longleftrightarrow> (f \<longlongrightarrow> c) F"
   unfolding tendsto_iff by simp
 
 lemma tendsto_add_const_iff:
-  "((\<lambda>x. c + f x :: 'a :: real_normed_vector) \<longlongrightarrow> c + d) F \<longleftrightarrow> (f \<longlongrightarrow> d) F"
+  "((\<lambda>x. c + f x :: 'a::real_normed_vector) \<longlongrightarrow> c + d) F \<longleftrightarrow> (f \<longlongrightarrow> d) F"
   using tendsto_add[OF tendsto_const[of c], of f d]
-        tendsto_add[OF tendsto_const[of "-c"], of "\<lambda>x. c + f x" "c + d"] by auto
+    and tendsto_add[OF tendsto_const[of "-c"], of "\<lambda>x. c + f x" "c + d"] by auto
 
 
 subsubsection \<open>Inverse and division\<close>
 
 lemma (in bounded_bilinear) Zfun_prod_Bfun:
   assumes f: "Zfun f F"
-  assumes g: "Bfun g F"
+    and g: "Bfun g F"
   shows "Zfun (\<lambda>x. f x ** g x) F"
 proof -
   obtain K where K: "0 \<le> K"
@@ -860,8 +894,7 @@
     have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K"
       by (rule norm_le)
     also have "\<dots> \<le> norm (f x) * B * K"
-      by (intro mult_mono' order_refl norm_g norm_ge_zero
-                mult_nonneg_nonneg K elim)
+      by (intro mult_mono' order_refl norm_g norm_ge_zero mult_nonneg_nonneg K elim)
     also have "\<dots> = norm (f x) * (B * K)"
       by (rule mult.assoc)
     finally show "norm (f x ** g x) \<le> norm (f x) * (B * K)" .
@@ -872,14 +905,15 @@
 
 lemma (in bounded_bilinear) Bfun_prod_Zfun:
   assumes f: "Bfun f F"
-  assumes g: "Zfun g F"
+    and g: "Zfun g F"
   shows "Zfun (\<lambda>x. f x ** g x) F"
   using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun)
 
 lemma Bfun_inverse_lemma:
   fixes x :: "'a::real_normed_div_algebra"
-  shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
-  apply (subst nonzero_norm_inverse, clarsimp)
+  shows "r \<le> norm x \<Longrightarrow> 0 < r \<Longrightarrow> norm (inverse x) \<le> inverse r"
+  apply (subst nonzero_norm_inverse)
+  apply clarsimp
   apply (erule (1) le_imp_inverse_le)
   done
 
@@ -890,38 +924,40 @@
   shows "Bfun (\<lambda>x. inverse (f x)) F"
 proof -
   from a have "0 < norm a" by simp
-  hence "\<exists>r>0. r < norm a" by (rule dense)
-  then obtain r where r1: "0 < r" and r2: "r < norm a" by blast
+  then have "\<exists>r>0. r < norm a" by (rule dense)
+  then obtain r where r1: "0 < r" and r2: "r < norm a"
+    by blast
   have "eventually (\<lambda>x. dist (f x) a < r) F"
     using tendstoD [OF f r1] by blast
-  hence "eventually (\<lambda>x. norm (inverse (f x)) \<le> inverse (norm a - r)) F"
+  then have "eventually (\<lambda>x. norm (inverse (f x)) \<le> inverse (norm a - r)) F"
   proof eventually_elim
     case (elim x)
-    hence 1: "norm (f x - a) < r"
+    then have 1: "norm (f x - a) < r"
       by (simp add: dist_norm)
-    hence 2: "f x \<noteq> 0" using r2 by auto
-    hence "norm (inverse (f x)) = inverse (norm (f x))"
+    then have 2: "f x \<noteq> 0" using r2 by auto
+    then have "norm (inverse (f x)) = inverse (norm (f x))"
       by (rule nonzero_norm_inverse)
     also have "\<dots> \<le> inverse (norm a - r)"
     proof (rule le_imp_inverse_le)
-      show "0 < norm a - r" using r2 by simp
-    next
+      show "0 < norm a - r"
+        using r2 by simp
       have "norm a - norm (f x) \<le> norm (a - f x)"
         by (rule norm_triangle_ineq2)
       also have "\<dots> = norm (f x - a)"
         by (rule norm_minus_commute)
       also have "\<dots> < r" using 1 .
-      finally show "norm a - r \<le> norm (f x)" by simp
+      finally show "norm a - r \<le> norm (f x)"
+        by simp
     qed
     finally show "norm (inverse (f x)) \<le> inverse (norm a - r)" .
   qed
-  thus ?thesis by (rule BfunI)
+  then show ?thesis by (rule BfunI)
 qed
 
 lemma tendsto_inverse [tendsto_intros]:
   fixes a :: "'a::real_normed_div_algebra"
   assumes f: "(f \<longlongrightarrow> a) F"
-  assumes a: "a \<noteq> 0"
+    and a: "a \<noteq> 0"
   shows "((\<lambda>x. inverse (f x)) \<longlongrightarrow> inverse a) F"
 proof -
   from a have "0 < norm a" by simp
@@ -942,43 +978,49 @@
 
 lemma continuous_inverse:
   fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
-  assumes "continuous F f" and "f (Lim F (\<lambda>x. x)) \<noteq> 0"
+  assumes "continuous F f"
+    and "f (Lim F (\<lambda>x. x)) \<noteq> 0"
   shows "continuous F (\<lambda>x. inverse (f x))"
   using assms unfolding continuous_def by (rule tendsto_inverse)
 
 lemma continuous_at_within_inverse[continuous_intros]:
   fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
-  assumes "continuous (at a within s) f" and "f a \<noteq> 0"
+  assumes "continuous (at a within s) f"
+    and "f a \<noteq> 0"
   shows "continuous (at a within s) (\<lambda>x. inverse (f x))"
   using assms unfolding continuous_within by (rule tendsto_inverse)
 
 lemma isCont_inverse[continuous_intros, simp]:
   fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
-  assumes "isCont f a" and "f a \<noteq> 0"
+  assumes "isCont f a"
+    and "f a \<noteq> 0"
   shows "isCont (\<lambda>x. inverse (f x)) a"
   using assms unfolding continuous_at by (rule tendsto_inverse)
 
 lemma continuous_on_inverse[continuous_intros]:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra"
-  assumes "continuous_on s f" and "\<forall>x\<in>s. f x \<noteq> 0"
+  assumes "continuous_on s f"
+    and "\<forall>x\<in>s. f x \<noteq> 0"
   shows "continuous_on s (\<lambda>x. inverse (f x))"
   using assms unfolding continuous_on_def by (blast intro: tendsto_inverse)
 
 lemma tendsto_divide [tendsto_intros]:
   fixes a b :: "'a::real_normed_field"
-  shows "\<lbrakk>(f \<longlongrightarrow> a) F; (g \<longlongrightarrow> b) F; b \<noteq> 0\<rbrakk>
-    \<Longrightarrow> ((\<lambda>x. f x / g x) \<longlongrightarrow> a / b) F"
+  shows "(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> ((\<lambda>x. f x / g x) \<longlongrightarrow> a / b) F"
   by (simp add: tendsto_mult tendsto_inverse divide_inverse)
 
 lemma continuous_divide:
   fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_field"
-  assumes "continuous F f" and "continuous F g" and "g (Lim F (\<lambda>x. x)) \<noteq> 0"
+  assumes "continuous F f"
+    and "continuous F g"
+    and "g (Lim F (\<lambda>x. x)) \<noteq> 0"
   shows "continuous F (\<lambda>x. (f x) / (g x))"
   using assms unfolding continuous_def by (rule tendsto_divide)
 
 lemma continuous_at_within_divide[continuous_intros]:
   fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_field"
-  assumes "continuous (at a within s) f" "continuous (at a within s) g" and "g a \<noteq> 0"
+  assumes "continuous (at a within s) f" "continuous (at a within s) g"
+    and "g a \<noteq> 0"
   shows "continuous (at a within s) (\<lambda>x. (f x) / (g x))"
   using assms unfolding continuous_within by (rule tendsto_divide)
 
@@ -990,36 +1032,40 @@
 
 lemma continuous_on_divide[continuous_intros]:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_field"
-  assumes "continuous_on s f" "continuous_on s g" and "\<forall>x\<in>s. g x \<noteq> 0"
+  assumes "continuous_on s f" "continuous_on s g"
+    and "\<forall>x\<in>s. g x \<noteq> 0"
   shows "continuous_on s (\<lambda>x. (f x) / (g x))"
   using assms unfolding continuous_on_def by (blast intro: tendsto_divide)
 
-lemma tendsto_sgn [tendsto_intros]:
-  fixes l :: "'a::real_normed_vector"
-  shows "\<lbrakk>(f \<longlongrightarrow> l) F; l \<noteq> 0\<rbrakk> \<Longrightarrow> ((\<lambda>x. sgn (f x)) \<longlongrightarrow> sgn l) F"
+lemma tendsto_sgn [tendsto_intros]: "(f \<longlongrightarrow> l) F \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> ((\<lambda>x. sgn (f x)) \<longlongrightarrow> sgn l) F"
+  for l :: "'a::real_normed_vector"
   unfolding sgn_div_norm by (simp add: tendsto_intros)
 
 lemma continuous_sgn:
   fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
-  assumes "continuous F f" and "f (Lim F (\<lambda>x. x)) \<noteq> 0"
+  assumes "continuous F f"
+    and "f (Lim F (\<lambda>x. x)) \<noteq> 0"
   shows "continuous F (\<lambda>x. sgn (f x))"
   using assms unfolding continuous_def by (rule tendsto_sgn)
 
 lemma continuous_at_within_sgn[continuous_intros]:
   fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
-  assumes "continuous (at a within s) f" and "f a \<noteq> 0"
+  assumes "continuous (at a within s) f"
+    and "f a \<noteq> 0"
   shows "continuous (at a within s) (\<lambda>x. sgn (f x))"
   using assms unfolding continuous_within by (rule tendsto_sgn)
 
 lemma isCont_sgn[continuous_intros]:
   fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
-  assumes "isCont f a" and "f a \<noteq> 0"
+  assumes "isCont f a"
+    and "f a \<noteq> 0"
   shows "isCont (\<lambda>x. sgn (f x)) a"
   using assms unfolding continuous_at by (rule tendsto_sgn)
 
 lemma continuous_on_sgn[continuous_intros]:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
-  assumes "continuous_on s f" and "\<forall>x\<in>s. f x \<noteq> 0"
+  assumes "continuous_on s f"
+    and "\<forall>x\<in>s. f x \<noteq> 0"
   shows "continuous_on s (\<lambda>x. sgn (f x))"
   using assms unfolding continuous_on_def by (blast intro: tendsto_sgn)
 
@@ -1029,35 +1075,40 @@
   shows "(LIM x F. f x :> at_infinity) \<longleftrightarrow> (\<forall>r>c. eventually (\<lambda>x. r \<le> norm (f x)) F)"
   unfolding filterlim_iff eventually_at_infinity
 proof safe
-  fix P :: "'a \<Rightarrow> bool" and b
+  fix P :: "'a \<Rightarrow> bool"
+  fix b
   assume *: "\<forall>r>c. eventually (\<lambda>x. r \<le> norm (f x)) F"
-    and P: "\<forall>x. b \<le> norm x \<longrightarrow> P x"
+  assume P: "\<forall>x. b \<le> norm x \<longrightarrow> P x"
   have "max b (c + 1) > c" by auto
   with * have "eventually (\<lambda>x. max b (c + 1) \<le> norm (f x)) F"
     by auto
   then show "eventually (\<lambda>x. P (f x)) F"
   proof eventually_elim
-    fix x assume "max b (c + 1) \<le> norm (f x)"
+    case (elim x)
     with P show "P (f x)" by auto
   qed
 qed force
 
 lemma not_tendsto_and_filterlim_at_infinity:
+  fixes c :: "'a::real_normed_vector"
   assumes "F \<noteq> bot"
-  assumes "(f \<longlongrightarrow> (c :: 'a :: real_normed_vector)) F"
-  assumes "filterlim f at_infinity F"
-  shows   False
+    and "(f \<longlongrightarrow> c) F"
+    and "filterlim f at_infinity F"
+  shows False
 proof -
   from tendstoD[OF assms(2), of "1/2"]
-    have "eventually (\<lambda>x. dist (f x) c < 1/2) F" by simp
-  moreover from filterlim_at_infinity[of "norm c" f F] assms(3)
-    have "eventually (\<lambda>x. norm (f x) \<ge> norm c + 1) F" by simp
+  have "eventually (\<lambda>x. dist (f x) c < 1/2) F"
+    by simp
+  moreover
+  from filterlim_at_infinity[of "norm c" f F] assms(3)
+  have "eventually (\<lambda>x. norm (f x) \<ge> norm c + 1) F" by simp
   ultimately have "eventually (\<lambda>x. False) F"
   proof eventually_elim
-    fix x assume A: "dist (f x) c < 1/2" and B: "norm (f x) \<ge> norm c + 1"
-    note B
+    fix x
+    assume A: "dist (f x) c < 1/2"
+    assume "norm (f x) \<ge> norm c + 1"
     also have "norm (f x) = dist (f x) 0" by simp
-    also have "... \<le> dist (f x) c + dist c 0" by (rule dist_triangle)
+    also have "\<dots> \<le> dist (f x) c + dist c 0" by (rule dist_triangle)
     finally show False using A by simp
   qed
   with assms show False by simp
@@ -1065,83 +1116,97 @@
 
 lemma filterlim_at_infinity_imp_not_convergent:
   assumes "filterlim f at_infinity sequentially"
-  shows   "\<not>convergent f"
+  shows "\<not> convergent f"
   by (rule notI, rule not_tendsto_and_filterlim_at_infinity[OF _ _ assms])
      (simp_all add: convergent_LIMSEQ_iff)
 
 lemma filterlim_at_infinity_imp_eventually_ne:
   assumes "filterlim f at_infinity F"
-  shows   "eventually (\<lambda>z. f z \<noteq> c) F"
+  shows "eventually (\<lambda>z. f z \<noteq> c) F"
 proof -
-  have "norm c + 1 > 0" by (intro add_nonneg_pos) simp_all
+  have "norm c + 1 > 0"
+    by (intro add_nonneg_pos) simp_all
   with filterlim_at_infinity[OF order.refl, of f F] assms
-    have "eventually (\<lambda>z. norm (f z) \<ge> norm c + 1) F" by blast
-  thus ?thesis by eventually_elim auto
+  have "eventually (\<lambda>z. norm (f z) \<ge> norm c + 1) F"
+    by blast
+  then show ?thesis
+    by eventually_elim auto
 qed
 
 lemma tendsto_of_nat [tendsto_intros]:
-  "filterlim (of_nat :: nat \<Rightarrow> 'a :: real_normed_algebra_1) at_infinity sequentially"
+  "filterlim (of_nat :: nat \<Rightarrow> 'a::real_normed_algebra_1) at_infinity sequentially"
 proof (subst filterlim_at_infinity[OF order.refl], intro allI impI)
   fix r :: real
   assume r: "r > 0"
   define n where "n = nat \<lceil>r\<rceil>"
-  from r have n: "\<forall>m\<ge>n. of_nat m \<ge> r" unfolding n_def by linarith
+  from r have n: "\<forall>m\<ge>n. of_nat m \<ge> r"
+    unfolding n_def by linarith
   from eventually_ge_at_top[of n] show "eventually (\<lambda>m. norm (of_nat m :: 'a) \<ge> r) sequentially"
-    by eventually_elim (insert n, simp_all)
+    by eventually_elim (use n in simp_all)
 qed
 
 
 subsection \<open>Relate @{const at}, @{const at_left} and @{const at_right}\<close>
 
 text \<open>
-
-This lemmas are useful for conversion between @{term "at x"} to @{term "at_left x"} and
-@{term "at_right x"} and also @{term "at_right 0"}.
-
+  This lemmas are useful for conversion between @{term "at x"} to @{term "at_left x"} and
+  @{term "at_right x"} and also @{term "at_right 0"}.
 \<close>
 
 lemmas filterlim_split_at_real = filterlim_split_at[where 'a=real]
 
-lemma filtermap_nhds_shift: "filtermap (\<lambda>x. x - d) (nhds a) = nhds (a - d::'a::real_normed_vector)"
+lemma filtermap_nhds_shift: "filtermap (\<lambda>x. x - d) (nhds a) = nhds (a - d)"
+  for a d :: "'a::real_normed_vector"
   by (rule filtermap_fun_inverse[where g="\<lambda>x. x + d"])
-     (auto intro!: tendsto_eq_intros filterlim_ident)
-
-lemma filtermap_nhds_minus: "filtermap (\<lambda>x. - x) (nhds a) = nhds (- a::'a::real_normed_vector)"
+    (auto intro!: tendsto_eq_intros filterlim_ident)
+
+lemma filtermap_nhds_minus: "filtermap (\<lambda>x. - x) (nhds a) = nhds (- a)"
+  for a :: "'a::real_normed_vector"
   by (rule filtermap_fun_inverse[where g=uminus])
-     (auto intro!: tendsto_eq_intros filterlim_ident)
-
-lemma filtermap_at_shift: "filtermap (\<lambda>x. x - d) (at a) = at (a - d::'a::real_normed_vector)"
+    (auto intro!: tendsto_eq_intros filterlim_ident)
+
+lemma filtermap_at_shift: "filtermap (\<lambda>x. x - d) (at a) = at (a - d)"
+  for a d :: "'a::real_normed_vector"
   by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric])
 
-lemma filtermap_at_right_shift: "filtermap (\<lambda>x. x - d) (at_right a) = at_right (a - d::real)"
+lemma filtermap_at_right_shift: "filtermap (\<lambda>x. x - d) (at_right a) = at_right (a - d)"
+  for a d :: "real"
   by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric])
 
-lemma at_right_to_0: "at_right (a::real) = filtermap (\<lambda>x. x + a) (at_right 0)"
+lemma at_right_to_0: "at_right a = filtermap (\<lambda>x. x + a) (at_right 0)"
+  for a :: real
   using filtermap_at_right_shift[of "-a" 0] by simp
 
 lemma filterlim_at_right_to_0:
-  "filterlim f F (at_right (a::real)) \<longleftrightarrow> filterlim (\<lambda>x. f (x + a)) F (at_right 0)"
+  "filterlim f F (at_right a) \<longleftrightarrow> filterlim (\<lambda>x. f (x + a)) F (at_right 0)"
+  for a :: real
   unfolding filterlim_def filtermap_filtermap at_right_to_0[of a] ..
 
 lemma eventually_at_right_to_0:
-  "eventually P (at_right (a::real)) \<longleftrightarrow> eventually (\<lambda>x. P (x + a)) (at_right 0)"
+  "eventually P (at_right a) \<longleftrightarrow> eventually (\<lambda>x. P (x + a)) (at_right 0)"
+  for a :: real
   unfolding at_right_to_0[of a] by (simp add: eventually_filtermap)
 
-lemma filtermap_at_minus: "filtermap (\<lambda>x. - x) (at a) = at (- a::'a::real_normed_vector)"
+lemma filtermap_at_minus: "filtermap (\<lambda>x. - x) (at a) = at (- a)"
+  for a :: "'a::real_normed_vector"
   by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric])
 
-lemma at_left_minus: "at_left (a::real) = filtermap (\<lambda>x. - x) (at_right (- a))"
+lemma at_left_minus: "at_left a = filtermap (\<lambda>x. - x) (at_right (- a))"
+  for a :: real
   by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric])
 
-lemma at_right_minus: "at_right (a::real) = filtermap (\<lambda>x. - x) (at_left (- a))"
+lemma at_right_minus: "at_right a = filtermap (\<lambda>x. - x) (at_left (- a))"
+  for a :: real
   by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric])
 
 lemma filterlim_at_left_to_right:
-  "filterlim f F (at_left (a::real)) \<longleftrightarrow> filterlim (\<lambda>x. f (- x)) F (at_right (-a))"
+  "filterlim f F (at_left a) \<longleftrightarrow> filterlim (\<lambda>x. f (- x)) F (at_right (-a))"
+  for a :: real
   unfolding filterlim_def filtermap_filtermap at_left_minus[of a] ..
 
 lemma eventually_at_left_to_right:
-  "eventually P (at_left (a::real)) \<longleftrightarrow> eventually (\<lambda>x. P (- x)) (at_right (-a))"
+  "eventually P (at_left a) \<longleftrightarrow> eventually (\<lambda>x. P (- x)) (at_right (-a))"
+  for a :: real
   unfolding at_left_minus[of a] by (simp add: eventually_filtermap)
 
 lemma filterlim_uminus_at_top_at_bot: "LIM x at_bot. - x :: real :> at_top"
@@ -1167,7 +1232,7 @@
 
 lemma filterlim_uminus_at_top: "(LIM x F. f x :> at_top) \<longleftrightarrow> (LIM x F. - (f x) :: real :> at_bot)"
   using filterlim_compose[OF filterlim_uminus_at_bot_at_top, of f F]
-  using filterlim_compose[OF filterlim_uminus_at_top_at_bot, of "\<lambda>x. - f x" F]
+    and filterlim_compose[OF filterlim_uminus_at_top_at_bot, of "\<lambda>x. - f x" F]
   by auto
 
 lemma filterlim_uminus_at_bot: "(LIM x F. f x :> at_bot) \<longleftrightarrow> (LIM x F. - (f x) :: real :> at_top)"
@@ -1176,7 +1241,8 @@
 lemma filterlim_inverse_at_top_right: "LIM x at_right (0::real). inverse x :> at_top"
   unfolding filterlim_at_top_gt[where c=0] eventually_at_filter
 proof safe
-  fix Z :: real assume [arith]: "0 < Z"
+  fix Z :: real
+  assume [arith]: "0 < Z"
   then have "eventually (\<lambda>x. x < inverse Z) (nhds 0)"
     by (auto simp add: eventually_nhds_metric dist_real_def intro!: exI[of _ "\<bar>inverse Z\<bar>"])
   then show "eventually (\<lambda>x. x \<noteq> 0 \<longrightarrow> x \<in> {0<..} \<longrightarrow> Z \<le> inverse x) (nhds 0)"
@@ -1188,41 +1254,56 @@
   shows "(inverse \<longlongrightarrow> (0::'a)) at_infinity"
   unfolding tendsto_Zfun_iff diff_0_right Zfun_def eventually_at_infinity
 proof safe
-  fix r :: real assume "0 < r"
+  fix r :: real
+  assume "0 < r"
   show "\<exists>b. \<forall>x. b \<le> norm x \<longrightarrow> norm (inverse x :: 'a) < r"
   proof (intro exI[of _ "inverse (r / 2)"] allI impI)
     fix x :: 'a
     from \<open>0 < r\<close> have "0 < inverse (r / 2)" by simp
     also assume *: "inverse (r / 2) \<le> norm x"
     finally show "norm (inverse x) < r"
-      using * \<open>0 < r\<close> by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps)
+      using * \<open>0 < r\<close>
+      by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps)
   qed
 qed
 
 lemma tendsto_add_filterlim_at_infinity:
-  assumes "(f \<longlongrightarrow> (c :: 'b :: real_normed_vector)) (F :: 'a filter)"
-  assumes "filterlim g at_infinity F"
-  shows   "filterlim (\<lambda>x. f x + g x) at_infinity F"
+  fixes c :: "'b::real_normed_vector"
+    and F :: "'a filter"
+  assumes "(f \<longlongrightarrow> c) F"
+    and "filterlim g at_infinity F"
+  shows "filterlim (\<lambda>x. f x + g x) at_infinity F"
 proof (subst filterlim_at_infinity[OF order_refl], safe)
-  fix r :: real assume r: "r > 0"
-  from assms(1) have "((\<lambda>x. norm (f x)) \<longlongrightarrow> norm c) F" by (rule tendsto_norm)
-  hence "eventually (\<lambda>x. norm (f x) < norm c + 1) F" by (rule order_tendstoD) simp_all
-  moreover from r have "r + norm c + 1 > 0" by (intro add_pos_nonneg) simp_all
+  fix r :: real
+  assume r: "r > 0"
+  from assms(1) have "((\<lambda>x. norm (f x)) \<longlongrightarrow> norm c) F"
+    by (rule tendsto_norm)
+  then have "eventually (\<lambda>x. norm (f x) < norm c + 1) F"
+    by (rule order_tendstoD) simp_all
+  moreover from r have "r + norm c + 1 > 0"
+    by (intro add_pos_nonneg) simp_all
   with assms(2) have "eventually (\<lambda>x. norm (g x) \<ge> r + norm c + 1) F"
-    unfolding filterlim_at_infinity[OF order_refl] by (elim allE[of _ "r + norm c + 1"]) simp_all
+    unfolding filterlim_at_infinity[OF order_refl]
+    by (elim allE[of _ "r + norm c + 1"]) simp_all
   ultimately show "eventually (\<lambda>x. norm (f x + g x) \<ge> r) F"
   proof eventually_elim
-    fix x :: 'a assume A: "norm (f x) < norm c + 1" and B: "r + norm c + 1 \<le> norm (g x)"
-    from A B have "r \<le> norm (g x) - norm (f x)" by simp
-    also have "norm (g x) - norm (f x) \<le> norm (g x + f x)" by (rule norm_diff_ineq)
-    finally show "r \<le> norm (f x + g x)" by (simp add: add_ac)
+    fix x :: 'a
+    assume A: "norm (f x) < norm c + 1" and B: "r + norm c + 1 \<le> norm (g x)"
+    from A B have "r \<le> norm (g x) - norm (f x)"
+      by simp
+    also have "norm (g x) - norm (f x) \<le> norm (g x + f x)"
+      by (rule norm_diff_ineq)
+    finally show "r \<le> norm (f x + g x)"
+      by (simp add: add_ac)
   qed
 qed
 
 lemma tendsto_add_filterlim_at_infinity':
+  fixes c :: "'b::real_normed_vector"
+    and F :: "'a filter"
   assumes "filterlim f at_infinity F"
-  assumes "(g \<longlongrightarrow> (c :: 'b :: real_normed_vector)) (F :: 'a filter)"
-  shows   "filterlim (\<lambda>x. f x + g x) at_infinity F"
+    and "(g \<longlongrightarrow> c) F"
+  shows "filterlim (\<lambda>x. f x + g x) at_infinity F"
   by (subst add.commute) (rule tendsto_add_filterlim_at_infinity assms)+
 
 lemma filterlim_inverse_at_right_top: "LIM x at_top. inverse x :> at_right (0::real)"
@@ -1272,7 +1353,8 @@
   shows "filterlim inverse at_infinity (at (0::'a))"
   unfolding filterlim_at_infinity[OF order_refl]
 proof safe
-  fix r :: real assume "0 < r"
+  fix r :: real
+  assume "0 < r"
   then show "eventually (\<lambda>x::'a. r \<le> norm (inverse x)) (at 0)"
     unfolding eventually_at norm_inverse
     by (intro exI[of _ "inverse r"])
@@ -1290,7 +1372,7 @@
   also have "\<dots> \<le> at 0"
     using tendsto_inverse_0[where 'a='b]
     by (auto intro!: exI[of _ 1]
-             simp: le_principal eventually_filtermap filterlim_def at_within_def eventually_at_infinity)
+        simp: le_principal eventually_filtermap filterlim_def at_within_def eventually_at_infinity)
   finally show "filtermap inverse (filtermap g F) \<le> at 0" .
 next
   assume "filtermap inverse (filtermap g F) \<le> at 0"
@@ -1301,36 +1383,40 @@
 qed
 
 lemma tendsto_mult_filterlim_at_infinity:
-  assumes "F \<noteq> bot" "(f \<longlongrightarrow> (c :: 'a :: real_normed_field)) F" "c \<noteq> 0"
+  fixes c :: "'a::real_normed_field"
+  assumes "F \<noteq> bot" "(f \<longlongrightarrow> c) F" "c \<noteq> 0"
   assumes "filterlim g at_infinity F"
-  shows   "filterlim (\<lambda>x. f x * g x) at_infinity F"
+  shows "filterlim (\<lambda>x. f x * g x) at_infinity F"
 proof -
   have "((\<lambda>x. inverse (f x) * inverse (g x)) \<longlongrightarrow> inverse c * 0) F"
     by (intro tendsto_mult tendsto_inverse assms filterlim_compose[OF tendsto_inverse_0])
-  hence "filterlim (\<lambda>x. inverse (f x) * inverse (g x)) (at (inverse c * 0)) F"
-    unfolding filterlim_at using assms
+  then have "filterlim (\<lambda>x. inverse (f x) * inverse (g x)) (at (inverse c * 0)) F"
+    unfolding filterlim_at
+    using assms
     by (auto intro: filterlim_at_infinity_imp_eventually_ne tendsto_imp_eventually_ne eventually_conj)
-  thus ?thesis by (subst filterlim_inverse_at_iff[symmetric]) simp_all
+  then show ?thesis
+    by (subst filterlim_inverse_at_iff[symmetric]) simp_all
 qed
 
 lemma tendsto_inverse_0_at_top: "LIM x F. f x :> at_top \<Longrightarrow> ((\<lambda>x. inverse (f x) :: real) \<longlongrightarrow> 0) F"
  by (metis filterlim_at filterlim_mono[OF _ at_top_le_at_infinity order_refl] filterlim_inverse_at_iff)
 
-lemma mult_nat_left_at_top: "c > 0 \<Longrightarrow> filterlim (\<lambda>x. c * x :: nat) at_top sequentially"
-  by (rule filterlim_subseq) (auto simp: subseq_def)
-
-lemma mult_nat_right_at_top: "c > 0 \<Longrightarrow> filterlim (\<lambda>x. x * c :: nat) at_top sequentially"
+lemma mult_nat_left_at_top: "c > 0 \<Longrightarrow> filterlim (\<lambda>x. c * x) at_top sequentially"
+  for c :: nat
   by (rule filterlim_subseq) (auto simp: subseq_def)
 
-lemma at_to_infinity:
-  fixes x :: "'a :: {real_normed_field,field}"
-  shows "(at (0::'a)) = filtermap inverse at_infinity"
+lemma mult_nat_right_at_top: "c > 0 \<Longrightarrow> filterlim (\<lambda>x. x * c) at_top sequentially"
+  for c :: nat
+  by (rule filterlim_subseq) (auto simp: subseq_def)
+
+lemma at_to_infinity: "(at (0::'a::{real_normed_field,field})) = filtermap inverse at_infinity"
 proof (rule antisym)
   have "(inverse \<longlongrightarrow> (0::'a)) at_infinity"
     by (fact tendsto_inverse_0)
   then show "filtermap inverse at_infinity \<le> at (0::'a)"
     apply (simp add: le_principal eventually_filtermap eventually_at_infinity filterlim_def at_within_def)
-    apply (rule_tac x="1" in exI, auto)
+    apply (rule_tac x="1" in exI)
+    apply auto
     done
 next
   have "filtermap inverse (filtermap inverse (at (0::'a))) \<le> filtermap inverse at_infinity"
@@ -1341,38 +1427,39 @@
 qed
 
 lemma lim_at_infinity_0:
-  fixes l :: "'a :: {real_normed_field,field}"
-  shows "(f \<longlongrightarrow> l) at_infinity \<longleftrightarrow> ((f o inverse) \<longlongrightarrow> l) (at (0::'a))"
-by (simp add: tendsto_compose_filtermap at_to_infinity filtermap_filtermap)
+  fixes l :: "'a::{real_normed_field,field}"
+  shows "(f \<longlongrightarrow> l) at_infinity \<longleftrightarrow> ((f \<circ> inverse) \<longlongrightarrow> l) (at (0::'a))"
+  by (simp add: tendsto_compose_filtermap at_to_infinity filtermap_filtermap)
 
 lemma lim_zero_infinity:
-  fixes l :: "'a :: {real_normed_field,field}"
+  fixes l :: "'a::{real_normed_field,field}"
   shows "((\<lambda>x. f(1 / x)) \<longlongrightarrow> l) (at (0::'a)) \<Longrightarrow> (f \<longlongrightarrow> l) at_infinity"
-by (simp add: inverse_eq_divide lim_at_infinity_0 comp_def)
+  by (simp add: inverse_eq_divide lim_at_infinity_0 comp_def)
 
 
 text \<open>
-
-We only show rules for multiplication and addition when the functions are either against a real
-value or against infinity. Further rules are easy to derive by using @{thm filterlim_uminus_at_top}.
-
+  We only show rules for multiplication and addition when the functions are either against a real
+  value or against infinity. Further rules are easy to derive by using @{thm
+  filterlim_uminus_at_top}.
 \<close>
 
 lemma filterlim_tendsto_pos_mult_at_top:
-  assumes f: "(f \<longlongrightarrow> c) F" and c: "0 < c"
-  assumes g: "LIM x F. g x :> at_top"
+  assumes f: "(f \<longlongrightarrow> c) F"
+    and c: "0 < c"
+    and g: "LIM x F. g x :> at_top"
   shows "LIM x F. (f x * g x :: real) :> at_top"
   unfolding filterlim_at_top_gt[where c=0]
 proof safe
-  fix Z :: real assume "0 < Z"
+  fix Z :: real
+  assume "0 < Z"
   from f \<open>0 < c\<close> have "eventually (\<lambda>x. c / 2 < f x) F"
     by (auto dest!: tendstoD[where e="c / 2"] elim!: eventually_mono
-             simp: dist_real_def abs_real_def split: if_split_asm)
+        simp: dist_real_def abs_real_def split: if_split_asm)
   moreover from g have "eventually (\<lambda>x. (Z / c * 2) \<le> g x) F"
     unfolding filterlim_at_top by auto
   ultimately show "eventually (\<lambda>x. Z \<le> f x * g x) F"
   proof eventually_elim
-    fix x assume "c / 2 < f x" "Z / c * 2 \<le> g x"
+    case (elim x)
     with \<open>0 < Z\<close> \<open>0 < c\<close> have "c / 2 * (Z / c * 2) \<le> f x * g x"
       by (intro mult_mono) (auto simp: zero_le_divide_iff)
     with \<open>0 < c\<close> show "Z \<le> f x * g x"
@@ -1382,18 +1469,19 @@
 
 lemma filterlim_at_top_mult_at_top:
   assumes f: "LIM x F. f x :> at_top"
-  assumes g: "LIM x F. g x :> at_top"
+    and g: "LIM x F. g x :> at_top"
   shows "LIM x F. (f x * g x :: real) :> at_top"
   unfolding filterlim_at_top_gt[where c=0]
 proof safe
-  fix Z :: real assume "0 < Z"
+  fix Z :: real
+  assume "0 < Z"
   from f have "eventually (\<lambda>x. 1 \<le> f x) F"
     unfolding filterlim_at_top by auto
   moreover from g have "eventually (\<lambda>x. Z \<le> g x) F"
     unfolding filterlim_at_top by auto
   ultimately show "eventually (\<lambda>x. Z \<le> f x * g x) F"
   proof eventually_elim
-    fix x assume "1 \<le> f x" "Z \<le> g x"
+    case (elim x)
     with \<open>0 < Z\<close> have "1 * Z \<le> f x * g x"
       by (intro mult_mono) (auto simp: zero_le_divide_iff)
     then show "Z \<le> f x * g x"
@@ -1402,25 +1490,32 @@
 qed
 
 lemma filterlim_tendsto_pos_mult_at_bot:
-  assumes "(f \<longlongrightarrow> c) F" "0 < (c::real)" "filterlim g at_bot F"
+  fixes c :: real
+  assumes "(f \<longlongrightarrow> c) F" "0 < c" "filterlim g at_bot F"
   shows "LIM x F. f x * g x :> at_bot"
   using filterlim_tendsto_pos_mult_at_top[OF assms(1,2), of "\<lambda>x. - g x"] assms(3)
   unfolding filterlim_uminus_at_bot by simp
 
 lemma filterlim_tendsto_neg_mult_at_bot:
-  assumes c: "(f \<longlongrightarrow> c) F" "(c::real) < 0" and g: "filterlim g at_top F"
+  fixes c :: real
+  assumes c: "(f \<longlongrightarrow> c) F" "c < 0" and g: "filterlim g at_top F"
   shows "LIM x F. f x * g x :> at_bot"
   using c filterlim_tendsto_pos_mult_at_top[of "\<lambda>x. - f x" "- c" F, OF _ _ g]
   unfolding filterlim_uminus_at_bot tendsto_minus_cancel_left by simp
 
 lemma filterlim_pow_at_top:
   fixes f :: "real \<Rightarrow> real"
-  assumes "0 < n" and f: "LIM x F. f x :> at_top"
+  assumes "0 < n"
+    and f: "LIM x F. f x :> at_top"
   shows "LIM x F. (f x)^n :: real :> at_top"
-using \<open>0 < n\<close> proof (induct n)
+  using \<open>0 < n\<close>
+proof (induct n)
+  case 0
+  then show ?case by simp
+next
   case (Suc n) with f show ?case
     by (cases "n = 0") (auto intro!: filterlim_at_top_mult_at_top)
-qed simp
+qed
 
 lemma filterlim_pow_at_bot_even:
   fixes f :: "real \<Rightarrow> real"
@@ -1434,11 +1529,12 @@
 
 lemma filterlim_tendsto_add_at_top:
   assumes f: "(f \<longlongrightarrow> c) F"
-  assumes g: "LIM x F. g x :> at_top"
+    and g: "LIM x F. g x :> at_top"
   shows "LIM x F. (f x + g x :: real) :> at_top"
   unfolding filterlim_at_top_gt[where c=0]
 proof safe
-  fix Z :: real assume "0 < Z"
+  fix Z :: real
+  assume "0 < Z"
   from f have "eventually (\<lambda>x. c - 1 < f x) F"
     by (auto dest!: tendstoD[where e=1] elim!: eventually_mono simp: dist_real_def)
   moreover from g have "eventually (\<lambda>x. Z - (c - 1) \<le> g x) F"
@@ -1450,18 +1546,19 @@
 lemma LIM_at_top_divide:
   fixes f g :: "'a \<Rightarrow> real"
   assumes f: "(f \<longlongrightarrow> a) F" "0 < a"
-  assumes g: "(g \<longlongrightarrow> 0) F" "eventually (\<lambda>x. 0 < g x) F"
+    and g: "(g \<longlongrightarrow> 0) F" "eventually (\<lambda>x. 0 < g x) F"
   shows "LIM x F. f x / g x :> at_top"
   unfolding divide_inverse
   by (rule filterlim_tendsto_pos_mult_at_top[OF f]) (rule filterlim_inverse_at_top[OF g])
 
 lemma filterlim_at_top_add_at_top:
   assumes f: "LIM x F. f x :> at_top"
-  assumes g: "LIM x F. g x :> at_top"
+    and g: "LIM x F. g x :> at_top"
   shows "LIM x F. (f x + g x :: real) :> at_top"
   unfolding filterlim_at_top_gt[where c=0]
 proof safe
-  fix Z :: real assume "0 < Z"
+  fix Z :: real
+  assume "0 < Z"
   from f have "eventually (\<lambda>x. 0 \<le> f x) F"
     unfolding filterlim_at_top by auto
   moreover from g have "eventually (\<lambda>x. Z \<le> g x) F"
@@ -1473,34 +1570,43 @@
 lemma tendsto_divide_0:
   fixes f :: "_ \<Rightarrow> 'a::{real_normed_div_algebra, division_ring}"
   assumes f: "(f \<longlongrightarrow> c) F"
-  assumes g: "LIM x F. g x :> at_infinity"
+    and g: "LIM x F. g x :> at_infinity"
   shows "((\<lambda>x. f x / g x) \<longlongrightarrow> 0) F"
-  using tendsto_mult[OF f filterlim_compose[OF tendsto_inverse_0 g]] by (simp add: divide_inverse)
+  using tendsto_mult[OF f filterlim_compose[OF tendsto_inverse_0 g]]
+  by (simp add: divide_inverse)
 
 lemma linear_plus_1_le_power:
   fixes x :: real
   assumes x: "0 \<le> x"
   shows "real n * x + 1 \<le> (x + 1) ^ n"
 proof (induct n)
+  case 0
+  then show ?case by simp
+next
   case (Suc n)
-  have "real (Suc n) * x + 1 \<le> (x + 1) * (real n * x + 1)"
-    by (simp add: field_simps of_nat_Suc x)
+  from x have "real (Suc n) * x + 1 \<le> (x + 1) * (real n * x + 1)"
+    by (simp add: field_simps)
   also have "\<dots> \<le> (x + 1)^Suc n"
     using Suc x by (simp add: mult_left_mono)
   finally show ?case .
-qed simp
+qed
 
 lemma filterlim_realpow_sequentially_gt1:
   fixes x :: "'a :: real_normed_div_algebra"
   assumes x[arith]: "1 < norm x"
   shows "LIM n sequentially. x ^ n :> at_infinity"
 proof (intro filterlim_at_infinity[THEN iffD2] allI impI)
-  fix y :: real assume "0 < y"
+  fix y :: real
+  assume "0 < y"
   have "0 < norm x - 1" by simp
-  then obtain N::nat where "y < real N * (norm x - 1)" by (blast dest: reals_Archimedean3)
-  also have "\<dots> \<le> real N * (norm x - 1) + 1" by simp
-  also have "\<dots> \<le> (norm x - 1 + 1) ^ N" by (rule linear_plus_1_le_power) simp
-  also have "\<dots> = norm x ^ N" by simp
+  then obtain N :: nat where "y < real N * (norm x - 1)"
+    by (blast dest: reals_Archimedean3)
+  also have "\<dots> \<le> real N * (norm x - 1) + 1"
+    by simp
+  also have "\<dots> \<le> (norm x - 1 + 1) ^ N"
+    by (rule linear_plus_1_le_power) simp
+  also have "\<dots> = norm x ^ N"
+    by simp
   finally have "\<forall>n\<ge>N. y \<le> norm x ^ n"
     by (metis order_less_le_trans power_increasing order_less_imp_le x)
   then show "eventually (\<lambda>n. y \<le> norm (x ^ n)) sequentially"
@@ -1512,48 +1618,48 @@
 subsection \<open>Floor and Ceiling\<close>
 
 lemma eventually_floor_less:
-  fixes f::"'a \<Rightarrow> 'b::{order_topology, floor_ceiling}"
+  fixes f :: "'a \<Rightarrow> 'b::{order_topology,floor_ceiling}"
   assumes f: "(f \<longlongrightarrow> l) F"
-  assumes l: "l \<notin> \<int>"
+    and l: "l \<notin> \<int>"
   shows "\<forall>\<^sub>F x in F. of_int (floor l) < f x"
   by (intro order_tendstoD[OF f]) (metis Ints_of_int antisym_conv2 floor_correct l)
 
 lemma eventually_less_ceiling:
-  fixes f::"'a \<Rightarrow> 'b::{order_topology, floor_ceiling}"
+  fixes f :: "'a \<Rightarrow> 'b::{order_topology,floor_ceiling}"
   assumes f: "(f \<longlongrightarrow> l) F"
-  assumes l: "l \<notin> \<int>"
+    and l: "l \<notin> \<int>"
   shows "\<forall>\<^sub>F x in F. f x < of_int (ceiling l)"
   by (intro order_tendstoD[OF f]) (metis Ints_of_int l le_of_int_ceiling less_le)
 
 lemma eventually_floor_eq:
-  fixes f::"'a \<Rightarrow> 'b::{order_topology, floor_ceiling}"
+  fixes f::"'a \<Rightarrow> 'b::{order_topology,floor_ceiling}"
   assumes f: "(f \<longlongrightarrow> l) F"
-  assumes l: "l \<notin> \<int>"
+    and l: "l \<notin> \<int>"
   shows "\<forall>\<^sub>F x in F. floor (f x) = floor l"
   using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms]
   by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq)
 
 lemma eventually_ceiling_eq:
-  fixes f::"'a \<Rightarrow> 'b::{order_topology, floor_ceiling}"
+  fixes f::"'a \<Rightarrow> 'b::{order_topology,floor_ceiling}"
   assumes f: "(f \<longlongrightarrow> l) F"
-  assumes l: "l \<notin> \<int>"
+    and l: "l \<notin> \<int>"
   shows "\<forall>\<^sub>F x in F. ceiling (f x) = ceiling l"
   using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms]
   by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq)
 
 lemma tendsto_of_int_floor:
-  fixes f::"'a \<Rightarrow> 'b::{order_topology, floor_ceiling}"
+  fixes f::"'a \<Rightarrow> 'b::{order_topology,floor_ceiling}"
   assumes "(f \<longlongrightarrow> l) F"
-  assumes "l \<notin> \<int>"
-  shows "((\<lambda>x. of_int (floor (f x))::'c::{ring_1, topological_space}) \<longlongrightarrow> of_int (floor l)) F"
+    and "l \<notin> \<int>"
+  shows "((\<lambda>x. of_int (floor (f x)) :: 'c::{ring_1,topological_space}) \<longlongrightarrow> of_int (floor l)) F"
   using eventually_floor_eq[OF assms]
   by (simp add: eventually_mono topological_tendstoI)
 
 lemma tendsto_of_int_ceiling:
-  fixes f::"'a \<Rightarrow> 'b::{order_topology, floor_ceiling}"
+  fixes f::"'a \<Rightarrow> 'b::{order_topology,floor_ceiling}"
   assumes "(f \<longlongrightarrow> l) F"
-  assumes "l \<notin> \<int>"
-  shows "((\<lambda>x. of_int (ceiling (f x))::'c::{ring_1, topological_space}) \<longlongrightarrow> of_int (ceiling l)) F"
+    and "l \<notin> \<int>"
+  shows "((\<lambda>x. of_int (ceiling (f x)):: 'c::{ring_1,topological_space}) \<longlongrightarrow> of_int (ceiling l)) F"
   using eventually_ceiling_eq[OF assms]
   by (simp add: eventually_mono topological_tendstoI)
 
@@ -1580,69 +1686,64 @@
   shows "(X \<longlonglongrightarrow> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)"
 unfolding lim_sequentially dist_norm ..
 
-lemma LIMSEQ_I:
-  fixes L :: "'a::real_normed_vector"
-  shows "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X \<longlonglongrightarrow> L"
-by (simp add: LIMSEQ_iff)
-
-lemma LIMSEQ_D:
-  fixes L :: "'a::real_normed_vector"
-  shows "\<lbrakk>X \<longlonglongrightarrow> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
-by (simp add: LIMSEQ_iff)
-
-lemma LIMSEQ_linear: "\<lbrakk> X \<longlonglongrightarrow> x ; l > 0 \<rbrakk> \<Longrightarrow> (\<lambda> n. X (n * l)) \<longlonglongrightarrow> x"
+lemma LIMSEQ_I: "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X \<longlonglongrightarrow> L"
+  for L :: "'a::real_normed_vector"
+  by (simp add: LIMSEQ_iff)
+
+lemma LIMSEQ_D: "X \<longlonglongrightarrow> L \<Longrightarrow> 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
+  for L :: "'a::real_normed_vector"
+  by (simp add: LIMSEQ_iff)
+
+lemma LIMSEQ_linear: "X \<longlonglongrightarrow> x \<Longrightarrow> l > 0 \<Longrightarrow> (\<lambda> n. X (n * l)) \<longlonglongrightarrow> x"
   unfolding tendsto_def eventually_sequentially
   by (metis div_le_dividend div_mult_self1_is_m le_trans mult.commute)
 
-lemma Bseq_inverse_lemma:
-  fixes x :: "'a::real_normed_div_algebra"
-  shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
-apply (subst nonzero_norm_inverse, clarsimp)
-apply (erule (1) le_imp_inverse_le)
-done
-
-lemma Bseq_inverse:
-  fixes a :: "'a::real_normed_div_algebra"
-  shows "\<lbrakk>X \<longlonglongrightarrow> a; a \<noteq> 0\<rbrakk> \<Longrightarrow> Bseq (\<lambda>n. inverse (X n))"
+lemma Bseq_inverse_lemma: "r \<le> norm x \<Longrightarrow> 0 < r \<Longrightarrow> norm (inverse x) \<le> inverse r"
+  for x :: "'a::real_normed_div_algebra"
+  apply (subst nonzero_norm_inverse, clarsimp)
+  apply (erule (1) le_imp_inverse_le)
+  done
+
+lemma Bseq_inverse: "X \<longlonglongrightarrow> a \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> Bseq (\<lambda>n. inverse (X n))"
+  for a :: "'a::real_normed_div_algebra"
   by (rule Bfun_inverse)
 
-text\<open>Transformation of limit.\<close>
-
-lemma Lim_transform:
-  fixes a b :: "'a::real_normed_vector"
-  shows "\<lbrakk>(g \<longlongrightarrow> a) F; ((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F\<rbrakk> \<Longrightarrow> (f \<longlongrightarrow> a) F"
+
+text \<open>Transformation of limit.\<close>
+
+lemma Lim_transform: "(g \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> a) F"
+  for a b :: "'a::real_normed_vector"
   using tendsto_add [of g a F "\<lambda>x. f x - g x" 0] by simp
 
-lemma Lim_transform2:
-  fixes a b :: "'a::real_normed_vector"
-  shows "\<lbrakk>(f \<longlongrightarrow> a) F; ((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F\<rbrakk> \<Longrightarrow> (g \<longlongrightarrow> a) F"
+lemma Lim_transform2: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F \<Longrightarrow> (g \<longlongrightarrow> a) F"
+  for a b :: "'a::real_normed_vector"
   by (erule Lim_transform) (simp add: tendsto_minus_cancel)
 
-proposition Lim_transform_eq:
-  fixes a :: "'a::real_normed_vector"
-  shows "((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> a) F \<longleftrightarrow> (g \<longlongrightarrow> a) F"
-using Lim_transform Lim_transform2 by blast
+proposition Lim_transform_eq: "((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> a) F \<longleftrightarrow> (g \<longlongrightarrow> a) F"
+  for a :: "'a::real_normed_vector"
+  using Lim_transform Lim_transform2 by blast
 
 lemma Lim_transform_eventually:
   "eventually (\<lambda>x. f x = g x) net \<Longrightarrow> (f \<longlongrightarrow> l) net \<Longrightarrow> (g \<longlongrightarrow> l) net"
   apply (rule topological_tendstoI)
   apply (drule (2) topological_tendstoD)
-  apply (erule (1) eventually_elim2, simp)
+  apply (erule (1) eventually_elim2)
+  apply simp
   done
 
 lemma Lim_transform_within:
   assumes "(f \<longlongrightarrow> l) (at x within S)"
     and "0 < d"
-    and "\<And>x'. \<lbrakk>x'\<in>S; 0 < dist x' x; dist x' x < d\<rbrakk> \<Longrightarrow> f x' = g x'"
+    and "\<And>x'. x'\<in>S \<Longrightarrow> 0 < dist x' x \<Longrightarrow> dist x' x < d \<Longrightarrow> f x' = g x'"
   shows "(g \<longlongrightarrow> l) (at x within S)"
 proof (rule Lim_transform_eventually)
   show "eventually (\<lambda>x. f x = g x) (at x within S)"
     using assms by (auto simp: eventually_at)
-  show "(f \<longlongrightarrow> l) (at x within S)" by fact
+  show "(f \<longlongrightarrow> l) (at x within S)"
+    by fact
 qed
 
-text\<open>Common case assuming being away from some crucial point like 0.\<close>
-
+text \<open>Common case assuming being away from some crucial point like 0.\<close>
 lemma Lim_transform_away_within:
   fixes a b :: "'a::t1_space"
   assumes "a \<noteq> b"
@@ -1650,26 +1751,26 @@
     and "(f \<longlongrightarrow> l) (at a within S)"
   shows "(g \<longlongrightarrow> l) (at a within S)"
 proof (rule Lim_transform_eventually)
-  show "(f \<longlongrightarrow> l) (at a within S)" by fact
+  show "(f \<longlongrightarrow> l) (at a within S)"
+    by fact
   show "eventually (\<lambda>x. f x = g x) (at a within S)"
     unfolding eventually_at_topological
-    by (rule exI [where x="- {b}"], simp add: open_Compl assms)
+    by (rule exI [where x="- {b}"]) (simp add: open_Compl assms)
 qed
 
 lemma Lim_transform_away_at:
   fixes a b :: "'a::t1_space"
-  assumes ab: "a\<noteq>b"
+  assumes ab: "a \<noteq> b"
     and fg: "\<forall>x. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
     and fl: "(f \<longlongrightarrow> l) (at a)"
   shows "(g \<longlongrightarrow> l) (at a)"
   using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl by simp
 
-text\<open>Alternatively, within an open set.\<close>
-
+text \<open>Alternatively, within an open set.\<close>
 lemma Lim_transform_within_open:
   assumes "(f \<longlongrightarrow> l) (at a within T)"
     and "open s" and "a \<in> s"
-    and "\<And>x. \<lbrakk>x\<in>s; x \<noteq> a\<rbrakk> \<Longrightarrow> f x = g x"
+    and "\<And>x. x\<in>s \<Longrightarrow> x \<noteq> a \<Longrightarrow> f x = g x"
   shows "(g \<longlongrightarrow> l) (at a within T)"
 proof (rule Lim_transform_eventually)
   show "eventually (\<lambda>x. f x = g x) (at a within T)"
@@ -1678,7 +1779,8 @@
   show "(f \<longlongrightarrow> l) (at a within T)" by fact
 qed
 
-text\<open>A congruence rule allowing us to transform limits assuming not at point.\<close>
+
+text \<open>A congruence rule allowing us to transform limits assuming not at point.\<close>
 
 (* FIXME: Only one congruence rule for tendsto can be used at a time! *)
 
@@ -1697,35 +1799,32 @@
   shows "((\<lambda>x. f x) \<longlongrightarrow> x) (at a) \<longleftrightarrow> ((g \<longlongrightarrow> y) (at a))"
   unfolding tendsto_def eventually_at_topological
   using assms by simp
-text\<open>An unbounded sequence's inverse tends to 0\<close>
-
-lemma LIMSEQ_inverse_zero:
-  "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) \<longlonglongrightarrow> 0"
+
+text \<open>An unbounded sequence's inverse tends to 0.\<close>
+lemma LIMSEQ_inverse_zero: "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) \<longlonglongrightarrow> 0"
   apply (rule filterlim_compose[OF tendsto_inverse_0])
   apply (simp add: filterlim_at_infinity[OF order_refl] eventually_sequentially)
   apply (metis abs_le_D1 linorder_le_cases linorder_not_le)
   done
 
-text\<open>The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity\<close>
-
-lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) \<longlonglongrightarrow> 0"
+text \<open>The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity.\<close>
+lemma LIMSEQ_inverse_real_of_nat: "(\<lambda>n. inverse (real (Suc n))) \<longlonglongrightarrow> 0"
   by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc
-            filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity)
-
-text\<open>The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
-infinity is now easily proved\<close>
-
-lemma LIMSEQ_inverse_real_of_nat_add:
-     "(%n. r + inverse(real(Suc n))) \<longlonglongrightarrow> r"
+      filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity)
+
+text \<open>
+  The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
+  infinity is now easily proved.
+\<close>
+
+lemma LIMSEQ_inverse_real_of_nat_add: "(\<lambda>n. r + inverse (real (Suc n))) \<longlonglongrightarrow> r"
   using tendsto_add [OF tendsto_const LIMSEQ_inverse_real_of_nat] by auto
 
-lemma LIMSEQ_inverse_real_of_nat_add_minus:
-     "(%n. r + -inverse(real(Suc n))) \<longlonglongrightarrow> r"
+lemma LIMSEQ_inverse_real_of_nat_add_minus: "(\<lambda>n. r + -inverse (real (Suc n))) \<longlonglongrightarrow> r"
   using tendsto_add [OF tendsto_const tendsto_minus [OF LIMSEQ_inverse_real_of_nat]]
   by auto
 
-lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
-     "(%n. r*( 1 + -inverse(real(Suc n)))) \<longlonglongrightarrow> r"
+lemma LIMSEQ_inverse_real_of_nat_add_minus_mult: "(\<lambda>n. r * (1 + - inverse (real (Suc n)))) \<longlonglongrightarrow> r"
   using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]]
   by auto
 
@@ -1735,46 +1834,57 @@
 lemma LIMSEQ_Suc_n_over_n: "(\<lambda>n. of_nat (Suc n) / of_nat n :: 'a :: real_normed_field) \<longlonglongrightarrow> 1"
 proof (rule Lim_transform_eventually)
   show "eventually (\<lambda>n. 1 + inverse (of_nat n :: 'a) = of_nat (Suc n) / of_nat n) sequentially"
-    using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: field_simps)
+    using eventually_gt_at_top[of "0::nat"]
+    by eventually_elim (simp add: field_simps)
   have "(\<lambda>n. 1 + inverse (of_nat n) :: 'a) \<longlonglongrightarrow> 1 + 0"
     by (intro tendsto_add tendsto_const lim_inverse_n)
-  thus "(\<lambda>n. 1 + inverse (of_nat n) :: 'a) \<longlonglongrightarrow> 1" by simp
+  then show "(\<lambda>n. 1 + inverse (of_nat n) :: 'a) \<longlonglongrightarrow> 1"
+    by simp
 qed
 
 lemma LIMSEQ_n_over_Suc_n: "(\<lambda>n. of_nat n / of_nat (Suc n) :: 'a :: real_normed_field) \<longlonglongrightarrow> 1"
 proof (rule Lim_transform_eventually)
   show "eventually (\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a) =
-                        of_nat n / of_nat (Suc n)) sequentially"
+      of_nat n / of_nat (Suc n)) sequentially"
     using eventually_gt_at_top[of "0::nat"]
     by eventually_elim (simp add: field_simps del: of_nat_Suc)
   have "(\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \<longlonglongrightarrow> inverse 1"
     by (intro tendsto_inverse LIMSEQ_Suc_n_over_n) simp_all
-  thus "(\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \<longlonglongrightarrow> 1" by simp
+  then show "(\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \<longlonglongrightarrow> 1"
+    by simp
 qed
 
+
 subsection \<open>Convergence on sequences\<close>
 
 lemma convergent_cong:
   assumes "eventually (\<lambda>x. f x = g x) sequentially"
-  shows   "convergent f \<longleftrightarrow> convergent g"
-  unfolding convergent_def by (subst filterlim_cong[OF refl refl assms]) (rule refl)
+  shows "convergent f \<longleftrightarrow> convergent g"
+  unfolding convergent_def
+  by (subst filterlim_cong[OF refl refl assms]) (rule refl)
 
 lemma convergent_Suc_iff: "convergent (\<lambda>n. f (Suc n)) \<longleftrightarrow> convergent f"
   by (auto simp: convergent_def LIMSEQ_Suc_iff)
 
 lemma convergent_ignore_initial_segment: "convergent (\<lambda>n. f (n + m)) = convergent f"
-proof (induction m arbitrary: f)
+proof (induct m arbitrary: f)
+  case 0
+  then show ?case by simp
+next
   case (Suc m)
-  have "convergent (\<lambda>n. f (n + Suc m)) \<longleftrightarrow> convergent (\<lambda>n. f (Suc n + m))" by simp
-  also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. f (n + m))" by (rule convergent_Suc_iff)
-  also have "\<dots> \<longleftrightarrow> convergent f" by (rule Suc)
+  have "convergent (\<lambda>n. f (n + Suc m)) \<longleftrightarrow> convergent (\<lambda>n. f (Suc n + m))"
+    by simp
+  also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. f (n + m))"
+    by (rule convergent_Suc_iff)
+  also have "\<dots> \<longleftrightarrow> convergent f"
+    by (rule Suc)
   finally show ?case .
-qed simp_all
+qed
 
 lemma convergent_add:
   fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
   assumes "convergent (\<lambda>n. X n)"
-  assumes "convergent (\<lambda>n. Y n)"
+    and "convergent (\<lambda>n. Y n)"
   shows "convergent (\<lambda>n. X n + Y n)"
   using assms unfolding convergent_def by (blast intro: tendsto_add)
 
@@ -1783,9 +1893,14 @@
   assumes "\<And>i. i \<in> A \<Longrightarrow> convergent (\<lambda>n. X i n)"
   shows "convergent (\<lambda>n. \<Sum>i\<in>A. X i n)"
 proof (cases "finite A")
-  case True from this and assms show ?thesis
-    by (induct A set: finite) (simp_all add: convergent_const convergent_add)
-qed (simp add: convergent_const)
+  case True
+  then show ?thesis
+    using assms by (induct A set: finite) (simp_all add: convergent_const convergent_add)
+next
+  case False
+  then show ?thesis
+    by (simp add: convergent_const)
+qed
 
 lemma (in bounded_linear) convergent:
   assumes "convergent (\<lambda>n. X n)"
@@ -1793,17 +1908,18 @@
   using assms unfolding convergent_def by (blast intro: tendsto)
 
 lemma (in bounded_bilinear) convergent:
-  assumes "convergent (\<lambda>n. X n)" and "convergent (\<lambda>n. Y n)"
+  assumes "convergent (\<lambda>n. X n)"
+    and "convergent (\<lambda>n. Y n)"
   shows "convergent (\<lambda>n. X n ** Y n)"
   using assms unfolding convergent_def by (blast intro: tendsto)
 
-lemma convergent_minus_iff:
-  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
-  shows "convergent X \<longleftrightarrow> convergent (\<lambda>n. - X n)"
-apply (simp add: convergent_def)
-apply (auto dest: tendsto_minus)
-apply (drule tendsto_minus, auto)
-done
+lemma convergent_minus_iff: "convergent X \<longleftrightarrow> convergent (\<lambda>n. - X n)"
+  for X :: "nat \<Rightarrow> 'a::real_normed_vector"
+  apply (simp add: convergent_def)
+  apply (auto dest: tendsto_minus)
+  apply (drule tendsto_minus)
+  apply auto
+  done
 
 lemma convergent_diff:
   fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
@@ -1814,57 +1930,64 @@
 
 lemma convergent_norm:
   assumes "convergent f"
-  shows   "convergent (\<lambda>n. norm (f n))"
+  shows "convergent (\<lambda>n. norm (f n))"
 proof -
-  from assms have "f \<longlonglongrightarrow> lim f" by (simp add: convergent_LIMSEQ_iff)
-  hence "(\<lambda>n. norm (f n)) \<longlonglongrightarrow> norm (lim f)" by (rule tendsto_norm)
-  thus ?thesis by (auto simp: convergent_def)
+  from assms have "f \<longlonglongrightarrow> lim f"
+    by (simp add: convergent_LIMSEQ_iff)
+  then have "(\<lambda>n. norm (f n)) \<longlonglongrightarrow> norm (lim f)"
+    by (rule tendsto_norm)
+  then show ?thesis
+    by (auto simp: convergent_def)
 qed
 
 lemma convergent_of_real:
-  "convergent f \<Longrightarrow> convergent (\<lambda>n. of_real (f n) :: 'a :: real_normed_algebra_1)"
+  "convergent f \<Longrightarrow> convergent (\<lambda>n. of_real (f n) :: 'a::real_normed_algebra_1)"
   unfolding convergent_def by (blast intro!: tendsto_of_real)
 
 lemma convergent_add_const_iff:
-  "convergent (\<lambda>n. c + f n :: 'a :: real_normed_vector) \<longleftrightarrow> convergent f"
+  "convergent (\<lambda>n. c + f n :: 'a::real_normed_vector) \<longleftrightarrow> convergent f"
 proof
   assume "convergent (\<lambda>n. c + f n)"
-  from convergent_diff[OF this convergent_const[of c]] show "convergent f" by simp
+  from convergent_diff[OF this convergent_const[of c]] show "convergent f"
+    by simp
 next
   assume "convergent f"
-  from convergent_add[OF convergent_const[of c] this] show "convergent (\<lambda>n. c + f n)" by simp
+  from convergent_add[OF convergent_const[of c] this] show "convergent (\<lambda>n. c + f n)"
+    by simp
 qed
 
 lemma convergent_add_const_right_iff:
-  "convergent (\<lambda>n. f n + c :: 'a :: real_normed_vector) \<longleftrightarrow> convergent f"
+  "convergent (\<lambda>n. f n + c :: 'a::real_normed_vector) \<longleftrightarrow> convergent f"
   using convergent_add_const_iff[of c f] by (simp add: add_ac)
 
 lemma convergent_diff_const_right_iff:
-  "convergent (\<lambda>n. f n - c :: 'a :: real_normed_vector) \<longleftrightarrow> convergent f"
+  "convergent (\<lambda>n. f n - c :: 'a::real_normed_vector) \<longleftrightarrow> convergent f"
   using convergent_add_const_right_iff[of f "-c"] by (simp add: add_ac)
 
 lemma convergent_mult:
   fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
   assumes "convergent (\<lambda>n. X n)"
-  assumes "convergent (\<lambda>n. Y n)"
+    and "convergent (\<lambda>n. Y n)"
   shows "convergent (\<lambda>n. X n * Y n)"
   using assms unfolding convergent_def by (blast intro: tendsto_mult)
 
 lemma convergent_mult_const_iff:
   assumes "c \<noteq> 0"
-  shows   "convergent (\<lambda>n. c * f n :: 'a :: real_normed_field) \<longleftrightarrow> convergent f"
+  shows "convergent (\<lambda>n. c * f n :: 'a::real_normed_field) \<longleftrightarrow> convergent f"
 proof
   assume "convergent (\<lambda>n. c * f n)"
   from assms convergent_mult[OF this convergent_const[of "inverse c"]]
     show "convergent f" by (simp add: field_simps)
 next
   assume "convergent f"
-  from convergent_mult[OF convergent_const[of c] this] show "convergent (\<lambda>n. c * f n)" by simp
+  from convergent_mult[OF convergent_const[of c] this] show "convergent (\<lambda>n. c * f n)"
+    by simp
 qed
 
 lemma convergent_mult_const_right_iff:
+  fixes c :: "'a::real_normed_field"
   assumes "c \<noteq> 0"
-  shows   "convergent (\<lambda>n. (f n :: 'a :: real_normed_field) * c) \<longleftrightarrow> convergent f"
+  shows "convergent (\<lambda>n. f n * c) \<longleftrightarrow> convergent f"
   using convergent_mult_const_iff[OF assms, of f] by (simp add: mult_ac)
 
 lemma convergent_imp_Bseq: "convergent f \<Longrightarrow> Bseq f"
@@ -1874,60 +1997,66 @@
 text \<open>A monotone sequence converges to its least upper bound.\<close>
 
 lemma LIMSEQ_incseq_SUP:
-  fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology}"
+  fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder,linorder_topology}"
   assumes u: "bdd_above (range X)"
-  assumes X: "incseq X"
+    and X: "incseq X"
   shows "X \<longlonglongrightarrow> (SUP i. X i)"
   by (rule order_tendstoI)
-     (auto simp: eventually_sequentially u less_cSUP_iff intro: X[THEN incseqD] less_le_trans cSUP_lessD[OF u])
+    (auto simp: eventually_sequentially u less_cSUP_iff
+      intro: X[THEN incseqD] less_le_trans cSUP_lessD[OF u])
 
 lemma LIMSEQ_decseq_INF:
   fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology}"
   assumes u: "bdd_below (range X)"
-  assumes X: "decseq X"
+    and X: "decseq X"
   shows "X \<longlonglongrightarrow> (INF i. X i)"
   by (rule order_tendstoI)
-     (auto simp: eventually_sequentially u cINF_less_iff intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u])
-
-text\<open>Main monotonicity theorem\<close>
-
-lemma Bseq_monoseq_convergent: "Bseq X \<Longrightarrow> monoseq X \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
-  by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP dest: Bseq_bdd_above Bseq_bdd_below)
-
-lemma Bseq_mono_convergent: "Bseq X \<Longrightarrow> (\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n) \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
+     (auto simp: eventually_sequentially u cINF_less_iff
+       intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u])
+
+text \<open>Main monotonicity theorem.\<close>
+
+lemma Bseq_monoseq_convergent: "Bseq X \<Longrightarrow> monoseq X \<Longrightarrow> convergent X"
+  for X :: "nat \<Rightarrow> real"
+  by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP
+      dest: Bseq_bdd_above Bseq_bdd_below)
+
+lemma Bseq_mono_convergent: "Bseq X \<Longrightarrow> (\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n) \<Longrightarrow> convergent X"
+  for X :: "nat \<Rightarrow> real"
   by (auto intro!: Bseq_monoseq_convergent incseq_imp_monoseq simp: incseq_def)
 
-lemma monoseq_imp_convergent_iff_Bseq: "monoseq (f :: nat \<Rightarrow> real) \<Longrightarrow> convergent f \<longleftrightarrow> Bseq f"
+lemma monoseq_imp_convergent_iff_Bseq: "monoseq f \<Longrightarrow> convergent f \<longleftrightarrow> Bseq f"
+  for f :: "nat \<Rightarrow> real"
   using Bseq_monoseq_convergent[of f] convergent_imp_Bseq[of f] by blast
 
 lemma Bseq_monoseq_convergent'_inc:
-  "Bseq (\<lambda>n. f (n + M) :: real) \<Longrightarrow> (\<And>m n. M \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m \<le> f n) \<Longrightarrow> convergent f"
+  fixes f :: "nat \<Rightarrow> real"
+  shows "Bseq (\<lambda>n. f (n + M)) \<Longrightarrow> (\<And>m n. M \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m \<le> f n) \<Longrightarrow> convergent f"
   by (subst convergent_ignore_initial_segment [symmetric, of _ M])
      (auto intro!: Bseq_monoseq_convergent simp: monoseq_def)
 
 lemma Bseq_monoseq_convergent'_dec:
-  "Bseq (\<lambda>n. f (n + M) :: real) \<Longrightarrow> (\<And>m n. M \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m \<ge> f n) \<Longrightarrow> convergent f"
+  fixes f :: "nat \<Rightarrow> real"
+  shows "Bseq (\<lambda>n. f (n + M)) \<Longrightarrow> (\<And>m n. M \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m \<ge> f n) \<Longrightarrow> convergent f"
   by (subst convergent_ignore_initial_segment [symmetric, of _ M])
-     (auto intro!: Bseq_monoseq_convergent simp: monoseq_def)
-
-lemma Cauchy_iff:
-  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
-  shows "Cauchy X \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e)"
+    (auto intro!: Bseq_monoseq_convergent simp: monoseq_def)
+
+lemma Cauchy_iff: "Cauchy X \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e)"
+  for X :: "nat \<Rightarrow> 'a::real_normed_vector"
   unfolding Cauchy_def dist_norm ..
 
-lemma CauchyI:
-  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
-  shows "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
-by (simp add: Cauchy_iff)
-
-lemma CauchyD:
-  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
-  shows "\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
-by (simp add: Cauchy_iff)
+lemma CauchyI: "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
+  for X :: "nat \<Rightarrow> 'a::real_normed_vector"
+  by (simp add: Cauchy_iff)
+
+lemma CauchyD: "Cauchy X \<Longrightarrow> 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
+  for X :: "nat \<Rightarrow> 'a::real_normed_vector"
+  by (simp add: Cauchy_iff)
 
 lemma incseq_convergent:
   fixes X :: "nat \<Rightarrow> real"
-  assumes "incseq X" and "\<forall>i. X i \<le> B"
+  assumes "incseq X"
+    and "\<forall>i. X i \<le> B"
   obtains L where "X \<longlonglongrightarrow> L" "\<forall>i. X i \<le> L"
 proof atomize_elim
   from incseq_bounded[OF assms] \<open>incseq X\<close> Bseq_monoseq_convergent[of X]
@@ -1939,7 +2068,8 @@
 
 lemma decseq_convergent:
   fixes X :: "nat \<Rightarrow> real"
-  assumes "decseq X" and "\<forall>i. B \<le> X i"
+  assumes "decseq X"
+    and "\<forall>i. B \<le> X i"
   obtains L where "X \<longlonglongrightarrow> L" "\<forall>i. L \<le> X i"
 proof atomize_elim
   from decseq_bounded[OF assms] \<open>decseq X\<close> Bseq_monoseq_convergent[of X]
@@ -1949,69 +2079,85 @@
     by (auto intro!: exI[of _ L] decseq_le)
 qed
 
+
 subsubsection \<open>Cauchy Sequences are Bounded\<close>
 
-text\<open>A Cauchy sequence is bounded -- this is the standard
-  proof mechanization rather than the nonstandard proof\<close>
-
-lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
-          ==>  \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
-apply (clarify, drule spec, drule (1) mp)
-apply (simp only: norm_minus_commute)
-apply (drule order_le_less_trans [OF norm_triangle_ineq2])
-apply simp
-done
+text \<open>
+  A Cauchy sequence is bounded -- this is the standard
+  proof mechanization rather than the nonstandard proof.
+\<close>
+
+lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real) \<Longrightarrow>
+  \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
+  apply clarify
+  apply (drule spec)
+  apply (drule (1) mp)
+  apply (simp only: norm_minus_commute)
+  apply (drule order_le_less_trans [OF norm_triangle_ineq2])
+  apply simp
+  done
+
 
 subsection \<open>Power Sequences\<close>
 
-text\<open>The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
-"x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
-  also fact that bounded and monotonic sequence converges.\<close>
-
-lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
-apply (simp add: Bseq_def)
-apply (rule_tac x = 1 in exI)
-apply (simp add: power_abs)
-apply (auto dest: power_mono)
-done
-
-lemma monoseq_realpow: fixes x :: real shows "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
-apply (clarify intro!: mono_SucI2)
-apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
-done
-
-lemma convergent_realpow:
-  "[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)"
-by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
-
-lemma LIMSEQ_inverse_realpow_zero: "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) \<longlonglongrightarrow> 0"
+text \<open>
+  The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
+  "x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
+  also fact that bounded and monotonic sequence converges.
+\<close>
+
+lemma Bseq_realpow: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> Bseq (\<lambda>n. x ^ n)"
+  for x :: real
+  apply (simp add: Bseq_def)
+  apply (rule_tac x = 1 in exI)
+  apply (simp add: power_abs)
+  apply (auto dest: power_mono)
+  done
+
+lemma monoseq_realpow: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> monoseq (\<lambda>n. x ^ n)"
+  for x :: real
+  apply (clarify intro!: mono_SucI2)
+  apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing)
+     apply auto
+  done
+
+lemma convergent_realpow: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> convergent (\<lambda>n. x ^ n)"
+  for x :: real
+  by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
+
+lemma LIMSEQ_inverse_realpow_zero: "1 < x \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) \<longlonglongrightarrow> 0"
+  for x :: real
   by (rule filterlim_compose[OF tendsto_inverse_0 filterlim_realpow_sequentially_gt1]) simp
 
 lemma LIMSEQ_realpow_zero:
-  "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) \<longlonglongrightarrow> 0"
-proof cases
-  assume "0 \<le> x" and "x \<noteq> 0"
-  hence x0: "0 < x" by simp
-  assume x1: "x < 1"
-  from x0 x1 have "1 < inverse x"
-    by (rule one_less_inverse)
-  hence "(\<lambda>n. inverse (inverse x ^ n)) \<longlonglongrightarrow> 0"
+  fixes x :: real
+  assumes "0 \<le> x" "x < 1"
+  shows "(\<lambda>n. x ^ n) \<longlonglongrightarrow> 0"
+proof (cases "x = 0")
+  case False
+  with \<open>0 \<le> x\<close> have x0: "0 < x" by simp
+  then have "1 < inverse x"
+    using \<open>x < 1\<close> by (rule one_less_inverse)
+  then have "(\<lambda>n. inverse (inverse x ^ n)) \<longlonglongrightarrow> 0"
     by (rule LIMSEQ_inverse_realpow_zero)
-  thus ?thesis by (simp add: power_inverse)
-qed (rule LIMSEQ_imp_Suc, simp)
-
-lemma LIMSEQ_power_zero:
-  fixes x :: "'a::{real_normed_algebra_1}"
-  shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) \<longlonglongrightarrow> 0"
-apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
-apply (simp only: tendsto_Zfun_iff, erule Zfun_le)
-apply (simp add: power_abs norm_power_ineq)
-done
+  then show ?thesis by (simp add: power_inverse)
+next
+  case True
+  show ?thesis
+    by (rule LIMSEQ_imp_Suc) (simp add: True)
+qed
+
+lemma LIMSEQ_power_zero: "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) \<longlonglongrightarrow> 0"
+  for x :: "'a::real_normed_algebra_1"
+  apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
+  apply (simp only: tendsto_Zfun_iff, erule Zfun_le)
+  apply (simp add: power_abs norm_power_ineq)
+  done
 
 lemma LIMSEQ_divide_realpow_zero: "1 < x \<Longrightarrow> (\<lambda>n. a / (x ^ n) :: real) \<longlonglongrightarrow> 0"
   by (rule tendsto_divide_0 [OF tendsto_const filterlim_realpow_sequentially_gt1]) simp
 
-text\<open>Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}\<close>
+text \<open>Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}.\<close>
 
 lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. \<bar>c\<bar> ^ n :: real) \<longlonglongrightarrow> 0"
   by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
@@ -2022,92 +2168,81 @@
 
 subsection \<open>Limits of Functions\<close>
 
-lemma LIM_eq:
-  fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
-  shows "f \<midarrow>a\<rightarrow> L =
-     (\<forall>r>0.\<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)"
-by (simp add: LIM_def dist_norm)
+lemma LIM_eq: "f \<midarrow>a\<rightarrow> L = (\<forall>r>0. \<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - L) < r)"
+  for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
+  by (simp add: LIM_def dist_norm)
 
 lemma LIM_I:
-  fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
-  shows "(!!r. 0<r ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)
-      ==> f \<midarrow>a\<rightarrow> L"
-by (simp add: LIM_eq)
-
-lemma LIM_D:
-  fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
-  shows "[| f \<midarrow>a\<rightarrow> L; 0<r |]
-      ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r"
-by (simp add: LIM_eq)
-
-lemma LIM_offset:
-  fixes a :: "'a::real_normed_vector"
-  shows "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. f (x + k)) \<midarrow>(a - k)\<rightarrow> L"
-  unfolding filtermap_at_shift[symmetric, of a k] filterlim_def filtermap_filtermap by simp
-
-lemma LIM_offset_zero:
-  fixes a :: "'a::real_normed_vector"
-  shows "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L"
-by (drule_tac k="a" in LIM_offset, simp add: add.commute)
-
-lemma LIM_offset_zero_cancel:
-  fixes a :: "'a::real_normed_vector"
-  shows "(\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L \<Longrightarrow> f \<midarrow>a\<rightarrow> L"
-by (drule_tac k="- a" in LIM_offset, simp)
-
-lemma LIM_offset_zero_iff:
-  fixes f :: "'a :: real_normed_vector \<Rightarrow> _"
-  shows  "f \<midarrow>a\<rightarrow> L \<longleftrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L"
+  "(\<And>r. 0 < r \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - L) < r) \<Longrightarrow> f \<midarrow>a\<rightarrow> L"
+  for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
+  by (simp add: LIM_eq)
+
+lemma LIM_D: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> 0 < r \<Longrightarrow> \<exists>s>0.\<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - L) < r"
+  for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
+  by (simp add: LIM_eq)
+
+lemma LIM_offset: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. f (x + k)) \<midarrow>(a - k)\<rightarrow> L"
+  for a :: "'a::real_normed_vector"
+  by (simp add: filtermap_at_shift[symmetric, of a k] filterlim_def filtermap_filtermap)
+
+lemma LIM_offset_zero: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L"
+  for a :: "'a::real_normed_vector"
+  by (drule LIM_offset [where k = a]) (simp add: add.commute)
+
+lemma LIM_offset_zero_cancel: "(\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L \<Longrightarrow> f \<midarrow>a\<rightarrow> L"
+  for a :: "'a::real_normed_vector"
+  by (drule LIM_offset [where k = "- a"]) simp
+
+lemma LIM_offset_zero_iff: "f \<midarrow>a\<rightarrow> L \<longleftrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L"
+  for f :: "'a :: real_normed_vector \<Rightarrow> _"
   using LIM_offset_zero_cancel[of f a L] LIM_offset_zero[of f L a] by auto
 
-lemma LIM_zero:
-  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
-  shows "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. f x - l) \<longlongrightarrow> 0) F"
-unfolding tendsto_iff dist_norm by simp
+lemma LIM_zero: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. f x - l) \<longlongrightarrow> 0) F"
+  for f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
+  unfolding tendsto_iff dist_norm by simp
 
 lemma LIM_zero_cancel:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
   shows "((\<lambda>x. f x - l) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> l) F"
 unfolding tendsto_iff dist_norm by simp
 
-lemma LIM_zero_iff:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
-  shows "((\<lambda>x. f x - l) \<longlongrightarrow> 0) F = (f \<longlongrightarrow> l) F"
-unfolding tendsto_iff dist_norm by simp
+lemma LIM_zero_iff: "((\<lambda>x. f x - l) \<longlongrightarrow> 0) F = (f \<longlongrightarrow> l) F"
+  for f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  unfolding tendsto_iff dist_norm by simp
 
 lemma LIM_imp_LIM:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
   fixes g :: "'a::topological_space \<Rightarrow> 'c::real_normed_vector"
   assumes f: "f \<midarrow>a\<rightarrow> l"
-  assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> norm (g x - m) \<le> norm (f x - l)"
+    and le: "\<And>x. x \<noteq> a \<Longrightarrow> norm (g x - m) \<le> norm (f x - l)"
   shows "g \<midarrow>a\<rightarrow> m"
-  by (rule metric_LIM_imp_LIM [OF f],
-    simp add: dist_norm le)
+  by (rule metric_LIM_imp_LIM [OF f]) (simp add: dist_norm le)
 
 lemma LIM_equal2:
   fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
-  assumes 1: "0 < R"
-  assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < R\<rbrakk> \<Longrightarrow> f x = g x"
+  assumes "0 < R"
+    and "\<And>x. x \<noteq> a \<Longrightarrow> norm (x - a) < R \<Longrightarrow> f x = g x"
   shows "g \<midarrow>a\<rightarrow> l \<Longrightarrow> f \<midarrow>a\<rightarrow> l"
-by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm)
+  by (rule metric_LIM_equal2 [OF assms]) (simp_all add: dist_norm)
 
 lemma LIM_compose2:
   fixes a :: "'a::real_normed_vector"
   assumes f: "f \<midarrow>a\<rightarrow> b"
-  assumes g: "g \<midarrow>b\<rightarrow> c"
-  assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b"
+    and g: "g \<midarrow>b\<rightarrow> c"
+    and inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b"
   shows "(\<lambda>x. g (f x)) \<midarrow>a\<rightarrow> c"
-by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]])
+  by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]])
 
 lemma real_LIM_sandwich_zero:
   fixes f g :: "'a::topological_space \<Rightarrow> real"
   assumes f: "f \<midarrow>a\<rightarrow> 0"
-  assumes 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x"
-  assumes 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x"
+    and 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x"
+    and 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x"
   shows "g \<midarrow>a\<rightarrow> 0"
 proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *)
-  fix x assume x: "x \<noteq> a"
-  have "norm (g x - 0) = g x" by (simp add: 1 x)
+  fix x
+  assume x: "x \<noteq> a"
+  with 1 have "norm (g x - 0) = g x" by simp
   also have "g x \<le> f x" by (rule 2 [OF x])
   also have "f x \<le> \<bar>f x\<bar>" by (rule abs_ge_self)
   also have "\<bar>f x\<bar> = norm (f x - 0)" by simp
@@ -2117,61 +2252,50 @@
 
 subsection \<open>Continuity\<close>
 
-lemma LIM_isCont_iff:
-  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
-  shows "(f \<midarrow>a\<rightarrow> f a) = ((\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> f a)"
-by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel])
-
-lemma isCont_iff:
-  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
-  shows "isCont f x = (\<lambda>h. f (x + h)) \<midarrow>0\<rightarrow> f x"
-by (simp add: isCont_def LIM_isCont_iff)
+lemma LIM_isCont_iff: "(f \<midarrow>a\<rightarrow> f a) = ((\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> f a)"
+  for f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
+  by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel])
+
+lemma isCont_iff: "isCont f x = (\<lambda>h. f (x + h)) \<midarrow>0\<rightarrow> f x"
+  for f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
+  by (simp add: isCont_def LIM_isCont_iff)
 
 lemma isCont_LIM_compose2:
   fixes a :: "'a::real_normed_vector"
   assumes f [unfolded isCont_def]: "isCont f a"
-  assumes g: "g \<midarrow>f a\<rightarrow> l"
-  assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a"
+    and g: "g \<midarrow>f a\<rightarrow> l"
+    and inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a"
   shows "(\<lambda>x. g (f x)) \<midarrow>a\<rightarrow> l"
-by (rule LIM_compose2 [OF f g inj])
-
-
-lemma isCont_norm [simp]:
-  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
-  shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a"
+  by (rule LIM_compose2 [OF f g inj])
+
+lemma isCont_norm [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a"
+  for f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
   by (fact continuous_norm)
 
-lemma isCont_rabs [simp]:
-  fixes f :: "'a::t2_space \<Rightarrow> real"
-  shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a"
+lemma isCont_rabs [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a"
+  for f :: "'a::t2_space \<Rightarrow> real"
   by (fact continuous_rabs)
 
-lemma isCont_add [simp]:
-  fixes f :: "'a::t2_space \<Rightarrow> 'b::topological_monoid_add"
-  shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x + g x) a"
+lemma isCont_add [simp]: "isCont f a \<Longrightarrow> isCont g a \<Longrightarrow> isCont (\<lambda>x. f x + g x) a"
+  for f :: "'a::t2_space \<Rightarrow> 'b::topological_monoid_add"
   by (fact continuous_add)
 
-lemma isCont_minus [simp]:
-  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
-  shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a"
+lemma isCont_minus [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a"
+  for f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
   by (fact continuous_minus)
 
-lemma isCont_diff [simp]:
-  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
-  shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x - g x) a"
+lemma isCont_diff [simp]: "isCont f a \<Longrightarrow> isCont g a \<Longrightarrow> isCont (\<lambda>x. f x - g x) a"
+  for f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
   by (fact continuous_diff)
 
-lemma isCont_mult [simp]:
-  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra"
-  shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x * g x) a"
+lemma isCont_mult [simp]: "isCont f a \<Longrightarrow> isCont g a \<Longrightarrow> isCont (\<lambda>x. f x * g x) a"
+  for f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra"
   by (fact continuous_mult)
 
-lemma (in bounded_linear) isCont:
-  "isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a"
+lemma (in bounded_linear) isCont: "isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a"
   by (fact continuous)
 
-lemma (in bounded_bilinear) isCont:
-  "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a"
+lemma (in bounded_bilinear) isCont: "isCont f a \<Longrightarrow> isCont g a \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a"
   by (fact continuous)
 
 lemmas isCont_scaleR [simp] =
@@ -2180,16 +2304,15 @@
 lemmas isCont_of_real [simp] =
   bounded_linear.isCont [OF bounded_linear_of_real]
 
-lemma isCont_power [simp]:
-  fixes f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
-  shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a"
+lemma isCont_power [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a"
+  for f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
   by (fact continuous_power)
 
-lemma isCont_setsum [simp]:
-  fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::topological_comm_monoid_add"
-  shows "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a"
+lemma isCont_setsum [simp]: "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a"
+  for f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::topological_comm_monoid_add"
   by (auto intro: continuous_setsum)
 
+
 subsection \<open>Uniform Continuity\<close>
 
 lemma uniformly_continuous_on_def:
@@ -2200,37 +2323,39 @@
     uniformity_dist filterlim_INF filterlim_principal eventually_inf_principal
   by (force simp: Ball_def uniformity_dist[symmetric] eventually_uniformity_metric)
 
-abbreviation isUCont :: "['a::metric_space \<Rightarrow> 'b::metric_space] \<Rightarrow> bool" where
-  "isUCont f \<equiv> uniformly_continuous_on UNIV f"
-
-lemma isUCont_def: "isUCont f = (\<forall>r>0. \<exists>s>0. \<forall>x y. dist x y < s \<longrightarrow> dist (f x) (f y) < r)"
+abbreviation isUCont :: "['a::metric_space \<Rightarrow> 'b::metric_space] \<Rightarrow> bool"
+  where "isUCont f \<equiv> uniformly_continuous_on UNIV f"
+
+lemma isUCont_def: "isUCont f \<longleftrightarrow> (\<forall>r>0. \<exists>s>0. \<forall>x y. dist x y < s \<longrightarrow> dist (f x) (f y) < r)"
   by (auto simp: uniformly_continuous_on_def dist_commute)
 
-lemma isUCont_isCont: "isUCont f ==> isCont f x"
+lemma isUCont_isCont: "isUCont f \<Longrightarrow> isCont f x"
   by (drule uniformly_continuous_imp_continuous) (simp add: continuous_on_eq_continuous_at)
 
 lemma uniformly_continuous_on_Cauchy:
-  fixes f::"'a::metric_space \<Rightarrow> 'b::metric_space"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space"
   assumes "uniformly_continuous_on S f" "Cauchy X" "\<And>n. X n \<in> S"
   shows "Cauchy (\<lambda>n. f (X n))"
   using assms
-  unfolding uniformly_continuous_on_def
-  apply -
+  apply (simp only: uniformly_continuous_on_def)
   apply (rule metric_CauchyI)
-  apply (drule_tac x=e in spec, safe)
-  apply (drule_tac e=d in metric_CauchyD, safe)
-  apply (rule_tac x=M in exI, simp)
+  apply (drule_tac x=e in spec)
+  apply safe
+  apply (drule_tac e=d in metric_CauchyD)
+   apply safe
+  apply (rule_tac x=M in exI)
+  apply simp
   done
 
-lemma isUCont_Cauchy:
-  "\<lbrakk>isUCont f; Cauchy X\<rbrakk> \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
+lemma isUCont_Cauchy: "isUCont f \<Longrightarrow> Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
   by (rule uniformly_continuous_on_Cauchy[where S=UNIV and f=f]) simp_all
 
 lemma (in bounded_linear) isUCont: "isUCont f"
-unfolding isUCont_def dist_norm
+  unfolding isUCont_def dist_norm
 proof (intro allI impI)
-  fix r::real assume r: "0 < r"
-  obtain K where K: "0 < K" and norm_le: "\<And>x. norm (f x) \<le> norm x * K"
+  fix r :: real
+  assume r: "0 < r"
+  obtain K where K: "0 < K" and norm_le: "norm (f x) \<le> norm x * K" for x
     using pos_bounded by blast
   show "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r"
   proof (rule exI, safe)
@@ -2246,7 +2371,7 @@
 qed
 
 lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
-by (rule isUCont [THEN isUCont_Cauchy])
+  by (rule isUCont [THEN isUCont_Cauchy])
 
 lemma LIM_less_bound:
   fixes f :: "real \<Rightarrow> real"
@@ -2268,16 +2393,21 @@
 proof -
   have "incseq f" unfolding incseq_Suc_iff by fact
   have "decseq g" unfolding decseq_Suc_iff by fact
-
-  { fix n
-    from \<open>decseq g\<close> have "g n \<le> g 0" by (rule decseqD) simp
-    with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] have "f n \<le> g 0" by auto }
+  have "f n \<le> g 0" for n
+  proof -
+    from \<open>decseq g\<close> have "g n \<le> g 0"
+      by (rule decseqD) simp
+    with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] show ?thesis
+      by auto
+  qed
   then obtain u where "f \<longlonglongrightarrow> u" "\<forall>i. f i \<le> u"
     using incseq_convergent[OF \<open>incseq f\<close>] by auto
-  moreover
-  { fix n
+  moreover have "f 0 \<le> g n" for n
+  proof -
     from \<open>incseq f\<close> have "f 0 \<le> f n" by (rule incseqD) simp
-    with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] have "f 0 \<le> g n" by simp }
+    with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] show ?thesis
+      by simp
+  qed
   then obtain l where "g \<longlonglongrightarrow> l" "\<forall>i. l \<le> g i"
     using decseq_convergent[OF \<open>decseq g\<close>] by auto
   moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF \<open>f \<longlonglongrightarrow> u\<close> \<open>g \<longlonglongrightarrow> l\<close>]]
@@ -2287,8 +2417,8 @@
 lemma Bolzano[consumes 1, case_names trans local]:
   fixes P :: "real \<Rightarrow> real \<Rightarrow> bool"
   assumes [arith]: "a \<le> b"
-  assumes trans: "\<And>a b c. \<lbrakk>P a b; P b c; a \<le> b; b \<le> c\<rbrakk> \<Longrightarrow> P a c"
-  assumes local: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<exists>d>0. \<forall>a b. a \<le> x \<and> x \<le> b \<and> b - a < d \<longrightarrow> P a b"
+    and trans: "\<And>a b c. P a b \<Longrightarrow> P b c \<Longrightarrow> a \<le> b \<Longrightarrow> b \<le> c \<Longrightarrow> P a c"
+    and local: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<exists>d>0. \<forall>a b. a \<le> x \<and> x \<le> b \<and> b - a < d \<longrightarrow> P a b"
   shows "P a b"
 proof -
   define bisect where "bisect =
@@ -2298,57 +2428,73 @@
     and u[simp]: "u 0 = b" "\<And>n. u (Suc n) = (if P (l n) ((l n + u n) / 2) then u n else (l n + u n) / 2)"
     by (simp_all add: l_def u_def bisect_def split: prod.split)
 
-  { fix n have "l n \<le> u n" by (induct n) auto } note this[simp]
+  have [simp]: "l n \<le> u n" for n by (induct n) auto
 
   have "\<exists>x. ((\<forall>n. l n \<le> x) \<and> l \<longlonglongrightarrow> x) \<and> ((\<forall>n. x \<le> u n) \<and> u \<longlonglongrightarrow> x)"
   proof (safe intro!: nested_sequence_unique)
-    fix n show "l n \<le> l (Suc n)" "u (Suc n) \<le> u n" by (induct n) auto
+    show "l n \<le> l (Suc n)" "u (Suc n) \<le> u n" for n
+      by (induct n) auto
   next
-    { fix n have "l n - u n = (a - b) / 2^n" by (induct n) (auto simp: field_simps) }
-    then show "(\<lambda>n. l n - u n) \<longlonglongrightarrow> 0" by (simp add: LIMSEQ_divide_realpow_zero)
+    have "l n - u n = (a - b) / 2^n" for n
+      by (induct n) (auto simp: field_simps)
+    then show "(\<lambda>n. l n - u n) \<longlonglongrightarrow> 0"
+      by (simp add: LIMSEQ_divide_realpow_zero)
   qed fact
-  then obtain x where x: "\<And>n. l n \<le> x" "\<And>n. x \<le> u n" and "l \<longlonglongrightarrow> x" "u \<longlonglongrightarrow> x" by auto
-  obtain d where "0 < d" and d: "\<And>a b. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> b - a < d \<Longrightarrow> P a b"
+  then obtain x where x: "\<And>n. l n \<le> x" "\<And>n. x \<le> u n" and "l \<longlonglongrightarrow> x" "u \<longlonglongrightarrow> x"
+    by auto
+  obtain d where "0 < d" and d: "a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> b - a < d \<Longrightarrow> P a b" for a b
     using \<open>l 0 \<le> x\<close> \<open>x \<le> u 0\<close> local[of x] by auto
 
   show "P a b"
   proof (rule ccontr)
     assume "\<not> P a b"
-    { fix n have "\<not> P (l n) (u n)"
-      proof (induct n)
-        case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto
-      qed (simp add: \<open>\<not> P a b\<close>) }
+    have "\<not> P (l n) (u n)" for n
+    proof (induct n)
+      case 0
+      then show ?case
+        by (simp add: \<open>\<not> P a b\<close>)
+    next
+      case (Suc n)
+      with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case
+        by auto
+    qed
     moreover
-    { have "eventually (\<lambda>n. x - d / 2 < l n) sequentially"
+    {
+      have "eventually (\<lambda>n. x - d / 2 < l n) sequentially"
         using \<open>0 < d\<close> \<open>l \<longlonglongrightarrow> x\<close> by (intro order_tendstoD[of _ x]) auto
       moreover have "eventually (\<lambda>n. u n < x + d / 2) sequentially"
         using \<open>0 < d\<close> \<open>u \<longlonglongrightarrow> x\<close> by (intro order_tendstoD[of _ x]) auto
       ultimately have "eventually (\<lambda>n. P (l n) (u n)) sequentially"
       proof eventually_elim
-        fix n assume "x - d / 2 < l n" "u n < x + d / 2"
+        case (elim n)
         from add_strict_mono[OF this] have "u n - l n < d" by simp
         with x show "P (l n) (u n)" by (rule d)
-      qed }
+      qed
+    }
     ultimately show False by simp
   qed
 qed
 
 lemma compact_Icc[simp, intro]: "compact {a .. b::real}"
 proof (cases "a \<le> b", rule compactI)
-  fix C assume C: "a \<le> b" "\<forall>t\<in>C. open t" "{a..b} \<subseteq> \<Union>C"
+  fix C
+  assume C: "a \<le> b" "\<forall>t\<in>C. open t" "{a..b} \<subseteq> \<Union>C"
   define T where "T = {a .. b}"
   from C(1,3) show "\<exists>C'\<subseteq>C. finite C' \<and> {a..b} \<subseteq> \<Union>C'"
   proof (induct rule: Bolzano)
     case (trans a b c)
-    then have *: "{a .. c} = {a .. b} \<union> {b .. c}" by auto
-    from trans obtain C1 C2 where "C1\<subseteq>C \<and> finite C1 \<and> {a..b} \<subseteq> \<Union>C1" "C2\<subseteq>C \<and> finite C2 \<and> {b..c} \<subseteq> \<Union>C2"
-      by (auto simp: *)
+    then have *: "{a..c} = {a..b} \<union> {b..c}"
+      by auto
+    with trans obtain C1 C2
+      where "C1\<subseteq>C" "finite C1" "{a..b} \<subseteq> \<Union>C1" "C2\<subseteq>C" "finite C2" "{b..c} \<subseteq> \<Union>C2"
+      by auto
     with trans show ?case
       unfolding * by (intro exI[of _ "C1 \<union> C2"]) auto
   next
     case (local x)
-    then have "x \<in> \<Union>C" using C by auto
-    with C(2) obtain c where "x \<in> c" "open c" "c \<in> C" by auto
+    with C have "x \<in> \<Union>C" by auto
+    with C(2) obtain c where "x \<in> c" "open c" "c \<in> C"
+      by auto
     then obtain e where "0 < e" "{x - e <..< x + e} \<subseteq> c"
       by (auto simp: open_dist dist_real_def subset_eq Ball_def abs_less_iff)
     with \<open>c \<in> C\<close> show ?case
@@ -2378,17 +2524,18 @@
 qed
 
 lemma open_Collect_positive:
- fixes f :: "'a::t2_space \<Rightarrow> real"
- assumes f: "continuous_on s f"
- shows "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. 0 < f x}"
- using continuous_on_open_invariant[THEN iffD1, OF f, rule_format, of "{0 <..}"]
- by (auto simp: Int_def field_simps)
+  fixes f :: "'a::t2_space \<Rightarrow> real"
+  assumes f: "continuous_on s f"
+  shows "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. 0 < f x}"
+  using continuous_on_open_invariant[THEN iffD1, OF f, rule_format, of "{0 <..}"]
+  by (auto simp: Int_def field_simps)
 
 lemma open_Collect_less_Int:
- fixes f g :: "'a::t2_space \<Rightarrow> real"
- assumes f: "continuous_on s f" and g: "continuous_on s g"
- shows "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. f x < g x}"
- using open_Collect_positive[OF continuous_on_diff[OF g f]] by (simp add: field_simps)
+  fixes f g :: "'a::t2_space \<Rightarrow> real"
+  assumes f: "continuous_on s f"
+    and g: "continuous_on s g"
+  shows "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. f x < g x}"
+  using open_Collect_positive[OF continuous_on_diff[OF g f]] by (simp add: field_simps)
 
 
 subsection \<open>Boundedness of continuous functions\<close>
@@ -2399,14 +2546,14 @@
   fixes f :: "real \<Rightarrow> 'a::linorder_topology"
   shows "a \<le> b \<Longrightarrow> \<forall>x::real. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow>
     \<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)"
-  using continuous_attains_sup[of "{a .. b}" f]
+  using continuous_attains_sup[of "{a..b}" f]
   by (auto simp add: continuous_at_imp_continuous_on Ball_def Bex_def)
 
 lemma isCont_eq_Lb:
   fixes f :: "real \<Rightarrow> 'a::linorder_topology"
   shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow>
     \<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> M \<le> f x) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)"
-  using continuous_attains_inf[of "{a .. b}" f]
+  using continuous_attains_inf[of "{a..b}" f]
   by (auto simp add: continuous_at_imp_continuous_on Ball_def Bex_def)
 
 lemma isCont_bounded:
@@ -2421,21 +2568,23 @@
   using isCont_eq_Ub[of a b f] by auto
 
 (*HOL style here: object-level formulations*)
-lemma IVT_objl: "(f(a::real) \<le> (y::real) & y \<le> f(b) & a \<le> b &
-      (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
-      --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
+lemma IVT_objl:
+  "(f a \<le> y \<and> y \<le> f b \<and> a \<le> b \<and> (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x)) \<longrightarrow>
+    (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = y)"
+  for a y :: real
   by (blast intro: IVT)
 
-lemma IVT2_objl: "(f(b::real) \<le> (y::real) & y \<le> f(a) & a \<le> b &
-      (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
-      --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
+lemma IVT2_objl:
+  "(f b \<le> y \<and> y \<le> f a \<and> a \<le> b \<and> (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x)) \<longrightarrow>
+    (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = y)"
+  for b y :: real
   by (blast intro: IVT2)
 
 lemma isCont_Lb_Ub:
   fixes f :: "real \<Rightarrow> real"
   assumes "a \<le> b" "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x"
   shows "\<exists>L M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> L \<le> f x \<and> f x \<le> M) \<and>
-               (\<forall>y. L \<le> y \<and> y \<le> M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> (f x = y)))"
+    (\<forall>y. L \<le> y \<and> y \<le> M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> (f x = y)))"
 proof -
   obtain M where M: "a \<le> M" "M \<le> b" "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> f M"
     using isCont_eq_Ub[OF assms] by auto
@@ -2446,22 +2595,26 @@
     apply (rule_tac x="f L" in exI)
     apply (rule_tac x="f M" in exI)
     apply (cases "L \<le> M")
-    apply (simp, metis order_trans)
-    apply (simp, metis order_trans)
+     apply simp
+     apply (metis order_trans)
+    apply simp
+    apply (metis order_trans)
     done
 qed
 
 
-text\<open>Continuity of inverse function\<close>
+text \<open>Continuity of inverse function.\<close>
 
 lemma isCont_inverse_function:
   fixes f g :: "real \<Rightarrow> real"
   assumes d: "0 < d"
-      and inj: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> g (f z) = z"
-      and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> isCont f z"
+    and inj: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> g (f z) = z"
+    and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> isCont f z"
   shows "isCont g (f x)"
 proof -
-  let ?A = "f (x - d)" and ?B = "f (x + d)" and ?D = "{x - d..x + d}"
+  let ?A = "f (x - d)"
+  let ?B = "f (x + d)"
+  let ?D = "{x - d..x + d}"
 
   have f: "continuous_on ?D f"
     using cont by (intro continuous_at_imp_continuous_on ballI) auto
@@ -2483,45 +2636,42 @@
 qed
 
 lemma isCont_inverse_function2:
-  fixes f g :: "real \<Rightarrow> real" shows
-  "\<lbrakk>a < x; x < b;
-    \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> g (f z) = z;
-    \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> isCont f z\<rbrakk>
-   \<Longrightarrow> isCont g (f x)"
-apply (rule isCont_inverse_function
-       [where f=f and d="min (x - a) (b - x)"])
-apply (simp_all add: abs_le_iff)
-done
+  fixes f g :: "real \<Rightarrow> real"
+  shows
+    "a < x \<Longrightarrow> x < b \<Longrightarrow>
+      \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> g (f z) = z \<Longrightarrow>
+      \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> isCont f z \<Longrightarrow> isCont g (f x)"
+  apply (rule isCont_inverse_function [where f=f and d="min (x - a) (b - x)"])
+  apply (simp_all add: abs_le_iff)
+  done
 
 (* need to rename second isCont_inverse *)
-
 lemma isCont_inv_fun:
   fixes f g :: "real \<Rightarrow> real"
-  shows "[| 0 < d; \<forall>z. \<bar>z - x\<bar> \<le> d --> g(f(z)) = z;
-         \<forall>z. \<bar>z - x\<bar> \<le> d --> isCont f z |]
-      ==> isCont g (f x)"
-by (rule isCont_inverse_function)
-
-text\<open>Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110\<close>
-lemma LIM_fun_gt_zero:
-  fixes f :: "real \<Rightarrow> real"
-  shows "f \<midarrow>c\<rightarrow> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> 0 < f x)"
-apply (drule (1) LIM_D, clarify)
-apply (rule_tac x = s in exI)
-apply (simp add: abs_less_iff)
-done
-
-lemma LIM_fun_less_zero:
-  fixes f :: "real \<Rightarrow> real"
-  shows "f \<midarrow>c\<rightarrow> l \<Longrightarrow> l < 0 \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> f x < 0)"
-apply (drule LIM_D [where r="-l"], simp, clarify)
-apply (rule_tac x = s in exI)
-apply (simp add: abs_less_iff)
-done
-
-lemma LIM_fun_not_zero:
-  fixes f :: "real \<Rightarrow> real"
-  shows "f \<midarrow>c\<rightarrow> l \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> f x \<noteq> 0)"
+  shows "0 < d \<Longrightarrow> (\<forall>z. \<bar>z - x\<bar> \<le> d \<longrightarrow> g (f z) = z) \<Longrightarrow>
+    \<forall>z. \<bar>z - x\<bar> \<le> d \<longrightarrow> isCont f z \<Longrightarrow> isCont g (f x)"
+  by (rule isCont_inverse_function)
+
+text \<open>Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110.\<close>
+lemma LIM_fun_gt_zero: "f \<midarrow>c\<rightarrow> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> 0 < f x)"
+  for f :: "real \<Rightarrow> real"
+  apply (drule (1) LIM_D)
+  apply clarify
+  apply (rule_tac x = s in exI)
+  apply (simp add: abs_less_iff)
+  done
+
+lemma LIM_fun_less_zero: "f \<midarrow>c\<rightarrow> l \<Longrightarrow> l < 0 \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> f x < 0)"
+  for f :: "real \<Rightarrow> real"
+  apply (drule LIM_D [where r="-l"])
+   apply simp
+  apply clarify
+  apply (rule_tac x = s in exI)
+  apply (simp add: abs_less_iff)
+  done
+
+lemma LIM_fun_not_zero: "f \<midarrow>c\<rightarrow> l \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> f x \<noteq> 0)"
+  for f :: "real \<Rightarrow> real"
   using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp add: neq_iff)
 
 end
--- a/src/HOL/List.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/List.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -3542,10 +3542,9 @@
 
     have "length ys = card (f ` {0 ..< size [x]})"
       using f_img by auto
-    then have "length ys = 1" by auto
-    moreover
+    then have *: "length ys = 1" by auto
     then have "f 0 = 0" using f_img by auto
-    ultimately show ?case using f_nth by (cases ys) auto
+    with * show ?case using f_nth by (cases ys) auto
   next
     case (3 x1 x2 xs)
     from "3.prems" obtain f where f_mono: "mono f"
@@ -5499,10 +5498,10 @@
   next
     let ?k = "length abs"
     case eq
-    hence "abs = bcs" "b = b'" using bs bs' by auto
-    moreover hence "(a, c') \<in> r"
+    hence *: "abs = bcs" "b = b'" using bs bs' by auto
+    hence "(a, c') \<in> r"
       using abr b'c'r assms unfolding trans_def by blast
-    ultimately show ?thesis using n n' unfolding lexn_conv as bs cs by auto
+    with * show ?thesis using n n' unfolding lexn_conv as bs cs by auto
   qed
 qed
 
@@ -5840,8 +5839,8 @@
   (is "?lhs \<longleftrightarrow> ?rhs")
 proof
   assume ?lhs
-  moreover hence "\<not> lexordp_eq ys xs" by induct simp_all
-  ultimately show ?rhs by(simp add: lexordp_into_lexordp_eq)
+  hence "\<not> lexordp_eq ys xs" by induct simp_all
+  with \<open>?lhs\<close> show ?rhs by (simp add: lexordp_into_lexordp_eq)
 next
   assume ?rhs
   hence "lexordp_eq xs ys" "\<not> lexordp_eq ys xs" by simp_all
--- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -71,9 +71,9 @@
 proof (induct s rule: finite_ne_induct)
   case (insert b s)
   assume *: "\<forall>x\<in>insert b s. \<forall>y\<in>insert b s. x \<le> y \<or> y \<le> x"
-  moreover then obtain u l where "l \<in> s" "\<forall>b\<in>s. l \<le> b" "u \<in> s" "\<forall>b\<in>s. b \<le> u"
+  then obtain u l where "l \<in> s" "\<forall>b\<in>s. l \<le> b" "u \<in> s" "\<forall>b\<in>s. b \<le> u"
     using insert by auto
-  ultimately show "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. a \<le> x" "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. x \<le> a"
+  with * show "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. a \<le> x" "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. x \<le> a"
     using *[rule_format, of b u] *[rule_format, of b l] by (metis insert_iff order.trans)+
 qed auto
 
--- a/src/HOL/Multivariate_Analysis/Conformal_Mappings.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Conformal_Mappings.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -1411,7 +1411,7 @@
   show no_df0: "norm(deriv f 0) \<le> 1"
     by (simp add: \<open>\<And>z. cmod z < 1 \<Longrightarrow> cmod (h z) \<le> 1\<close> df0)
   show "?Q" if "?P"
-  using that
+    using that
   proof
     assume "\<exists>z. cmod z < 1 \<and> z \<noteq> 0 \<and> cmod (f z) = cmod z"
     then obtain \<gamma> where \<gamma>: "cmod \<gamma> < 1" "\<gamma> \<noteq> 0" "cmod (f \<gamma>) = cmod \<gamma>" by blast
@@ -1424,9 +1424,9 @@
       by (metis add_diff_cancel_left' diff_diff_eq2 le_less_trans noh_le norm_triangle_ineq4)
     ultimately obtain c where c: "\<And>z. norm z < 1 \<Longrightarrow> h z = c"
       using Schwarz2 [OF holh, of "1 - norm \<gamma>" \<gamma>, unfolded constant_on_def] \<gamma> by auto
-    moreover then have "norm c = 1"
+    then have "norm c = 1"
       using \<gamma> by force
-    ultimately show ?thesis
+    with c show ?thesis
       using fz_eq by auto
   next
     assume [simp]: "cmod (deriv f 0) = 1"
--- a/src/HOL/Multivariate_Analysis/Gamma.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Gamma.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -1470,10 +1470,10 @@
   finally show "(rGamma x) = 0 \<longleftrightarrow> (\<exists>n. x = - real_of_nat n)" by simp
 next
   fix x :: real assume "\<And>n. x \<noteq> - of_nat n"
-  hence "complex_of_real x \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  hence x: "complex_of_real x \<notin> \<int>\<^sub>\<le>\<^sub>0"
     by (subst of_real_in_nonpos_Ints_iff) (auto elim!: nonpos_Ints_cases')
-  moreover from this have "x \<noteq> 0" by auto
-  ultimately have "(rGamma has_field_derivative - rGamma x * Digamma x) (at x)"
+  then have "x \<noteq> 0" by auto
+  with x have "(rGamma has_field_derivative - rGamma x * Digamma x) (at x)"
     by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex
                   simp: Polygamma_of_real rGamma_real_def [abs_def])
   thus "let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (x + of_nat k))
@@ -2248,8 +2248,8 @@
   hence "(g has_field_derivative 0) (at z)" for z using g_g'[of z] by simp
   hence "\<exists>c'. \<forall>z\<in>UNIV. g z = c'" by (intro has_field_derivative_zero_constant) simp_all
   then obtain c' where c: "\<And>z. g z = c'" by (force simp: dist_0_norm)
-  moreover from this[of 0] have "c' = pi" unfolding g_def by simp
-  ultimately have "g z = pi" by simp
+  from this[of 0] have "c' = pi" unfolding g_def by simp
+  with c have "g z = pi" by simp
 
   show ?thesis
   proof (cases "z \<in> \<int>")
--- a/src/HOL/Multivariate_Analysis/Integration.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Integration.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -10168,14 +10168,14 @@
   then have x: "{integral s (f k) |k. True} = range x"
     by auto
 
-  have "g integrable_on s \<and> (\<lambda>k. integral s (f k)) \<longlonglongrightarrow> integral s g"
+  have *: "g integrable_on s \<and> (\<lambda>k. integral s (f k)) \<longlonglongrightarrow> integral s g"
   proof (intro monotone_convergence_increasing allI ballI assms)
     show "bounded {integral s (f k) |k. True}"
       unfolding x by (rule convergent_imp_bounded) fact
   qed (auto intro: f)
-  moreover then have "integral s g = x'"
+  then have "integral s g = x'"
     by (intro LIMSEQ_unique[OF _ \<open>x \<longlonglongrightarrow> x'\<close>]) (simp add: x_eq)
-  ultimately show ?thesis
+  with * show ?thesis
     by (simp add: has_integral_integral)
 qed
 
@@ -11559,18 +11559,16 @@
     with \<open>open W\<close>
     have "\<exists>X0 Y. open X0 \<and> open Y \<and> x0 \<in> X0 \<and> y \<in> Y \<and> X0 \<times> Y \<subseteq> W"
       by (rule open_prod_elim) blast
-  } then obtain X0 Y where
-    "\<forall>y \<in> K. open (X0 y) \<and> open (Y y) \<and> x0 \<in> X0 y \<and> y \<in> Y y \<and> X0 y \<times> Y y \<subseteq> W"
+  }
+  then obtain X0 Y where
+    *: "\<forall>y \<in> K. open (X0 y) \<and> open (Y y) \<and> x0 \<in> X0 y \<and> y \<in> Y y \<and> X0 y \<times> Y y \<subseteq> W"
     by metis
-  moreover
-  then have "\<forall>t\<in>Y ` K. open t" "K \<subseteq> \<Union>(Y ` K)" by auto
-  with \<open>compact K\<close> obtain CC where "CC \<subseteq> Y ` K" "finite CC" "K \<subseteq> \<Union>CC"
+  from * have "\<forall>t\<in>Y ` K. open t" "K \<subseteq> \<Union>(Y ` K)" by auto
+  with \<open>compact K\<close> obtain CC where CC: "CC \<subseteq> Y ` K" "finite CC" "K \<subseteq> \<Union>CC"
     by (rule compactE)
-  moreover
-  then obtain c where c:
-    "\<And>C. C \<in> CC \<Longrightarrow> c C \<in> K \<and> C = Y (c C)"
+  then obtain c where c: "\<And>C. C \<in> CC \<Longrightarrow> c C \<in> K \<and> C = Y (c C)"
     by (force intro!: choice)
-  ultimately show ?thesis
+  with * CC show ?thesis
     by (force intro!: exI[where x="\<Inter>C\<in>CC. X0 (c C)"])
 qed
 
--- a/src/HOL/Multivariate_Analysis/Path_Connected.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -3057,11 +3057,11 @@
     then have ftendsw: "((\<lambda>n. f (z n)) \<circ> K) \<longlonglongrightarrow> w"
       by (metis LIMSEQ_subseq_LIMSEQ fun.map_cong0 fz tendsw)
     have zKs: "\<And>n. (z o K) n \<in> S" by (simp add: zs)
-    have "f \<circ> z = \<xi>"  "(\<lambda>n. f (z n)) = \<xi>"
+    have fz: "f \<circ> z = \<xi>"  "(\<lambda>n. f (z n)) = \<xi>"
       using fz by auto
-    moreover then have "(\<xi> \<circ> K) \<longlonglongrightarrow> f y"
+    then have "(\<xi> \<circ> K) \<longlonglongrightarrow> f y"
       by (metis (no_types) Klim zKs y contf comp_assoc continuous_on_closure_sequentially)
-    ultimately have wy: "w = f y" using fz LIMSEQ_unique ftendsw by auto
+    with fz have wy: "w = f y" using fz LIMSEQ_unique ftendsw by auto
     have rle: "r \<le> norm (f y - f a)"
       apply (rule le_no)
       using w wy oint
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -4343,10 +4343,11 @@
   have "F \<noteq> bot"
     unfolding F_def
   proof (rule INF_filter_not_bot)
-    fix X assume "X \<subseteq> insert U A" "finite X"
-    moreover with A(2)[THEN spec, of "X - {U}"] have "U \<inter> \<Inter>(X - {U}) \<noteq> {}"
+    fix X
+    assume X: "X \<subseteq> insert U A" "finite X"
+    with A(2)[THEN spec, of "X - {U}"] have "U \<inter> \<Inter>(X - {U}) \<noteq> {}"
       by auto
-    ultimately show "(INF a:X. principal a) \<noteq> bot"
+    with X show "(INF a:X. principal a) \<noteq> bot"
       by (auto simp add: INF_principal_finite principal_eq_bot_iff)
   qed
   moreover
@@ -6601,13 +6602,12 @@
       then obtain xs where xs: "xs \<longlonglongrightarrow> x" "\<And>n. xs n \<in> X"
         by (auto simp: closure_sequential)
       from continuous_on_tendsto_compose[OF cont_h xs(1)] xs(2) Y
-      have "(\<lambda>x. f (xs x)) \<longlonglongrightarrow> h x"
+      have hx: "(\<lambda>x. f (xs x)) \<longlonglongrightarrow> h x"
         by (auto simp: set_mp extension)
-      moreover
       then have "(\<lambda>x. f (xs x)) \<longlonglongrightarrow> y x"
         using \<open>x \<notin> X\<close> not_eventuallyD xs(2)
         by (force intro!: filterlim_compose[OF y[OF \<open>x \<in> closure X\<close>]] simp: filterlim_at xs)
-      ultimately have "h x = y x" by (rule LIMSEQ_unique)
+      with hx have "h x = y x" by (rule LIMSEQ_unique)
     } then
     have "h x = ?g x"
       using extension by auto
--- a/src/HOL/Number_Theory/Factorial_Ring.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Number_Theory/Factorial_Ring.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -732,9 +732,9 @@
   from prime_factorization_exists'[OF assms(1)] guess A . note A = this
   moreover from A and assms have "A \<noteq> {#}" by auto
   then obtain x where "x \<in># A" by blast
-  with A(1) have "x dvd msetprod A" "is_prime x" by (auto simp: dvd_msetprod)
-  moreover from this A have "x dvd a" by simp
-  ultimately show ?thesis by blast
+  with A(1) have *: "x dvd msetprod A" "is_prime x" by (auto simp: dvd_msetprod)
+  with A have "x dvd a" by simp
+  with * show ?thesis by blast
 qed
 
 lemma prime_divisors_induct [case_names zero unit factor]:
--- a/src/HOL/Number_Theory/Polynomial_Factorial.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Number_Theory/Polynomial_Factorial.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -1207,7 +1207,7 @@
   shows "irreducible [:a,b:]"
 proof (rule irreducibleI)
   fix p q assume pq: "[:a,b:] = p * q"
-  also from this assms have "degree \<dots> = degree p + degree q" 
+  also from pq assms have "degree \<dots> = degree p + degree q" 
     by (intro degree_mult_eq) auto
   finally have "degree p = 0 \<or> degree q = 0" using assms by auto
   with assms pq show "is_unit p \<or> is_unit q"
--- a/src/HOL/Probability/Distributions.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/Distributions.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -164,8 +164,8 @@
 lemma nn_integral_erlang_density:
   assumes [arith]: "0 < l"
   shows "(\<integral>\<^sup>+ x. ennreal (erlang_density k l x) * indicator {.. a} x \<partial>lborel) = erlang_CDF k l a"
-proof cases
-  assume [arith]: "0 \<le> a"
+proof (cases "0 \<le> a")
+  case [arith]: True
   have eq: "\<And>x. indicator {0..a} (x / l) = indicator {0..a*l} x"
     by (simp add: field_simps split: split_indicator)
   have "(\<integral>\<^sup>+x. ennreal (erlang_density k l x) * indicator {.. a} x \<partial>lborel) =
@@ -182,10 +182,10 @@
     by (auto simp: erlang_CDF_def)
   finally show ?thesis .
 next
-  assume "\<not> 0 \<le> a"
-  moreover then have "(\<integral>\<^sup>+ x. ennreal (erlang_density k l x) * indicator {.. a} x \<partial>lborel) = (\<integral>\<^sup>+x. 0 \<partial>(lborel::real measure))"
+  case False
+  then have "(\<integral>\<^sup>+ x. ennreal (erlang_density k l x) * indicator {.. a} x \<partial>lborel) = (\<integral>\<^sup>+x. 0 \<partial>(lborel::real measure))"
     by (intro nn_integral_cong) (auto simp: erlang_density_def)
-  ultimately show ?thesis
+  with False show ?thesis
     by (simp add: erlang_CDF_def)
 qed
 
--- a/src/HOL/Probability/Embed_Measure.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/Embed_Measure.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -33,10 +33,10 @@
   finally show "f ` space M - s \<in> {f ` A |A. A \<in> sets M}" .
 next
   fix A :: "nat \<Rightarrow> _" assume "range A \<subseteq> {f ` A |A. A \<in> sets M}"
-  then obtain A' where "\<And>i. A i = f ` A' i" "\<And>i. A' i \<in> sets M"
+  then obtain A' where A': "\<And>i. A i = f ` A' i" "\<And>i. A' i \<in> sets M"
     by (auto simp: subset_eq choice_iff)
-  moreover from this have "(\<Union>x. f ` A' x) = f ` (\<Union>x. A' x)" by blast
-  ultimately show "(\<Union>i. A i) \<in> {f ` A |A. A \<in> sets M}"
+  then have "(\<Union>x. f ` A' x) = f ` (\<Union>x. A' x)" by blast
+  with A' show "(\<Union>i. A i) \<in> {f ` A |A. A \<in> sets M}"
     by simp blast
 qed (auto dest: sets.sets_into_space)
 
@@ -81,7 +81,7 @@
 proof-
   {
     fix A assume A: "A \<in> sets M"
-    also from this have "A = A \<inter> space M" by auto
+    also from A have "A = A \<inter> space M" by auto
     also have "... = f -` f ` A \<inter> space M" using A assms
       by (auto dest: inj_onD sets.sets_into_space)
     finally have "f -` f ` A \<inter> space M \<in> sets M" .
--- a/src/HOL/Probability/Finite_Product_Measure.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/Finite_Product_Measure.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -1133,10 +1133,10 @@
 lemma (in product_sigma_finite) distr_component:
   "distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x) = Pi\<^sub>M {i} M" (is "?D = ?P")
 proof (intro PiM_eqI)
-  fix A assume "\<And>ia. ia \<in> {i} \<Longrightarrow> A ia \<in> sets (M ia)"
-  moreover then have "(\<lambda>x. \<lambda>i\<in>{i}. x) -` Pi\<^sub>E {i} A \<inter> space (M i) = A i"
+  fix A assume A: "\<And>ia. ia \<in> {i} \<Longrightarrow> A ia \<in> sets (M ia)"
+  then have "(\<lambda>x. \<lambda>i\<in>{i}. x) -` Pi\<^sub>E {i} A \<inter> space (M i) = A i"
     by (auto dest: sets.sets_into_space)
-  ultimately show "emeasure (distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x)) (Pi\<^sub>E {i} A) = (\<Prod>i\<in>{i}. emeasure (M i) (A i))"
+  with A show "emeasure (distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x)) (Pi\<^sub>E {i} A) = (\<Prod>i\<in>{i}. emeasure (M i) (A i))"
     by (subst emeasure_distr) (auto intro!: sets_PiM_I_finite measurable_restrict)
 qed simp_all
 
--- a/src/HOL/Probability/Giry_Monad.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/Giry_Monad.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -272,28 +272,28 @@
   ultimately show ?case by simp
 next
   case (set B)
-  moreover then have "(\<lambda>M'. \<integral>\<^sup>+M''. indicator B M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. emeasure M' B) \<in> ?B"
+  then have "(\<lambda>M'. \<integral>\<^sup>+M''. indicator B M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. emeasure M' B) \<in> ?B"
     by (intro measurable_cong nn_integral_indicator) (simp add: space_subprob_algebra)
-  ultimately show ?case
+  with set show ?case
     by (simp add: measurable_emeasure_subprob_algebra)
 next
   case (mult f c)
-  moreover then have "(\<lambda>M'. \<integral>\<^sup>+M''. c * f M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. c * \<integral>\<^sup>+M''. f M'' \<partial>M') \<in> ?B"
+  then have "(\<lambda>M'. \<integral>\<^sup>+M''. c * f M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. c * \<integral>\<^sup>+M''. f M'' \<partial>M') \<in> ?B"
     by (intro measurable_cong nn_integral_cmult) (auto simp add: space_subprob_algebra)
-  ultimately show ?case
+  with mult show ?case
     by simp
 next
   case (add f g)
-  moreover then have "(\<lambda>M'. \<integral>\<^sup>+M''. f M'' + g M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. (\<integral>\<^sup>+M''. f M'' \<partial>M') + (\<integral>\<^sup>+M''. g M'' \<partial>M')) \<in> ?B"
+  then have "(\<lambda>M'. \<integral>\<^sup>+M''. f M'' + g M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. (\<integral>\<^sup>+M''. f M'' \<partial>M') + (\<integral>\<^sup>+M''. g M'' \<partial>M')) \<in> ?B"
     by (intro measurable_cong nn_integral_add) (auto simp add: space_subprob_algebra)
-  ultimately show ?case
+  with add show ?case
     by (simp add: ac_simps)
 next
   case (seq F)
-  moreover then have "(\<lambda>M'. \<integral>\<^sup>+M''. (SUP i. F i) M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. SUP i. (\<integral>\<^sup>+M''. F i M'' \<partial>M')) \<in> ?B"
+  then have "(\<lambda>M'. \<integral>\<^sup>+M''. (SUP i. F i) M'' \<partial>M') \<in> ?B \<longleftrightarrow> (\<lambda>M'. SUP i. (\<integral>\<^sup>+M''. F i M'' \<partial>M')) \<in> ?B"
     unfolding SUP_apply
     by (intro measurable_cong nn_integral_monotone_convergence_SUP) (auto simp add: space_subprob_algebra)
-  ultimately show ?case
+  with seq show ?case
     by (simp add: ac_simps)
 qed
 
@@ -793,10 +793,10 @@
     by simp
 next
   case (set A)
-  moreover with M have "(\<integral>\<^sup>+ M'. integral\<^sup>N M' (indicator A) \<partial>M) = (\<integral>\<^sup>+ M'. emeasure M' A \<partial>M)"
+  with M have "(\<integral>\<^sup>+ M'. integral\<^sup>N M' (indicator A) \<partial>M) = (\<integral>\<^sup>+ M'. emeasure M' A \<partial>M)"
     by (intro nn_integral_cong nn_integral_indicator)
        (auto simp: space_subprob_algebra dest!: sets_eq_imp_space_eq)
-  ultimately show ?case
+  with set show ?case
     using M by (simp add: emeasure_join)
 next
   case (mult f c)
--- a/src/HOL/Probability/Infinite_Product_Measure.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/Infinite_Product_Measure.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -29,14 +29,18 @@
   note * = this
 
   have "emeasure (PiM I M) (emb I J (PiE J X)) = (\<Prod>i\<in>J. M i (X i))"
-  proof cases
-    assume "\<not> (J \<noteq> {} \<or> I = {})"
-    then obtain i where "J = {}" "i \<in> I" by auto
-    moreover then have "emb I {} {\<lambda>x. undefined} = emb I {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))"
+  proof (cases "J \<noteq> {} \<or> I = {}")
+    case False
+    then obtain i where i: "J = {}" "i \<in> I" by auto
+    then have "emb I {} {\<lambda>x. undefined} = emb I {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))"
       by (auto simp: space_PiM prod_emb_def)
-    ultimately show ?thesis
+    with i show ?thesis
       by (simp add: * M.emeasure_space_1)
-  qed (simp add: *[OF _ assms X])
+  next
+    case True
+    then show ?thesis
+      by (simp add: *[OF _ assms X])
+  qed
   with assms show "emeasure (distr (Pi\<^sub>M I M) (Pi\<^sub>M J M) (\<lambda>x. restrict x J)) (Pi\<^sub>E J X) = (\<Prod>i\<in>J. emeasure (M i) (X i))"
     by (subst emeasure_distr_restrict[OF _ refl]) (auto intro!: sets_PiM_I_finite X)
 qed (insert assms, auto)
--- a/src/HOL/Probability/Lebesgue_Integral_Substitution.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/Lebesgue_Integral_Substitution.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -323,9 +323,9 @@
     and "(LBINT x. f x * indicator {g a..g b} x) = (LBINT x. f (g x) * g' x * indicator {a..b} x)"
 proof-
   from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on)
-  from this and contg' have Mg: "set_borel_measurable borel {a..b} g" and
-                             Mg': "set_borel_measurable borel {a..b} g'"
-      by (simp_all only: set_measurable_continuous_on_ivl)
+  with contg' have Mg: "set_borel_measurable borel {a..b} g"
+    and Mg': "set_borel_measurable borel {a..b} g'"
+    by (simp_all only: set_measurable_continuous_on_ivl)
   from derivg derivg_nonneg have monog: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
     by (rule deriv_nonneg_imp_mono) simp_all
 
@@ -341,18 +341,18 @@
           enn2real (\<integral>\<^sup>+ x. ennreal (f x) * indicator {g a..g b} x \<partial>lborel) -
           enn2real (\<integral>\<^sup>+ x. ennreal (- (f x)) * indicator {g a..g b} x \<partial>lborel)" using integrable
     by (subst real_lebesgue_integral_def) (simp_all add: nn_integral_set_ennreal mult.commute)
-  also have "(\<integral>\<^sup>+x. ennreal (f x) * indicator {g a..g b} x \<partial>lborel) =
-               (\<integral>\<^sup>+x. ennreal (f x * indicator {g a..g b} x) \<partial>lborel)"
+  also have *: "(\<integral>\<^sup>+x. ennreal (f x) * indicator {g a..g b} x \<partial>lborel) =
+      (\<integral>\<^sup>+x. ennreal (f x * indicator {g a..g b} x) \<partial>lborel)"
     by (intro nn_integral_cong) (simp split: split_indicator)
-  also with M1 have A: "(\<integral>\<^sup>+ x. ennreal (f x * indicator {g a..g b} x) \<partial>lborel) =
+  also from M1 * have A: "(\<integral>\<^sup>+ x. ennreal (f x * indicator {g a..g b} x) \<partial>lborel) =
                             (\<integral>\<^sup>+ x. ennreal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel)"
     by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \<open>a \<le> b\<close>])
        (auto simp: nn_integral_set_ennreal mult.commute)
-  also have "(\<integral>\<^sup>+ x. ennreal (- (f x)) * indicator {g a..g b} x \<partial>lborel) =
-               (\<integral>\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \<partial>lborel)"
+  also have **: "(\<integral>\<^sup>+ x. ennreal (- (f x)) * indicator {g a..g b} x \<partial>lborel) =
+      (\<integral>\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \<partial>lborel)"
     by (intro nn_integral_cong) (simp split: split_indicator)
-  also with M2 have B: "(\<integral>\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \<partial>lborel) =
-                            (\<integral>\<^sup>+ x. ennreal (- (f (g x)) * g' x * indicator {a..b} x) \<partial>lborel)"
+  also from M2 ** have B: "(\<integral>\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \<partial>lborel) =
+        (\<integral>\<^sup>+ x. ennreal (- (f (g x)) * g' x * indicator {a..b} x) \<partial>lborel)"
     by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \<open>a \<le> b\<close>])
        (auto simp: nn_integral_set_ennreal mult.commute)
 
--- a/src/HOL/Probability/Lebesgue_Measure.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/Lebesgue_Measure.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -75,9 +75,9 @@
           fix i assume "i \<in> S - {m}"
           then have i: "i \<in> S" "i \<noteq> m" by auto
           { assume i': "l i < r i" "l i < r m"
-            moreover with \<open>finite S\<close> i m have "l m \<le> l i"
+            with \<open>finite S\<close> i m have "l m \<le> l i"
               by auto
-            ultimately have "{l i <.. r i} \<inter> {l m <.. r m} \<noteq> {}"
+            with i' have "{l i <.. r i} \<inter> {l m <.. r m} \<noteq> {}"
               by auto
             then have False
               using disjoint_family_onD[OF disj, of i m] i by auto }
@@ -852,9 +852,9 @@
   shows "(f has_integral r) UNIV"
 using f proof (induct f arbitrary: r rule: borel_measurable_induct_real)
   case (set A)
-  moreover then have "((\<lambda>x. 1) has_integral measure lborel A) A"
+  then have "((\<lambda>x. 1) has_integral measure lborel A) A"
     by (intro has_integral_measure_lborel) (auto simp: ennreal_indicator)
-  ultimately show ?case
+  with set show ?case
     by (simp add: ennreal_indicator measure_def) (simp add: indicator_def)
 next
   case (mult g c)
@@ -868,13 +868,12 @@
     by (auto intro!: has_integral_cmult_real)
 next
   case (add g h)
-  moreover
   then have "(\<integral>\<^sup>+ x. h x + g x \<partial>lborel) = (\<integral>\<^sup>+ x. h x \<partial>lborel) + (\<integral>\<^sup>+ x. g x \<partial>lborel)"
     by (simp add: nn_integral_add)
   with add obtain a b where "0 \<le> a" "0 \<le> b" "(\<integral>\<^sup>+ x. h x \<partial>lborel) = ennreal a" "(\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal b" "r = a + b"
     by (cases "\<integral>\<^sup>+ x. h x \<partial>lborel" "\<integral>\<^sup>+ x. g x \<partial>lborel" rule: ennreal2_cases)
        (auto simp: add_top nn_integral_add top_add ennreal_plus[symmetric] simp del: ennreal_plus)
-  ultimately show ?case
+  with add show ?case
     by (auto intro!: has_integral_add)
 next
   case (seq U)
@@ -1020,8 +1019,8 @@
   fixes A :: "'a::euclidean_space set"
   assumes A[measurable]: "A \<in> sets borel" and [simp]: "0 \<le> r"
   shows "((\<lambda>x. 1) has_integral r) A \<longleftrightarrow> emeasure lborel A = ennreal r"
-proof cases
-  assume emeasure_A: "emeasure lborel A = \<infinity>"
+proof (cases "emeasure lborel A = \<infinity>")
+  case emeasure_A: True
   have "\<not> (\<lambda>x. 1::real) integrable_on A"
   proof
     assume int: "(\<lambda>x. 1::real) integrable_on A"
@@ -1035,10 +1034,10 @@
   with emeasure_A show ?thesis
     by auto
 next
-  assume "emeasure lborel A \<noteq> \<infinity>"
-  moreover then have "((\<lambda>x. 1) has_integral measure lborel A) A"
+  case False
+  then have "((\<lambda>x. 1) has_integral measure lborel A) A"
     by (simp add: has_integral_measure_lborel less_top)
-  ultimately show ?thesis
+  with False show ?thesis
     by (auto simp: emeasure_eq_ennreal_measure has_integral_unique)
 qed
 
--- a/src/HOL/Probability/Levy.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/Levy.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -391,23 +391,22 @@
                 continuous_intros ballI M'.isCont_char continuous_intros)
     have "cmod (CLBINT t:{-d/2..d/2}. 1 - char M' t) \<le> LBINT t:{-d/2..d/2}. cmod (1 - char M' t)"
       using integral_norm_bound[OF 2] by simp
-    also have "\<dots> \<le> LBINT t:{-d/2..d/2}. \<epsilon> / 4"
+    also have 4: "\<dots> \<le> LBINT t:{-d/2..d/2}. \<epsilon> / 4"
       apply (rule integral_mono [OF 3])
-      apply (simp add: emeasure_lborel_Icc_eq)
-      apply (case_tac "x \<in> {-d/2..d/2}", auto)
+       apply (simp add: emeasure_lborel_Icc_eq)
+      apply (case_tac "x \<in> {-d/2..d/2}")
+       apply auto
       apply (subst norm_minus_commute)
       apply (rule less_imp_le)
       apply (rule d1 [simplified])
-      using d0 by auto
-    also with d0 have "\<dots> = d * \<epsilon> / 4"
+      using d0 apply auto
+      done
+    also from d0 4 have "\<dots> = d * \<epsilon> / 4"
       by simp
     finally have bound: "cmod (CLBINT t:{-d/2..d/2}. 1 - char M' t) \<le> d * \<epsilon> / 4" .
-    { fix n x
-      have "cmod (1 - char (M n) x) \<le> 2"
-        by (rule order_trans [OF norm_triangle_ineq4], auto simp add: Mn.cmod_char_le_1)
-    } note bd1 = this
-    have "(\<lambda>n. CLBINT t:{-d/2..d/2}. 1 - char (M n) t) \<longlonglongrightarrow> (CLBINT t:{-d/2..d/2}. 1 - char M' t)"
-      using bd1
+    have "cmod (1 - char (M n) x) \<le> 2" for n x
+      by (rule order_trans [OF norm_triangle_ineq4], auto simp add: Mn.cmod_char_le_1)
+    then have "(\<lambda>n. CLBINT t:{-d/2..d/2}. 1 - char (M n) t) \<longlonglongrightarrow> (CLBINT t:{-d/2..d/2}. 1 - char M' t)"
       apply (intro integral_dominated_convergence[where w="\<lambda>x. indicator {-d/2..d/2} x *\<^sub>R 2"])
       apply (auto intro!: char_conv tendsto_intros
                   simp: emeasure_lborel_Icc_eq
--- a/src/HOL/Probability/Measure_Space.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/Measure_Space.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -1862,7 +1862,7 @@
     using f unfolding Pi_def bij_betw_def by auto
   fix X assume "X \<in> sets (distr (count_space A) (count_space B) f)"
   then have X: "X \<in> sets (count_space B)" by auto
-  moreover then have "f -` X \<inter> A = the_inv_into A f ` X"
+  moreover from X have "f -` X \<inter> A = the_inv_into A f ` X"
     using f by (auto simp: bij_betw_def subset_image_iff image_iff the_inv_into_f_f intro: the_inv_into_f_f[symmetric])
   moreover have "inj_on (the_inv_into A f) B"
     using X f by (auto simp: bij_betw_def inj_on_the_inv_into)
@@ -1932,8 +1932,8 @@
 lemma emeasure_restrict_space:
   assumes "\<Omega> \<inter> space M \<in> sets M" "A \<subseteq> \<Omega>"
   shows "emeasure (restrict_space M \<Omega>) A = emeasure M A"
-proof cases
-  assume "A \<in> sets M"
+proof (cases "A \<in> sets M")
+  case True
   show ?thesis
   proof (rule emeasure_measure_of[OF restrict_space_def])
     show "op \<inter> \<Omega> ` sets M \<subseteq> Pow (\<Omega> \<inter> space M)" "A \<in> sets (restrict_space M \<Omega>)"
@@ -1951,10 +1951,10 @@
     qed
   qed
 next
-  assume "A \<notin> sets M"
-  moreover with assms have "A \<notin> sets (restrict_space M \<Omega>)"
+  case False
+  with assms have "A \<notin> sets (restrict_space M \<Omega>)"
     by (simp add: sets_restrict_space_iff)
-  ultimately show ?thesis
+  with False show ?thesis
     by (simp add: emeasure_notin_sets)
 qed
 
--- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -1734,7 +1734,7 @@
 proof (rule measure_eqI)
   fix X assume "X \<in> sets M"
   then have X: "X \<subseteq> A" by auto
-  moreover with A have "countable X" by (auto dest: countable_subset)
+  moreover from A X have "countable X" by (auto dest: countable_subset)
   ultimately have
     "emeasure M X = (\<integral>\<^sup>+a. emeasure M {a} \<partial>count_space X)"
     "emeasure N X = (\<integral>\<^sup>+a. emeasure N {a} \<partial>count_space X)"
--- a/src/HOL/Probability/PMF_Impl.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/PMF_Impl.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -402,9 +402,9 @@
                (\<lambda>x. listsum (map snd (filter (\<lambda>z. fst z = x) xs)))"
 proof -
   from assms have wf: "pmf_of_list_wf xs" by (intro pmf_of_list_wfI) force
-  moreover from this assms have "set_pmf (pmf_of_list xs) = fst ` set xs"
+  with assms have "set_pmf (pmf_of_list xs) = fst ` set xs"
     by (intro set_pmf_of_list_eq) auto
-  ultimately show ?thesis
+  with wf show ?thesis
     by (intro mapping_of_pmfI) (auto simp: lookup_tabulate pmf_pmf_of_list)
 qed
 
--- a/src/HOL/Probability/Probability_Mass_Function.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/Probability_Mass_Function.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -1104,17 +1104,17 @@
   assume eq: "\<forall>x\<in>set_pmf p. \<forall>y\<in>set_pmf q. R x y \<longrightarrow> measure p {x. R x y} = measure q {y. R x y}"
 
   show "measure p C = measure q C"
-  proof cases
-    assume "p \<inter> C = {}"
-    moreover then have "q \<inter> C = {}"
+  proof (cases "p \<inter> C = {}")
+    case True
+    then have "q \<inter> C = {}"
       using quotient_rel_set_disjoint[OF assms C R] by simp
-    ultimately show ?thesis
+    with True show ?thesis
       unfolding measure_pmf_zero_iff[symmetric] by simp
   next
-    assume "p \<inter> C \<noteq> {}"
-    moreover then have "q \<inter> C \<noteq> {}"
+    case False
+    then have "q \<inter> C \<noteq> {}"
       using quotient_rel_set_disjoint[OF assms C R] by simp
-    ultimately obtain x y where in_set: "x \<in> set_pmf p" "y \<in> set_pmf q" and in_C: "x \<in> C" "y \<in> C"
+    with False obtain x y where in_set: "x \<in> set_pmf p" "y \<in> set_pmf q" and in_C: "x \<in> C" "y \<in> C"
       by auto
     then have "R x y"
       using in_quotient_imp_in_rel[of UNIV ?R C x y] C assms
--- a/src/HOL/Probability/Radon_Nikodym.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/Radon_Nikodym.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -840,11 +840,11 @@
   assumes "f \<in> borel_measurable M" "density M f = N"
   shows "density M (RN_deriv M N) = N"
 proof -
-  have "\<exists>f. f \<in> borel_measurable M \<and> density M f = N"
+  have *: "\<exists>f. f \<in> borel_measurable M \<and> density M f = N"
     using assms by auto
-  moreover then have "density M (SOME f. f \<in> borel_measurable M \<and> density M f = N) = N"
+  then have "density M (SOME f. f \<in> borel_measurable M \<and> density M f = N) = N"
     by (rule someI2_ex) auto
-  ultimately show ?thesis
+  with * show ?thesis
     by (auto simp: RN_deriv_def)
 qed
 
--- a/src/HOL/Probability/Random_Permutations.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/Random_Permutations.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -52,9 +52,9 @@
 by (force, simp_all)
 termination proof (relation "Wellfounded.measure (\<lambda>(_,_,A). card A)")
   fix A :: "'a set" and f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" and x :: 'b and y :: 'a
-  assume "finite A" "A \<noteq> {}" "y \<in> set_pmf (pmf_of_set A)"
-  moreover from this have "card A > 0" by (simp add: card_gt_0_iff)
-  ultimately show "((f, f y x, A - {y}), f, x, A) \<in> Wellfounded.measure (\<lambda>(_, _, A). card A)"
+  assume A: "finite A" "A \<noteq> {}" "y \<in> set_pmf (pmf_of_set A)"
+  then have "card A > 0" by (simp add: card_gt_0_iff)
+  with A show "((f, f y x, A - {y}), f, x, A) \<in> Wellfounded.measure (\<lambda>(_, _, A). card A)"
     by simp
 qed simp_all
 
@@ -128,9 +128,9 @@
 termination proof (relation "Wellfounded.measure (\<lambda>(_,_,_,A). card A)")
   fix A :: "'a set" and f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" and x :: 'b 
     and y :: 'a and g :: "'b \<Rightarrow> 'c pmf"
-  assume "finite A" "A \<noteq> {}" "y \<in> set_pmf (pmf_of_set A)"
-  moreover from this have "card A > 0" by (simp add: card_gt_0_iff)
-  ultimately show "((f, g, f y x, A - {y}), f, g, x, A) \<in> Wellfounded.measure (\<lambda>(_, _, _, A). card A)"
+  assume A: "finite A" "A \<noteq> {}" "y \<in> set_pmf (pmf_of_set A)"
+  then have "card A > 0" by (simp add: card_gt_0_iff)
+  with A show "((f, g, f y x, A - {y}), f, g, x, A) \<in> Wellfounded.measure (\<lambda>(_, _, _, A). card A)"
     by simp
 qed simp_all
 
--- a/src/HOL/Probability/SPMF.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Probability/SPMF.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -2520,8 +2520,8 @@
     ultimately have "(weight_spmf p * weight_spmf q) * (weight_spmf p * weight_spmf q) = 1" by simp
     hence *: "weight_spmf p * weight_spmf q = 1"
       by(metis antisym_conv less_le mult_less_cancel_left1 weight_pair_spmf weight_spmf_le_1 weight_spmf_nonneg)
-    hence "weight_spmf p = 1" by(metis antisym_conv mult_left_le weight_spmf_le_1 weight_spmf_nonneg)
-    moreover with * have "weight_spmf q = 1" by simp
+    hence **: "weight_spmf p = 1" by(metis antisym_conv mult_left_le weight_spmf_le_1 weight_spmf_nonneg)
+    moreover from * ** have "weight_spmf q = 1" by simp
     moreover note calculation }
   note full = this
 
--- a/src/HOL/Quotient_Examples/Int_Pow.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Quotient_Examples/Int_Pow.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -41,8 +41,8 @@
   "[| x \<in> Units G; y \<in> Units G |] ==> inv (x \<otimes> y) = inv y \<otimes> inv x"
 proof -
   assume G: "x \<in> Units G"  "y \<in> Units G"
-  moreover then have "x \<in> carrier G"  "y \<in> carrier G" by auto
-  ultimately have "inv (x \<otimes> y) \<otimes> (x \<otimes> y) = (inv y \<otimes> inv x) \<otimes> (x \<otimes> y)"
+  then have "x \<in> carrier G"  "y \<in> carrier G" by auto
+  with G have "inv (x \<otimes> y) \<otimes> (x \<otimes> y) = (inv y \<otimes> inv x) \<otimes> (x \<otimes> y)"
     by (simp add: m_assoc) (simp add: m_assoc [symmetric])
   with G show ?thesis by (simp del: Units_l_inv)
 qed
--- a/src/HOL/Real_Vector_Spaces.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Real_Vector_Spaces.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -34,39 +34,36 @@
   using add [of x "- y"] by (simp add: minus)
 
 lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
-apply (cases "finite A")
-apply (induct set: finite)
-apply (simp add: zero)
-apply (simp add: add)
-apply (simp add: zero)
-done
+  apply (cases "finite A")
+   apply (induct set: finite)
+    apply (simp add: zero)
+   apply (simp add: add)
+  apply (simp add: zero)
+  done
 
 end
 
+
 subsection \<open>Vector spaces\<close>
 
 locale vector_space =
   fixes scale :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b"
-  assumes scale_right_distrib [algebra_simps]:
-    "scale a (x + y) = scale a x + scale a y"
-  and scale_left_distrib [algebra_simps]:
-    "scale (a + b) x = scale a x + scale b x"
-  and scale_scale [simp]: "scale a (scale b x) = scale (a * b) x"
-  and scale_one [simp]: "scale 1 x = x"
+  assumes scale_right_distrib [algebra_simps]: "scale a (x + y) = scale a x + scale a y"
+    and scale_left_distrib [algebra_simps]: "scale (a + b) x = scale a x + scale b x"
+    and scale_scale [simp]: "scale a (scale b x) = scale (a * b) x"
+    and scale_one [simp]: "scale 1 x = x"
 begin
 
-lemma scale_left_commute:
-  "scale a (scale b x) = scale b (scale a x)"
-by (simp add: mult.commute)
+lemma scale_left_commute: "scale a (scale b x) = scale b (scale a x)"
+  by (simp add: mult.commute)
 
 lemma scale_zero_left [simp]: "scale 0 x = 0"
   and scale_minus_left [simp]: "scale (- a) x = - (scale a x)"
-  and scale_left_diff_distrib [algebra_simps]:
-        "scale (a - b) x = scale a x - scale b x"
+  and scale_left_diff_distrib [algebra_simps]: "scale (a - b) x = scale a x - scale b x"
   and scale_setsum_left: "scale (setsum f A) x = (\<Sum>a\<in>A. scale (f a) x)"
 proof -
   interpret s: additive "\<lambda>a. scale a x"
-    proof qed (rule scale_left_distrib)
+    by standard (rule scale_left_distrib)
   show "scale 0 x = 0" by (rule s.zero)
   show "scale (- a) x = - (scale a x)" by (rule s.minus)
   show "scale (a - b) x = scale a x - scale b x" by (rule s.diff)
@@ -75,72 +72,70 @@
 
 lemma scale_zero_right [simp]: "scale a 0 = 0"
   and scale_minus_right [simp]: "scale a (- x) = - (scale a x)"
-  and scale_right_diff_distrib [algebra_simps]:
-        "scale a (x - y) = scale a x - scale a y"
+  and scale_right_diff_distrib [algebra_simps]: "scale a (x - y) = scale a x - scale a y"
   and scale_setsum_right: "scale a (setsum f A) = (\<Sum>x\<in>A. scale a (f x))"
 proof -
   interpret s: additive "\<lambda>x. scale a x"
-    proof qed (rule scale_right_distrib)
+    by standard (rule scale_right_distrib)
   show "scale a 0 = 0" by (rule s.zero)
   show "scale a (- x) = - (scale a x)" by (rule s.minus)
   show "scale a (x - y) = scale a x - scale a y" by (rule s.diff)
   show "scale a (setsum f A) = (\<Sum>x\<in>A. scale a (f x))" by (rule s.setsum)
 qed
 
-lemma scale_eq_0_iff [simp]:
-  "scale a x = 0 \<longleftrightarrow> a = 0 \<or> x = 0"
-proof cases
-  assume "a = 0" thus ?thesis by simp
+lemma scale_eq_0_iff [simp]: "scale a x = 0 \<longleftrightarrow> a = 0 \<or> x = 0"
+proof (cases "a = 0")
+  case True
+  then show ?thesis by simp
 next
-  assume anz [simp]: "a \<noteq> 0"
-  { assume "scale a x = 0"
-    hence "scale (inverse a) (scale a x) = 0" by simp
-    hence "x = 0" by simp }
-  thus ?thesis by force
+  case False
+  have "x = 0" if "scale a x = 0"
+  proof -
+    from False that have "scale (inverse a) (scale a x) = 0" by simp
+    with False show ?thesis by simp
+  qed
+  then show ?thesis by force
 qed
 
 lemma scale_left_imp_eq:
-  "\<lbrakk>a \<noteq> 0; scale a x = scale a y\<rbrakk> \<Longrightarrow> x = y"
+  assumes nonzero: "a \<noteq> 0"
+    and scale: "scale a x = scale a y"
+  shows "x = y"
 proof -
-  assume nonzero: "a \<noteq> 0"
-  assume "scale a x = scale a y"
-  hence "scale a (x - y) = 0"
+  from scale have "scale a (x - y) = 0"
      by (simp add: scale_right_diff_distrib)
-  hence "x - y = 0" by (simp add: nonzero)
-  thus "x = y" by (simp only: right_minus_eq)
+  with nonzero have "x - y = 0" by simp
+  then show "x = y" by (simp only: right_minus_eq)
 qed
 
 lemma scale_right_imp_eq:
-  "\<lbrakk>x \<noteq> 0; scale a x = scale b x\<rbrakk> \<Longrightarrow> a = b"
+  assumes nonzero: "x \<noteq> 0"
+    and scale: "scale a x = scale b x"
+  shows "a = b"
 proof -
-  assume nonzero: "x \<noteq> 0"
-  assume "scale a x = scale b x"
-  hence "scale (a - b) x = 0"
+  from scale have "scale (a - b) x = 0"
      by (simp add: scale_left_diff_distrib)
-  hence "a - b = 0" by (simp add: nonzero)
-  thus "a = b" by (simp only: right_minus_eq)
+  with nonzero have "a - b = 0" by simp
+  then show "a = b" by (simp only: right_minus_eq)
 qed
 
-lemma scale_cancel_left [simp]:
-  "scale a x = scale a y \<longleftrightarrow> x = y \<or> a = 0"
-by (auto intro: scale_left_imp_eq)
+lemma scale_cancel_left [simp]: "scale a x = scale a y \<longleftrightarrow> x = y \<or> a = 0"
+  by (auto intro: scale_left_imp_eq)
 
-lemma scale_cancel_right [simp]:
-  "scale a x = scale b x \<longleftrightarrow> a = b \<or> x = 0"
-by (auto intro: scale_right_imp_eq)
+lemma scale_cancel_right [simp]: "scale a x = scale b x \<longleftrightarrow> a = b \<or> x = 0"
+  by (auto intro: scale_right_imp_eq)
 
 end
 
+
 subsection \<open>Real vector spaces\<close>
 
 class scaleR =
   fixes scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "*\<^sub>R" 75)
 begin
 
-abbreviation
-  divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a" (infixl "'/\<^sub>R" 70)
-where
-  "x /\<^sub>R r == scaleR (inverse r) x"
+abbreviation divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a"  (infixl "'/\<^sub>R" 70)
+  where "x /\<^sub>R r \<equiv> scaleR (inverse r) x"
 
 end
 
@@ -150,14 +145,13 @@
   and scaleR_scaleR: "scaleR a (scaleR b x) = scaleR (a * b) x"
   and scaleR_one: "scaleR 1 x = x"
 
-interpretation real_vector:
-  vector_space "scaleR :: real \<Rightarrow> 'a \<Rightarrow> 'a::real_vector"
-apply unfold_locales
-apply (rule scaleR_add_right)
-apply (rule scaleR_add_left)
-apply (rule scaleR_scaleR)
-apply (rule scaleR_one)
-done
+interpretation real_vector: vector_space "scaleR :: real \<Rightarrow> 'a \<Rightarrow> 'a::real_vector"
+  apply unfold_locales
+     apply (rule scaleR_add_right)
+    apply (rule scaleR_add_left)
+   apply (rule scaleR_scaleR)
+  apply (rule scaleR_one)
+  done
 
 text \<open>Recover original theorem names\<close>
 
@@ -183,14 +177,13 @@
 lemmas scaleR_left_diff_distrib = scaleR_diff_left
 lemmas scaleR_right_diff_distrib = scaleR_diff_right
 
-lemma scaleR_minus1_left [simp]:
-  fixes x :: "'a::real_vector"
-  shows "scaleR (-1) x = - x"
+lemma scaleR_minus1_left [simp]: "scaleR (-1) x = - x"
+  for x :: "'a::real_vector"
   using scaleR_minus_left [of 1 x] by simp
 
 class real_algebra = real_vector + ring +
   assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)"
-  and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"
+    and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"
 
 class real_algebra_1 = real_algebra + ring_1
 
@@ -201,125 +194,122 @@
 instantiation real :: real_field
 begin
 
-definition
-  real_scaleR_def [simp]: "scaleR a x = a * x"
+definition real_scaleR_def [simp]: "scaleR a x = a * x"
 
-instance proof
-qed (simp_all add: algebra_simps)
+instance
+  by standard (simp_all add: algebra_simps)
 
 end
 
-interpretation scaleR_left: additive "(\<lambda>a. scaleR a x::'a::real_vector)"
-proof qed (rule scaleR_left_distrib)
+interpretation scaleR_left: additive "(\<lambda>a. scaleR a x :: 'a::real_vector)"
+  by standard (rule scaleR_left_distrib)
 
-interpretation scaleR_right: additive "(\<lambda>x. scaleR a x::'a::real_vector)"
-proof qed (rule scaleR_right_distrib)
+interpretation scaleR_right: additive "(\<lambda>x. scaleR a x :: 'a::real_vector)"
+  by standard (rule scaleR_right_distrib)
 
 lemma nonzero_inverse_scaleR_distrib:
-  fixes x :: "'a::real_div_algebra" shows
-  "\<lbrakk>a \<noteq> 0; x \<noteq> 0\<rbrakk> \<Longrightarrow> inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
-by (rule inverse_unique, simp)
+  "a \<noteq> 0 \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
+  for x :: "'a::real_div_algebra"
+  by (rule inverse_unique) simp
 
-lemma inverse_scaleR_distrib:
-  fixes x :: "'a::{real_div_algebra, division_ring}"
-  shows "inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
-apply (case_tac "a = 0", simp)
-apply (case_tac "x = 0", simp)
-apply (erule (1) nonzero_inverse_scaleR_distrib)
-done
-
-lemma setsum_constant_scaleR:
-  fixes y :: "'a::real_vector"
-  shows "(\<Sum>x\<in>A. y) = of_nat (card A) *\<^sub>R y"
-  apply (cases "finite A")
-  apply (induct set: finite)
-  apply (simp_all add: algebra_simps)
+lemma inverse_scaleR_distrib: "inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
+  for x :: "'a::{real_div_algebra,division_ring}"
+  apply (cases "a = 0")
+   apply simp
+  apply (cases "x = 0")
+   apply simp
+  apply (erule (1) nonzero_inverse_scaleR_distrib)
   done
 
-lemma vector_add_divide_simps :
-  fixes v :: "'a :: real_vector"
-  shows "v + (b / z) *\<^sub>R w = (if z = 0 then v else (z *\<^sub>R v + b *\<^sub>R w) /\<^sub>R z)"
-        "a *\<^sub>R v + (b / z) *\<^sub>R w = (if z = 0 then a *\<^sub>R v else ((a * z) *\<^sub>R v + b *\<^sub>R w) /\<^sub>R z)"
-        "(a / z) *\<^sub>R v + w = (if z = 0 then w else (a *\<^sub>R v + z *\<^sub>R w) /\<^sub>R z)"
-        "(a / z) *\<^sub>R v + b *\<^sub>R w = (if z = 0 then b *\<^sub>R w else (a *\<^sub>R v + (b * z) *\<^sub>R w) /\<^sub>R z)"
-        "v - (b / z) *\<^sub>R w = (if z = 0 then v else (z *\<^sub>R v - b *\<^sub>R w) /\<^sub>R z)"
-        "a *\<^sub>R v - (b / z) *\<^sub>R w = (if z = 0 then a *\<^sub>R v else ((a * z) *\<^sub>R v - b *\<^sub>R w) /\<^sub>R z)"
-        "(a / z) *\<^sub>R v - w = (if z = 0 then -w else (a *\<^sub>R v - z *\<^sub>R w) /\<^sub>R z)"
-        "(a / z) *\<^sub>R v - b *\<^sub>R w = (if z = 0 then -b *\<^sub>R w else (a *\<^sub>R v - (b * z) *\<^sub>R w) /\<^sub>R z)"
-by (simp_all add: divide_inverse_commute scaleR_add_right real_vector.scale_right_diff_distrib)
+lemma setsum_constant_scaleR: "(\<Sum>x\<in>A. y) = of_nat (card A) *\<^sub>R y"
+  for y :: "'a::real_vector"
+  apply (cases "finite A")
+   apply (induct set: finite)
+    apply (simp_all add: algebra_simps)
+  done
+
+lemma vector_add_divide_simps:
+  "v + (b / z) *\<^sub>R w = (if z = 0 then v else (z *\<^sub>R v + b *\<^sub>R w) /\<^sub>R z)"
+  "a *\<^sub>R v + (b / z) *\<^sub>R w = (if z = 0 then a *\<^sub>R v else ((a * z) *\<^sub>R v + b *\<^sub>R w) /\<^sub>R z)"
+  "(a / z) *\<^sub>R v + w = (if z = 0 then w else (a *\<^sub>R v + z *\<^sub>R w) /\<^sub>R z)"
+  "(a / z) *\<^sub>R v + b *\<^sub>R w = (if z = 0 then b *\<^sub>R w else (a *\<^sub>R v + (b * z) *\<^sub>R w) /\<^sub>R z)"
+  "v - (b / z) *\<^sub>R w = (if z = 0 then v else (z *\<^sub>R v - b *\<^sub>R w) /\<^sub>R z)"
+  "a *\<^sub>R v - (b / z) *\<^sub>R w = (if z = 0 then a *\<^sub>R v else ((a * z) *\<^sub>R v - b *\<^sub>R w) /\<^sub>R z)"
+  "(a / z) *\<^sub>R v - w = (if z = 0 then -w else (a *\<^sub>R v - z *\<^sub>R w) /\<^sub>R z)"
+  "(a / z) *\<^sub>R v - b *\<^sub>R w = (if z = 0 then -b *\<^sub>R w else (a *\<^sub>R v - (b * z) *\<^sub>R w) /\<^sub>R z)"
+  for v :: "'a :: real_vector"
+  by (simp_all add: divide_inverse_commute scaleR_add_right real_vector.scale_right_diff_distrib)
 
 lemma real_vector_affinity_eq:
   fixes x :: "'a :: real_vector"
   assumes m0: "m \<noteq> 0"
   shows "m *\<^sub>R x + c = y \<longleftrightarrow> x = inverse m *\<^sub>R y - (inverse m *\<^sub>R c)"
+    (is "?lhs \<longleftrightarrow> ?rhs")
 proof
-  assume h: "m *\<^sub>R x + c = y"
-  hence "m *\<^sub>R x = y - c" by (simp add: field_simps)
-  hence "inverse m *\<^sub>R (m *\<^sub>R x) = inverse m *\<^sub>R (y - c)" by simp
+  assume ?lhs
+  then have "m *\<^sub>R x = y - c" by (simp add: field_simps)
+  then have "inverse m *\<^sub>R (m *\<^sub>R x) = inverse m *\<^sub>R (y - c)" by simp
   then show "x = inverse m *\<^sub>R y - (inverse m *\<^sub>R c)"
     using m0
   by (simp add: real_vector.scale_right_diff_distrib)
 next
-  assume h: "x = inverse m *\<^sub>R y - (inverse m *\<^sub>R c)"
-  show "m *\<^sub>R x + c = y" unfolding h
-    using m0  by (simp add: real_vector.scale_right_diff_distrib)
+  assume ?rhs
+  with m0 show "m *\<^sub>R x + c = y"
+    by (simp add: real_vector.scale_right_diff_distrib)
 qed
 
-lemma real_vector_eq_affinity:
-  fixes x :: "'a :: real_vector"
-  shows "m \<noteq> 0 ==> (y = m *\<^sub>R x + c \<longleftrightarrow> inverse m *\<^sub>R y - (inverse m *\<^sub>R c) = x)"
+lemma real_vector_eq_affinity: "m \<noteq> 0 \<Longrightarrow> y = m *\<^sub>R x + c \<longleftrightarrow> inverse m *\<^sub>R y - (inverse m *\<^sub>R c) = x"
+  for x :: "'a::real_vector"
   using real_vector_affinity_eq[where m=m and x=x and y=y and c=c]
   by metis
 
-lemma scaleR_eq_iff [simp]:
-  fixes a :: "'a :: real_vector"
-  shows "b + u *\<^sub>R a = a + u *\<^sub>R b \<longleftrightarrow> a=b \<or> u=1"
-proof (cases "u=1")
-  case True then show ?thesis by auto
+lemma scaleR_eq_iff [simp]: "b + u *\<^sub>R a = a + u *\<^sub>R b \<longleftrightarrow> a = b \<or> u = 1"
+  for a :: "'a::real_vector"
+proof (cases "u = 1")
+  case True
+  then show ?thesis by auto
 next
   case False
-  { assume "b + u *\<^sub>R a = a + u *\<^sub>R b"
-    then have "(u - 1) *\<^sub>R a = (u - 1) *\<^sub>R b"
+  have "a = b" if "b + u *\<^sub>R a = a + u *\<^sub>R b"
+  proof -
+    from that have "(u - 1) *\<^sub>R a = (u - 1) *\<^sub>R b"
       by (simp add: algebra_simps)
-    with False have "a=b"
+    with False show ?thesis
       by auto
-  }
+  qed
   then show ?thesis by auto
 qed
 
-lemma scaleR_collapse [simp]:
-  fixes a :: "'a :: real_vector"
-  shows "(1 - u) *\<^sub>R a + u *\<^sub>R a = a"
-by (simp add: algebra_simps)
+lemma scaleR_collapse [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R a = a"
+  for a :: "'a::real_vector"
+  by (simp add: algebra_simps)
 
 
-subsection \<open>Embedding of the Reals into any \<open>real_algebra_1\<close>:
-@{term of_real}\<close>
+subsection \<open>Embedding of the Reals into any \<open>real_algebra_1\<close>: \<open>of_real\<close>\<close>
 
-definition
-  of_real :: "real \<Rightarrow> 'a::real_algebra_1" where
-  "of_real r = scaleR r 1"
+definition of_real :: "real \<Rightarrow> 'a::real_algebra_1"
+  where "of_real r = scaleR r 1"
 
 lemma scaleR_conv_of_real: "scaleR r x = of_real r * x"
-by (simp add: of_real_def)
+  by (simp add: of_real_def)
 
 lemma of_real_0 [simp]: "of_real 0 = 0"
-by (simp add: of_real_def)
+  by (simp add: of_real_def)
 
 lemma of_real_1 [simp]: "of_real 1 = 1"
-by (simp add: of_real_def)
+  by (simp add: of_real_def)
 
 lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y"
-by (simp add: of_real_def scaleR_left_distrib)
+  by (simp add: of_real_def scaleR_left_distrib)
 
 lemma of_real_minus [simp]: "of_real (- x) = - of_real x"
-by (simp add: of_real_def)
+  by (simp add: of_real_def)
 
 lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y"
-by (simp add: of_real_def scaleR_left_diff_distrib)
+  by (simp add: of_real_def scaleR_left_diff_distrib)
 
 lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y"
-by (simp add: of_real_def mult.commute)
+  by (simp add: of_real_def mult.commute)
 
 lemma of_real_setsum[simp]: "of_real (setsum f s) = (\<Sum>x\<in>s. of_real (f x))"
   by (induct s rule: infinite_finite_induct) auto
@@ -328,63 +318,56 @@
   by (induct s rule: infinite_finite_induct) auto
 
 lemma nonzero_of_real_inverse:
-  "x \<noteq> 0 \<Longrightarrow> of_real (inverse x) =
-   inverse (of_real x :: 'a::real_div_algebra)"
-by (simp add: of_real_def nonzero_inverse_scaleR_distrib)
+  "x \<noteq> 0 \<Longrightarrow> of_real (inverse x) = inverse (of_real x :: 'a::real_div_algebra)"
+  by (simp add: of_real_def nonzero_inverse_scaleR_distrib)
 
 lemma of_real_inverse [simp]:
-  "of_real (inverse x) =
-   inverse (of_real x :: 'a::{real_div_algebra, division_ring})"
-by (simp add: of_real_def inverse_scaleR_distrib)
+  "of_real (inverse x) = inverse (of_real x :: 'a::{real_div_algebra,division_ring})"
+  by (simp add: of_real_def inverse_scaleR_distrib)
 
 lemma nonzero_of_real_divide:
-  "y \<noteq> 0 \<Longrightarrow> of_real (x / y) =
-   (of_real x / of_real y :: 'a::real_field)"
-by (simp add: divide_inverse nonzero_of_real_inverse)
+  "y \<noteq> 0 \<Longrightarrow> of_real (x / y) = (of_real x / of_real y :: 'a::real_field)"
+  by (simp add: divide_inverse nonzero_of_real_inverse)
 
 lemma of_real_divide [simp]:
   "of_real (x / y) = (of_real x / of_real y :: 'a::real_div_algebra)"
-by (simp add: divide_inverse)
+  by (simp add: divide_inverse)
 
 lemma of_real_power [simp]:
   "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1}) ^ n"
-by (induct n) simp_all
+  by (induct n) simp_all
 
-lemma of_real_eq_iff [simp]: "(of_real x = of_real y) = (x = y)"
-by (simp add: of_real_def)
+lemma of_real_eq_iff [simp]: "of_real x = of_real y \<longleftrightarrow> x = y"
+  by (simp add: of_real_def)
 
-lemma inj_of_real:
-  "inj of_real"
+lemma inj_of_real: "inj of_real"
   by (auto intro: injI)
 
 lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified]
 
 lemma of_real_eq_id [simp]: "of_real = (id :: real \<Rightarrow> real)"
-proof
-  fix r
-  show "of_real r = id r"
-    by (simp add: of_real_def)
-qed
+  by (rule ext) (simp add: of_real_def)
 
-text\<open>Collapse nested embeddings\<close>
+text \<open>Collapse nested embeddings.\<close>
 lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n"
-by (induct n) auto
+  by (induct n) auto
 
 lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z"
-by (cases z rule: int_diff_cases, simp)
+  by (cases z rule: int_diff_cases) simp
 
 lemma of_real_numeral [simp]: "of_real (numeral w) = numeral w"
-using of_real_of_int_eq [of "numeral w"] by simp
+  using of_real_of_int_eq [of "numeral w"] by simp
 
 lemma of_real_neg_numeral [simp]: "of_real (- numeral w) = - numeral w"
-using of_real_of_int_eq [of "- numeral w"] by simp
+  using of_real_of_int_eq [of "- numeral w"] by simp
 
-text\<open>Every real algebra has characteristic zero\<close>
-
+text \<open>Every real algebra has characteristic zero.\<close>
 instance real_algebra_1 < ring_char_0
 proof
-  from inj_of_real inj_of_nat have "inj (of_real \<circ> of_nat)" by (rule inj_comp)
-  then show "inj (of_nat :: nat \<Rightarrow> 'a)" by (simp add: comp_def)
+  from inj_of_real inj_of_nat have "inj (of_real \<circ> of_nat)"
+    by (rule inj_comp)
+  then show "inj (of_nat :: nat \<Rightarrow> 'a)"
+    by (simp add: comp_def)
 qed
 
 instance real_field < field_char_0 ..
@@ -396,97 +379,91 @@
   where "\<real> = range of_real"
 
 lemma Reals_of_real [simp]: "of_real r \<in> \<real>"
-by (simp add: Reals_def)
+  by (simp add: Reals_def)
 
 lemma Reals_of_int [simp]: "of_int z \<in> \<real>"
-by (subst of_real_of_int_eq [symmetric], rule Reals_of_real)
+  by (subst of_real_of_int_eq [symmetric], rule Reals_of_real)
 
 lemma Reals_of_nat [simp]: "of_nat n \<in> \<real>"
-by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real)
+  by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real)
 
 lemma Reals_numeral [simp]: "numeral w \<in> \<real>"
-by (subst of_real_numeral [symmetric], rule Reals_of_real)
+  by (subst of_real_numeral [symmetric], rule Reals_of_real)
 
 lemma Reals_0 [simp]: "0 \<in> \<real>"
-apply (unfold Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_0 [symmetric])
-done
+  apply (unfold Reals_def)
+  apply (rule range_eqI)
+  apply (rule of_real_0 [symmetric])
+  done
 
 lemma Reals_1 [simp]: "1 \<in> \<real>"
-apply (unfold Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_1 [symmetric])
-done
+  apply (unfold Reals_def)
+  apply (rule range_eqI)
+  apply (rule of_real_1 [symmetric])
+  done
 
-lemma Reals_add [simp]: "\<lbrakk>a \<in> \<real>; b \<in> \<real>\<rbrakk> \<Longrightarrow> a + b \<in> \<real>"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_add [symmetric])
-done
+lemma Reals_add [simp]: "a \<in> \<real> \<Longrightarrow> b \<in> \<real> \<Longrightarrow> a + b \<in> \<real>"
+  apply (auto simp add: Reals_def)
+  apply (rule range_eqI)
+  apply (rule of_real_add [symmetric])
+  done
 
 lemma Reals_minus [simp]: "a \<in> \<real> \<Longrightarrow> - a \<in> \<real>"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_minus [symmetric])
-done
+  apply (auto simp add: Reals_def)
+  apply (rule range_eqI)
+  apply (rule of_real_minus [symmetric])
+  done
 
-lemma Reals_diff [simp]: "\<lbrakk>a \<in> \<real>; b \<in> \<real>\<rbrakk> \<Longrightarrow> a - b \<in> \<real>"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_diff [symmetric])
-done
+lemma Reals_diff [simp]: "a \<in> \<real> \<Longrightarrow> b \<in> \<real> \<Longrightarrow> a - b \<in> \<real>"
+  apply (auto simp add: Reals_def)
+  apply (rule range_eqI)
+  apply (rule of_real_diff [symmetric])
+  done
 
-lemma Reals_mult [simp]: "\<lbrakk>a \<in> \<real>; b \<in> \<real>\<rbrakk> \<Longrightarrow> a * b \<in> \<real>"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_mult [symmetric])
-done
+lemma Reals_mult [simp]: "a \<in> \<real> \<Longrightarrow> b \<in> \<real> \<Longrightarrow> a * b \<in> \<real>"
+  apply (auto simp add: Reals_def)
+  apply (rule range_eqI)
+  apply (rule of_real_mult [symmetric])
+  done
 
-lemma nonzero_Reals_inverse:
-  fixes a :: "'a::real_div_algebra"
-  shows "\<lbrakk>a \<in> \<real>; a \<noteq> 0\<rbrakk> \<Longrightarrow> inverse a \<in> \<real>"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (erule nonzero_of_real_inverse [symmetric])
-done
+lemma nonzero_Reals_inverse: "a \<in> \<real> \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> inverse a \<in> \<real>"
+  for a :: "'a::real_div_algebra"
+  apply (auto simp add: Reals_def)
+  apply (rule range_eqI)
+  apply (erule nonzero_of_real_inverse [symmetric])
+  done
 
-lemma Reals_inverse:
-  fixes a :: "'a::{real_div_algebra, division_ring}"
-  shows "a \<in> \<real> \<Longrightarrow> inverse a \<in> \<real>"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_inverse [symmetric])
-done
+lemma Reals_inverse: "a \<in> \<real> \<Longrightarrow> inverse a \<in> \<real>"
+  for a :: "'a::{real_div_algebra,division_ring}"
+  apply (auto simp add: Reals_def)
+  apply (rule range_eqI)
+  apply (rule of_real_inverse [symmetric])
+  done
 
-lemma Reals_inverse_iff [simp]:
-  fixes x:: "'a :: {real_div_algebra, division_ring}"
-  shows "inverse x \<in> \<real> \<longleftrightarrow> x \<in> \<real>"
-by (metis Reals_inverse inverse_inverse_eq)
+lemma Reals_inverse_iff [simp]: "inverse x \<in> \<real> \<longleftrightarrow> x \<in> \<real>"
+  for x :: "'a::{real_div_algebra,division_ring}"
+  by (metis Reals_inverse inverse_inverse_eq)
 
-lemma nonzero_Reals_divide:
-  fixes a b :: "'a::real_field"
-  shows "\<lbrakk>a \<in> \<real>; b \<in> \<real>; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / b \<in> \<real>"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (erule nonzero_of_real_divide [symmetric])
-done
+lemma nonzero_Reals_divide: "a \<in> \<real> \<Longrightarrow> b \<in> \<real> \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a / b \<in> \<real>"
+  for a b :: "'a::real_field"
+  apply (auto simp add: Reals_def)
+  apply (rule range_eqI)
+  apply (erule nonzero_of_real_divide [symmetric])
+  done
 
-lemma Reals_divide [simp]:
-  fixes a b :: "'a::{real_field, field}"
-  shows "\<lbrakk>a \<in> \<real>; b \<in> \<real>\<rbrakk> \<Longrightarrow> a / b \<in> \<real>"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_divide [symmetric])
-done
+lemma Reals_divide [simp]: "a \<in> \<real> \<Longrightarrow> b \<in> \<real> \<Longrightarrow> a / b \<in> \<real>"
+  for a b :: "'a::{real_field,field}"
+  apply (auto simp add: Reals_def)
+  apply (rule range_eqI)
+  apply (rule of_real_divide [symmetric])
+  done
 
-lemma Reals_power [simp]:
-  fixes a :: "'a::{real_algebra_1}"
-  shows "a \<in> \<real> \<Longrightarrow> a ^ n \<in> \<real>"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_power [symmetric])
-done
+lemma Reals_power [simp]: "a \<in> \<real> \<Longrightarrow> a ^ n \<in> \<real>"
+  for a :: "'a::real_algebra_1"
+  apply (auto simp add: Reals_def)
+  apply (rule range_eqI)
+  apply (rule of_real_power [symmetric])
+  done
 
 lemma Reals_cases [cases set: Reals]:
   assumes "q \<in> \<real>"
@@ -499,49 +476,56 @@
 qed
 
 lemma setsum_in_Reals [intro,simp]:
-  assumes "\<And>i. i \<in> s \<Longrightarrow> f i \<in> \<real>" shows "setsum f s \<in> \<real>"
+  assumes "\<And>i. i \<in> s \<Longrightarrow> f i \<in> \<real>"
+  shows "setsum f s \<in> \<real>"
 proof (cases "finite s")
-  case True then show ?thesis using assms
-    by (induct s rule: finite_induct) auto
+  case True
+  then show ?thesis
+    using assms by (induct s rule: finite_induct) auto
 next
-  case False then show ?thesis using assms
-    by (metis Reals_0 setsum.infinite)
+  case False
+  then show ?thesis
+    using assms by (metis Reals_0 setsum.infinite)
 qed
 
 lemma setprod_in_Reals [intro,simp]:
-  assumes "\<And>i. i \<in> s \<Longrightarrow> f i \<in> \<real>" shows "setprod f s \<in> \<real>"
+  assumes "\<And>i. i \<in> s \<Longrightarrow> f i \<in> \<real>"
+  shows "setprod f s \<in> \<real>"
 proof (cases "finite s")
-  case True then show ?thesis using assms
-    by (induct s rule: finite_induct) auto
+  case True
+  then show ?thesis
+    using assms by (induct s rule: finite_induct) auto
 next
-  case False then show ?thesis using assms
-    by (metis Reals_1 setprod.infinite)
+  case False
+  then show ?thesis
+    using assms by (metis Reals_1 setprod.infinite)
 qed
 
 lemma Reals_induct [case_names of_real, induct set: Reals]:
   "q \<in> \<real> \<Longrightarrow> (\<And>r. P (of_real r)) \<Longrightarrow> P q"
   by (rule Reals_cases) auto
 
+
 subsection \<open>Ordered real vector spaces\<close>
 
 class ordered_real_vector = real_vector + ordered_ab_group_add +
   assumes scaleR_left_mono: "x \<le> y \<Longrightarrow> 0 \<le> a \<Longrightarrow> a *\<^sub>R x \<le> a *\<^sub>R y"
-  assumes scaleR_right_mono: "a \<le> b \<Longrightarrow> 0 \<le> x \<Longrightarrow> a *\<^sub>R x \<le> b *\<^sub>R x"
+    and scaleR_right_mono: "a \<le> b \<Longrightarrow> 0 \<le> x \<Longrightarrow> a *\<^sub>R x \<le> b *\<^sub>R x"
 begin
 
-lemma scaleR_mono:
-  "a \<le> b \<Longrightarrow> x \<le> y \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> x \<Longrightarrow> a *\<^sub>R x \<le> b *\<^sub>R y"
-apply (erule scaleR_right_mono [THEN order_trans], assumption)
-apply (erule scaleR_left_mono, assumption)
-done
+lemma scaleR_mono: "a \<le> b \<Longrightarrow> x \<le> y \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> x \<Longrightarrow> a *\<^sub>R x \<le> b *\<^sub>R y"
+  apply (erule scaleR_right_mono [THEN order_trans])
+   apply assumption
+  apply (erule scaleR_left_mono)
+  apply assumption
+  done
 
-lemma scaleR_mono':
-  "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> c \<Longrightarrow> a *\<^sub>R c \<le> b *\<^sub>R d"
+lemma scaleR_mono': "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> c \<Longrightarrow> a *\<^sub>R c \<le> b *\<^sub>R d"
   by (rule scaleR_mono) (auto intro: order.trans)
 
 lemma pos_le_divideRI:
   assumes "0 < c"
-  assumes "c *\<^sub>R a \<le> b"
+    and "c *\<^sub>R a \<le> b"
   shows "a \<le> b /\<^sub>R c"
 proof -
   from scaleR_left_mono[OF assms(2)] assms(1)
@@ -554,20 +538,22 @@
 lemma pos_le_divideR_eq:
   assumes "0 < c"
   shows "a \<le> b /\<^sub>R c \<longleftrightarrow> c *\<^sub>R a \<le> b"
-proof rule
-  assume "a \<le> b /\<^sub>R c"
-  from scaleR_left_mono[OF this] assms
-  have "c *\<^sub>R a \<le> c *\<^sub>R (b /\<^sub>R c)"
+    (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume ?lhs
+  from scaleR_left_mono[OF this] assms have "c *\<^sub>R a \<le> c *\<^sub>R (b /\<^sub>R c)"
     by simp
-  with assms show "c *\<^sub>R a \<le> b"
+  with assms show ?rhs
     by (simp add: scaleR_one scaleR_scaleR inverse_eq_divide)
-qed (rule pos_le_divideRI[OF assms])
+next
+  assume ?rhs
+  with assms show ?lhs by (rule pos_le_divideRI)
+qed
 
-lemma scaleR_image_atLeastAtMost:
-  "c > 0 \<Longrightarrow> scaleR c ` {x..y} = {c *\<^sub>R x..c *\<^sub>R y}"
+lemma scaleR_image_atLeastAtMost: "c > 0 \<Longrightarrow> scaleR c ` {x..y} = {c *\<^sub>R x..c *\<^sub>R y}"
   apply (auto intro!: scaleR_left_mono)
   apply (rule_tac x = "inverse c *\<^sub>R xa" in image_eqI)
-  apply (simp_all add: pos_le_divideR_eq[symmetric] scaleR_scaleR scaleR_one)
+   apply (simp_all add: pos_le_divideR_eq[symmetric] scaleR_scaleR scaleR_one)
   done
 
 end
@@ -576,103 +562,105 @@
   fixes a :: "'a :: ordered_real_vector"
   assumes "c < 0"
   shows "a \<le> b /\<^sub>R c \<longleftrightarrow> b \<le> c *\<^sub>R a"
-  using pos_le_divideR_eq [of "-c" a "-b"] assms
-  by simp
+  using pos_le_divideR_eq [of "-c" a "-b"] assms by simp
 
-lemma scaleR_nonneg_nonneg: "0 \<le> a \<Longrightarrow> 0 \<le> (x::'a::ordered_real_vector) \<Longrightarrow> 0 \<le> a *\<^sub>R x"
-  using scaleR_left_mono [of 0 x a]
-  by simp
+lemma scaleR_nonneg_nonneg: "0 \<le> a \<Longrightarrow> 0 \<le> x \<Longrightarrow> 0 \<le> a *\<^sub>R x"
+  for x :: "'a::ordered_real_vector"
+  using scaleR_left_mono [of 0 x a] by simp
 
-lemma scaleR_nonneg_nonpos: "0 \<le> a \<Longrightarrow> (x::'a::ordered_real_vector) \<le> 0 \<Longrightarrow> a *\<^sub>R x \<le> 0"
+lemma scaleR_nonneg_nonpos: "0 \<le> a \<Longrightarrow> x \<le> 0 \<Longrightarrow> a *\<^sub>R x \<le> 0"
+  for x :: "'a::ordered_real_vector"
   using scaleR_left_mono [of x 0 a] by simp
 
-lemma scaleR_nonpos_nonneg: "a \<le> 0 \<Longrightarrow> 0 \<le> (x::'a::ordered_real_vector) \<Longrightarrow> a *\<^sub>R x \<le> 0"
+lemma scaleR_nonpos_nonneg: "a \<le> 0 \<Longrightarrow> 0 \<le> x \<Longrightarrow> a *\<^sub>R x \<le> 0"
+  for x :: "'a::ordered_real_vector"
   using scaleR_right_mono [of a 0 x] by simp
 
-lemma split_scaleR_neg_le: "(0 \<le> a & x \<le> 0) | (a \<le> 0 & 0 \<le> x) \<Longrightarrow>
-  a *\<^sub>R (x::'a::ordered_real_vector) \<le> 0"
+lemma split_scaleR_neg_le: "(0 \<le> a \<and> x \<le> 0) \<or> (a \<le> 0 \<and> 0 \<le> x) \<Longrightarrow> a *\<^sub>R x \<le> 0"
+  for x :: "'a::ordered_real_vector"
   by (auto simp add: scaleR_nonneg_nonpos scaleR_nonpos_nonneg)
 
-lemma le_add_iff1:
-  fixes c d e::"'a::ordered_real_vector"
-  shows "a *\<^sub>R e + c \<le> b *\<^sub>R e + d \<longleftrightarrow> (a - b) *\<^sub>R e + c \<le> d"
+lemma le_add_iff1: "a *\<^sub>R e + c \<le> b *\<^sub>R e + d \<longleftrightarrow> (a - b) *\<^sub>R e + c \<le> d"
+  for c d e :: "'a::ordered_real_vector"
   by (simp add: algebra_simps)
 
-lemma le_add_iff2:
-  fixes c d e::"'a::ordered_real_vector"
-  shows "a *\<^sub>R e + c \<le> b *\<^sub>R e + d \<longleftrightarrow> c \<le> (b - a) *\<^sub>R e + d"
+lemma le_add_iff2: "a *\<^sub>R e + c \<le> b *\<^sub>R e + d \<longleftrightarrow> c \<le> (b - a) *\<^sub>R e + d"
+  for c d e :: "'a::ordered_real_vector"
   by (simp add: algebra_simps)
 
-lemma scaleR_left_mono_neg:
-  fixes a b::"'a::ordered_real_vector"
-  shows "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> c *\<^sub>R a \<le> c *\<^sub>R b"
+lemma scaleR_left_mono_neg: "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> c *\<^sub>R a \<le> c *\<^sub>R b"
+  for a b :: "'a::ordered_real_vector"
   apply (drule scaleR_left_mono [of _ _ "- c"])
-  apply simp_all
+   apply simp_all
   done
 
-lemma scaleR_right_mono_neg:
-  fixes c::"'a::ordered_real_vector"
-  shows "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> a *\<^sub>R c \<le> b *\<^sub>R c"
+lemma scaleR_right_mono_neg: "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> a *\<^sub>R c \<le> b *\<^sub>R c"
+  for c :: "'a::ordered_real_vector"
   apply (drule scaleR_right_mono [of _ _ "- c"])
-  apply simp_all
+   apply simp_all
   done
 
-lemma scaleR_nonpos_nonpos: "a \<le> 0 \<Longrightarrow> (b::'a::ordered_real_vector) \<le> 0 \<Longrightarrow> 0 \<le> a *\<^sub>R b"
-using scaleR_right_mono_neg [of a 0 b] by simp
+lemma scaleR_nonpos_nonpos: "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> 0 \<le> a *\<^sub>R b"
+  for b :: "'a::ordered_real_vector"
+  using scaleR_right_mono_neg [of a 0 b] by simp
 
-lemma split_scaleR_pos_le:
-  fixes b::"'a::ordered_real_vector"
-  shows "(0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0) \<Longrightarrow> 0 \<le> a *\<^sub>R b"
+lemma split_scaleR_pos_le: "(0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0) \<Longrightarrow> 0 \<le> a *\<^sub>R b"
+  for b :: "'a::ordered_real_vector"
   by (auto simp add: scaleR_nonneg_nonneg scaleR_nonpos_nonpos)
 
 lemma zero_le_scaleR_iff:
-  fixes b::"'a::ordered_real_vector"
-  shows "0 \<le> a *\<^sub>R b \<longleftrightarrow> 0 < a \<and> 0 \<le> b \<or> a < 0 \<and> b \<le> 0 \<or> a = 0" (is "?lhs = ?rhs")
-proof cases
-  assume "a \<noteq> 0"
+  fixes b :: "'a::ordered_real_vector"
+  shows "0 \<le> a *\<^sub>R b \<longleftrightarrow> 0 < a \<and> 0 \<le> b \<or> a < 0 \<and> b \<le> 0 \<or> a = 0"
+    (is "?lhs = ?rhs")
+proof (cases "a = 0")
+  case True
+  then show ?thesis by simp
+next
+  case False
   show ?thesis
   proof
-    assume lhs: ?lhs
-    {
-      assume "0 < a"
-      with lhs have "inverse a *\<^sub>R 0 \<le> inverse a *\<^sub>R (a *\<^sub>R b)"
+    assume ?lhs
+    from \<open>a \<noteq> 0\<close> consider "a > 0" | "a < 0" by arith
+    then show ?rhs
+    proof cases
+      case 1
+      with \<open>?lhs\<close> have "inverse a *\<^sub>R 0 \<le> inverse a *\<^sub>R (a *\<^sub>R b)"
         by (intro scaleR_mono) auto
-      hence ?rhs using \<open>0 < a\<close>
+      with 1 show ?thesis
         by simp
-    } moreover {
-      assume "0 > a"
-      with lhs have "- inverse a *\<^sub>R 0 \<le> - inverse a *\<^sub>R (a *\<^sub>R b)"
+    next
+      case 2
+      with \<open>?lhs\<close> have "- inverse a *\<^sub>R 0 \<le> - inverse a *\<^sub>R (a *\<^sub>R b)"
         by (intro scaleR_mono) auto
-      hence ?rhs using \<open>0 > a\<close>
+      with 2 show ?thesis
         by simp
-    } ultimately show ?rhs using \<open>a \<noteq> 0\<close> by arith
-  qed (auto simp: not_le \<open>a \<noteq> 0\<close> intro!: split_scaleR_pos_le)
-qed simp
+    qed
+  next
+    assume ?rhs
+    then show ?lhs
+      by (auto simp: not_le \<open>a \<noteq> 0\<close> intro!: split_scaleR_pos_le)
+  qed
+qed
 
-lemma scaleR_le_0_iff:
-  fixes b::"'a::ordered_real_vector"
-  shows "a *\<^sub>R b \<le> 0 \<longleftrightarrow> 0 < a \<and> b \<le> 0 \<or> a < 0 \<and> 0 \<le> b \<or> a = 0"
+lemma scaleR_le_0_iff: "a *\<^sub>R b \<le> 0 \<longleftrightarrow> 0 < a \<and> b \<le> 0 \<or> a < 0 \<and> 0 \<le> b \<or> a = 0"
+  for b::"'a::ordered_real_vector"
   by (insert zero_le_scaleR_iff [of "-a" b]) force
 
-lemma scaleR_le_cancel_left:
-  fixes b::"'a::ordered_real_vector"
-  shows "c *\<^sub>R a \<le> c *\<^sub>R b \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
+lemma scaleR_le_cancel_left: "c *\<^sub>R a \<le> c *\<^sub>R b \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
+  for b :: "'a::ordered_real_vector"
   by (auto simp add: neq_iff scaleR_left_mono scaleR_left_mono_neg
-    dest: scaleR_left_mono[where a="inverse c"] scaleR_left_mono_neg[where c="inverse c"])
+      dest: scaleR_left_mono[where a="inverse c"] scaleR_left_mono_neg[where c="inverse c"])
 
-lemma scaleR_le_cancel_left_pos:
-  fixes b::"'a::ordered_real_vector"
-  shows "0 < c \<Longrightarrow> c *\<^sub>R a \<le> c *\<^sub>R b \<longleftrightarrow> a \<le> b"
+lemma scaleR_le_cancel_left_pos: "0 < c \<Longrightarrow> c *\<^sub>R a \<le> c *\<^sub>R b \<longleftrightarrow> a \<le> b"
+  for b :: "'a::ordered_real_vector"
   by (auto simp: scaleR_le_cancel_left)
 
-lemma scaleR_le_cancel_left_neg:
-  fixes b::"'a::ordered_real_vector"
-  shows "c < 0 \<Longrightarrow> c *\<^sub>R a \<le> c *\<^sub>R b \<longleftrightarrow> b \<le> a"
+lemma scaleR_le_cancel_left_neg: "c < 0 \<Longrightarrow> c *\<^sub>R a \<le> c *\<^sub>R b \<longleftrightarrow> b \<le> a"
+  for b :: "'a::ordered_real_vector"
   by (auto simp: scaleR_le_cancel_left)
 
-lemma scaleR_left_le_one_le:
-  fixes x::"'a::ordered_real_vector" and a::real
-  shows "0 \<le> x \<Longrightarrow> a \<le> 1 \<Longrightarrow> a *\<^sub>R x \<le> x"
+lemma scaleR_left_le_one_le: "0 \<le> x \<Longrightarrow> a \<le> 1 \<Longrightarrow> a *\<^sub>R x \<le> x"
+  for x :: "'a::ordered_real_vector" and a :: real
   using scaleR_right_mono[of a 1 x] by simp
 
 
@@ -704,8 +692,8 @@
 
 class real_normed_vector = real_vector + sgn_div_norm + dist_norm + uniformity_dist + open_uniformity +
   assumes norm_eq_zero [simp]: "norm x = 0 \<longleftrightarrow> x = 0"
-  and norm_triangle_ineq: "norm (x + y) \<le> norm x + norm y"
-  and norm_scaleR [simp]: "norm (scaleR a x) = \<bar>a\<bar> * norm x"
+    and norm_triangle_ineq: "norm (x + y) \<le> norm x + norm y"
+    and norm_scaleR [simp]: "norm (scaleR a x) = \<bar>a\<bar> * norm x"
 begin
 
 lemma norm_ge_zero [simp]: "0 \<le> norm x"
@@ -724,9 +712,8 @@
 class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra +
   assumes norm_one [simp]: "norm 1 = 1"
 
-lemma (in real_normed_algebra_1) scaleR_power [simp]:
-  "(scaleR x y) ^ n = scaleR (x^n) (y^n)"
-  by (induction n) (simp_all add: scaleR_one scaleR_scaleR mult_ac)
+lemma (in real_normed_algebra_1) scaleR_power [simp]: "(scaleR x y) ^ n = scaleR (x^n) (y^n)"
+  by (induct n) (simp_all add: scaleR_one scaleR_scaleR mult_ac)
 
 class real_normed_div_algebra = real_div_algebra + real_normed_vector +
   assumes norm_mult: "norm (x * y) = norm x * norm y"
@@ -735,36 +722,31 @@
 
 instance real_normed_div_algebra < real_normed_algebra_1
 proof
-  fix x y :: 'a
-  show "norm (x * y) \<le> norm x * norm y"
+  show "norm (x * y) \<le> norm x * norm y" for x y :: 'a
     by (simp add: norm_mult)
 next
   have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)"
     by (rule norm_mult)
-  thus "norm (1::'a) = 1" by simp
+  then show "norm (1::'a) = 1" by simp
 qed
 
 lemma norm_zero [simp]: "norm (0::'a::real_normed_vector) = 0"
-by simp
+  by simp
 
-lemma zero_less_norm_iff [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "(0 < norm x) = (x \<noteq> 0)"
-by (simp add: order_less_le)
+lemma zero_less_norm_iff [simp]: "norm x > 0 \<longleftrightarrow> x \<noteq> 0"
+  for x :: "'a::real_normed_vector"
+  by (simp add: order_less_le)
 
-lemma norm_not_less_zero [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "\<not> norm x < 0"
-by (simp add: linorder_not_less)
+lemma norm_not_less_zero [simp]: "\<not> norm x < 0"
+  for x :: "'a::real_normed_vector"
+  by (simp add: linorder_not_less)
 
-lemma norm_le_zero_iff [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "(norm x \<le> 0) = (x = 0)"
-by (simp add: order_le_less)
+lemma norm_le_zero_iff [simp]: "norm x \<le> 0 \<longleftrightarrow> x = 0"
+  for x :: "'a::real_normed_vector"
+  by (simp add: order_le_less)
 
-lemma norm_minus_cancel [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "norm (- x) = norm x"
+lemma norm_minus_cancel [simp]: "norm (- x) = norm x"
+  for x :: "'a::real_normed_vector"
 proof -
   have "norm (- x) = norm (scaleR (- 1) x)"
     by (simp only: scaleR_minus_left scaleR_one)
@@ -773,78 +755,68 @@
   finally show ?thesis by simp
 qed
 
-lemma norm_minus_commute:
-  fixes a b :: "'a::real_normed_vector"
-  shows "norm (a - b) = norm (b - a)"
+lemma norm_minus_commute: "norm (a - b) = norm (b - a)"
+  for a b :: "'a::real_normed_vector"
 proof -
   have "norm (- (b - a)) = norm (b - a)"
     by (rule norm_minus_cancel)
-  thus ?thesis by simp
+  then show ?thesis by simp
 qed
-  
-lemma dist_add_cancel [simp]:
-  fixes a :: "'a::real_normed_vector"
-  shows "dist (a + b) (a + c) = dist b c"
-by (simp add: dist_norm)
+
+lemma dist_add_cancel [simp]: "dist (a + b) (a + c) = dist b c"
+  for a :: "'a::real_normed_vector"
+  by (simp add: dist_norm)
 
-lemma dist_add_cancel2 [simp]:
-  fixes a :: "'a::real_normed_vector"
-  shows "dist (b + a) (c + a) = dist b c"
-by (simp add: dist_norm)
+lemma dist_add_cancel2 [simp]: "dist (b + a) (c + a) = dist b c"
+  for a :: "'a::real_normed_vector"
+  by (simp add: dist_norm)
 
-lemma dist_scaleR [simp]:
-  fixes a :: "'a::real_normed_vector"
-  shows "dist (x *\<^sub>R a) (y *\<^sub>R a) = abs (x-y) * norm a"
-by (metis dist_norm norm_scaleR scaleR_left.diff)
+lemma dist_scaleR [simp]: "dist (x *\<^sub>R a) (y *\<^sub>R a) = \<bar>x - y\<bar> * norm a"
+  for a :: "'a::real_normed_vector"
+  by (metis dist_norm norm_scaleR scaleR_left.diff)
 
-lemma norm_uminus_minus: "norm (-x - y :: 'a :: real_normed_vector) = norm (x + y)"
+lemma norm_uminus_minus: "norm (- x - y :: 'a :: real_normed_vector) = norm (x + y)"
   by (subst (2) norm_minus_cancel[symmetric], subst minus_add_distrib) simp
 
-lemma norm_triangle_ineq2:
-  fixes a b :: "'a::real_normed_vector"
-  shows "norm a - norm b \<le> norm (a - b)"
+lemma norm_triangle_ineq2: "norm a - norm b \<le> norm (a - b)"
+  for a b :: "'a::real_normed_vector"
 proof -
   have "norm (a - b + b) \<le> norm (a - b) + norm b"
     by (rule norm_triangle_ineq)
-  thus ?thesis by simp
+  then show ?thesis by simp
 qed
 
-lemma norm_triangle_ineq3:
-  fixes a b :: "'a::real_normed_vector"
-  shows "\<bar>norm a - norm b\<bar> \<le> norm (a - b)"
-apply (subst abs_le_iff)
-apply auto
-apply (rule norm_triangle_ineq2)
-apply (subst norm_minus_commute)
-apply (rule norm_triangle_ineq2)
-done
+lemma norm_triangle_ineq3: "\<bar>norm a - norm b\<bar> \<le> norm (a - b)"
+  for a b :: "'a::real_normed_vector"
+  apply (subst abs_le_iff)
+  apply auto
+   apply (rule norm_triangle_ineq2)
+  apply (subst norm_minus_commute)
+  apply (rule norm_triangle_ineq2)
+  done
 
-lemma norm_triangle_ineq4:
-  fixes a b :: "'a::real_normed_vector"
-  shows "norm (a - b) \<le> norm a + norm b"
+lemma norm_triangle_ineq4: "norm (a - b) \<le> norm a + norm b"
+  for a b :: "'a::real_normed_vector"
 proof -
   have "norm (a + - b) \<le> norm a + norm (- b)"
     by (rule norm_triangle_ineq)
   then show ?thesis by simp
 qed
 
-lemma norm_diff_ineq:
-  fixes a b :: "'a::real_normed_vector"
-  shows "norm a - norm b \<le> norm (a + b)"
+lemma norm_diff_ineq: "norm a - norm b \<le> norm (a + b)"
+  for a b :: "'a::real_normed_vector"
 proof -
   have "norm a - norm (- b) \<le> norm (a - - b)"
     by (rule norm_triangle_ineq2)
-  thus ?thesis by simp
+  then show ?thesis by simp
 qed
 
-lemma norm_add_leD:
-  fixes a b :: "'a::real_normed_vector"
-  shows "norm (a + b) \<le> c \<Longrightarrow> norm b \<le> norm a + c"
-    by (metis add.commute diff_le_eq norm_diff_ineq order.trans)
+lemma norm_add_leD: "norm (a + b) \<le> c \<Longrightarrow> norm b \<le> norm a + c"
+  for a b :: "'a::real_normed_vector"
+  by (metis add.commute diff_le_eq norm_diff_ineq order.trans)
 
-lemma norm_diff_triangle_ineq:
-  fixes a b c d :: "'a::real_normed_vector"
-  shows "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
+lemma norm_diff_triangle_ineq: "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
+  for a b c d :: "'a::real_normed_vector"
 proof -
   have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))"
     by (simp add: algebra_simps)
@@ -856,19 +828,19 @@
 lemma norm_diff_triangle_le:
   fixes x y z :: "'a::real_normed_vector"
   assumes "norm (x - y) \<le> e1"  "norm (y - z) \<le> e2"
-    shows "norm (x - z) \<le> e1 + e2"
+  shows "norm (x - z) \<le> e1 + e2"
   using norm_diff_triangle_ineq [of x y y z] assms by simp
 
 lemma norm_diff_triangle_less:
   fixes x y z :: "'a::real_normed_vector"
   assumes "norm (x - y) < e1"  "norm (y - z) < e2"
-    shows "norm (x - z) < e1 + e2"
+  shows "norm (x - z) < e1 + e2"
   using norm_diff_triangle_ineq [of x y y z] assms by simp
 
 lemma norm_triangle_mono:
   fixes a b :: "'a::real_normed_vector"
-  shows "\<lbrakk>norm a \<le> r; norm b \<le> s\<rbrakk> \<Longrightarrow> norm (a + b) \<le> r + s"
-by (metis add_mono_thms_linordered_semiring(1) norm_triangle_ineq order.trans)
+  shows "norm a \<le> r \<Longrightarrow> norm b \<le> s \<Longrightarrow> norm (a + b) \<le> r + s"
+  by (metis add_mono_thms_linordered_semiring(1) norm_triangle_ineq order.trans)
 
 lemma norm_setsum:
   fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
@@ -881,82 +853,68 @@
   shows "norm (setsum f S) \<le> setsum g S"
   by (rule order_trans [OF norm_setsum setsum_mono]) (simp add: fg)
 
-lemma abs_norm_cancel [simp]:
-  fixes a :: "'a::real_normed_vector"
-  shows "\<bar>norm a\<bar> = norm a"
-by (rule abs_of_nonneg [OF norm_ge_zero])
+lemma abs_norm_cancel [simp]: "\<bar>norm a\<bar> = norm a"
+  for a :: "'a::real_normed_vector"
+  by (rule abs_of_nonneg [OF norm_ge_zero])
 
-lemma norm_add_less:
-  fixes x y :: "'a::real_normed_vector"
-  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x + y) < r + s"
-by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono])
+lemma norm_add_less: "norm x < r \<Longrightarrow> norm y < s \<Longrightarrow> norm (x + y) < r + s"
+  for x y :: "'a::real_normed_vector"
+  by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono])
 
-lemma norm_mult_less:
-  fixes x y :: "'a::real_normed_algebra"
-  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x * y) < r * s"
-apply (rule order_le_less_trans [OF norm_mult_ineq])
-apply (simp add: mult_strict_mono')
-done
+lemma norm_mult_less: "norm x < r \<Longrightarrow> norm y < s \<Longrightarrow> norm (x * y) < r * s"
+  for x y :: "'a::real_normed_algebra"
+  by (rule order_le_less_trans [OF norm_mult_ineq]) (simp add: mult_strict_mono')
 
-lemma norm_of_real [simp]:
-  "norm (of_real r :: 'a::real_normed_algebra_1) = \<bar>r\<bar>"
-unfolding of_real_def by simp
+lemma norm_of_real [simp]: "norm (of_real r :: 'a::real_normed_algebra_1) = \<bar>r\<bar>"
+  by (simp add: of_real_def)
 
-lemma norm_numeral [simp]:
-  "norm (numeral w::'a::real_normed_algebra_1) = numeral w"
-by (subst of_real_numeral [symmetric], subst norm_of_real, simp)
+lemma norm_numeral [simp]: "norm (numeral w::'a::real_normed_algebra_1) = numeral w"
+  by (subst of_real_numeral [symmetric], subst norm_of_real, simp)
 
-lemma norm_neg_numeral [simp]:
-  "norm (- numeral w::'a::real_normed_algebra_1) = numeral w"
-by (subst of_real_neg_numeral [symmetric], subst norm_of_real, simp)
+lemma norm_neg_numeral [simp]: "norm (- numeral w::'a::real_normed_algebra_1) = numeral w"
+  by (subst of_real_neg_numeral [symmetric], subst norm_of_real, simp)
 
-lemma norm_of_real_add1 [simp]:
-     "norm (of_real x + 1 :: 'a :: real_normed_div_algebra) = abs (x + 1)"
+lemma norm_of_real_add1 [simp]: "norm (of_real x + 1 :: 'a :: real_normed_div_algebra) = \<bar>x + 1\<bar>"
   by (metis norm_of_real of_real_1 of_real_add)
 
 lemma norm_of_real_addn [simp]:
-     "norm (of_real x + numeral b :: 'a :: real_normed_div_algebra) = abs (x + numeral b)"
+  "norm (of_real x + numeral b :: 'a :: real_normed_div_algebra) = \<bar>x + numeral b\<bar>"
   by (metis norm_of_real of_real_add of_real_numeral)
 
-lemma norm_of_int [simp]:
-  "norm (of_int z::'a::real_normed_algebra_1) = \<bar>of_int z\<bar>"
-by (subst of_real_of_int_eq [symmetric], rule norm_of_real)
+lemma norm_of_int [simp]: "norm (of_int z::'a::real_normed_algebra_1) = \<bar>of_int z\<bar>"
+  by (subst of_real_of_int_eq [symmetric], rule norm_of_real)
 
-lemma norm_of_nat [simp]:
-  "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n"
-apply (subst of_real_of_nat_eq [symmetric])
-apply (subst norm_of_real, simp)
-done
+lemma norm_of_nat [simp]: "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n"
+  apply (subst of_real_of_nat_eq [symmetric])
+  apply (subst norm_of_real, simp)
+  done
 
-lemma nonzero_norm_inverse:
-  fixes a :: "'a::real_normed_div_algebra"
-  shows "a \<noteq> 0 \<Longrightarrow> norm (inverse a) = inverse (norm a)"
-apply (rule inverse_unique [symmetric])
-apply (simp add: norm_mult [symmetric])
-done
+lemma nonzero_norm_inverse: "a \<noteq> 0 \<Longrightarrow> norm (inverse a) = inverse (norm a)"
+  for a :: "'a::real_normed_div_algebra"
+  apply (rule inverse_unique [symmetric])
+  apply (simp add: norm_mult [symmetric])
+  done
 
-lemma norm_inverse:
-  fixes a :: "'a::{real_normed_div_algebra, division_ring}"
-  shows "norm (inverse a) = inverse (norm a)"
-apply (case_tac "a = 0", simp)
-apply (erule nonzero_norm_inverse)
-done
+lemma norm_inverse: "norm (inverse a) = inverse (norm a)"
+  for a :: "'a::{real_normed_div_algebra,division_ring}"
+  apply (cases "a = 0")
+   apply simp
+  apply (erule nonzero_norm_inverse)
+  done
 
-lemma nonzero_norm_divide:
-  fixes a b :: "'a::real_normed_field"
-  shows "b \<noteq> 0 \<Longrightarrow> norm (a / b) = norm a / norm b"
-by (simp add: divide_inverse norm_mult nonzero_norm_inverse)
+lemma nonzero_norm_divide: "b \<noteq> 0 \<Longrightarrow> norm (a / b) = norm a / norm b"
+  for a b :: "'a::real_normed_field"
+  by (simp add: divide_inverse norm_mult nonzero_norm_inverse)
 
-lemma norm_divide:
-  fixes a b :: "'a::{real_normed_field, field}"
-  shows "norm (a / b) = norm a / norm b"
-by (simp add: divide_inverse norm_mult norm_inverse)
+lemma norm_divide: "norm (a / b) = norm a / norm b"
+  for a b :: "'a::{real_normed_field,field}"
+  by (simp add: divide_inverse norm_mult norm_inverse)
 
-lemma norm_power_ineq:
-  fixes x :: "'a::{real_normed_algebra_1}"
-  shows "norm (x ^ n) \<le> norm x ^ n"
+lemma norm_power_ineq: "norm (x ^ n) \<le> norm x ^ n"
+  for x :: "'a::real_normed_algebra_1"
 proof (induct n)
-  case 0 show "norm (x ^ 0) \<le> norm x ^ 0" by simp
+  case 0
+  show "norm (x ^ 0) \<le> norm x ^ 0" by simp
 next
   case (Suc n)
   have "norm (x * x ^ n) \<le> norm x * norm (x ^ n)"
@@ -967,10 +925,9 @@
     by simp
 qed
 
-lemma norm_power:
-  fixes x :: "'a::real_normed_div_algebra"
-  shows "norm (x ^ n) = norm x ^ n"
-by (induct n) (simp_all add: norm_mult)
+lemma norm_power: "norm (x ^ n) = norm x ^ n"
+  for x :: "'a::real_normed_div_algebra"
+  by (induct n) (simp_all add: norm_mult)
 
 lemma power_eq_imp_eq_norm:
   fixes w :: "'a::real_normed_div_algebra"
@@ -983,34 +940,31 @@
     using assms by (force intro: power_eq_imp_eq_base)
 qed
 
-lemma norm_mult_numeral1 [simp]:
-  fixes a b :: "'a::{real_normed_field, field}"
-  shows "norm (numeral w * a) = numeral w * norm a"
-by (simp add: norm_mult)
+lemma norm_mult_numeral1 [simp]: "norm (numeral w * a) = numeral w * norm a"
+  for a b :: "'a::{real_normed_field,field}"
+  by (simp add: norm_mult)
 
-lemma norm_mult_numeral2 [simp]:
-  fixes a b :: "'a::{real_normed_field, field}"
-  shows "norm (a * numeral w) = norm a * numeral w"
-by (simp add: norm_mult)
+lemma norm_mult_numeral2 [simp]: "norm (a * numeral w) = norm a * numeral w"
+  for a b :: "'a::{real_normed_field,field}"
+  by (simp add: norm_mult)
 
-lemma norm_divide_numeral [simp]:
-  fixes a b :: "'a::{real_normed_field, field}"
-  shows "norm (a / numeral w) = norm a / numeral w"
-by (simp add: norm_divide)
+lemma norm_divide_numeral [simp]: "norm (a / numeral w) = norm a / numeral w"
+  for a b :: "'a::{real_normed_field,field}"
+  by (simp add: norm_divide)
 
 lemma norm_of_real_diff [simp]:
-    "norm (of_real b - of_real a :: 'a::real_normed_algebra_1) \<le> \<bar>b - a\<bar>"
+  "norm (of_real b - of_real a :: 'a::real_normed_algebra_1) \<le> \<bar>b - a\<bar>"
   by (metis norm_of_real of_real_diff order_refl)
 
-text\<open>Despite a superficial resemblance, \<open>norm_eq_1\<close> is not relevant.\<close>
+text \<open>Despite a superficial resemblance, \<open>norm_eq_1\<close> is not relevant.\<close>
 lemma square_norm_one:
   fixes x :: "'a::real_normed_div_algebra"
-  assumes "x^2 = 1" shows "norm x = 1"
+  assumes "x\<^sup>2 = 1"
+  shows "norm x = 1"
   by (metis assms norm_minus_cancel norm_one power2_eq_1_iff)
 
-lemma norm_less_p1:
-  fixes x :: "'a::real_normed_algebra_1"
-  shows "norm x < norm (of_real (norm x) + 1 :: 'a)"
+lemma norm_less_p1: "norm x < norm (of_real (norm x) + 1 :: 'a)"
+  for x :: "'a::real_normed_algebra_1"
 proof -
   have "norm x < norm (of_real (norm x + 1) :: 'a)"
     by (simp add: of_real_def)
@@ -1018,14 +972,16 @@
     by simp
 qed
 
-lemma setprod_norm:
-  fixes f :: "'a \<Rightarrow> 'b::{comm_semiring_1,real_normed_div_algebra}"
-  shows "setprod (\<lambda>x. norm(f x)) A = norm (setprod f A)"
+lemma setprod_norm: "setprod (\<lambda>x. norm (f x)) A = norm (setprod f A)"
+  for f :: "'a \<Rightarrow> 'b::{comm_semiring_1,real_normed_div_algebra}"
   by (induct A rule: infinite_finite_induct) (auto simp: norm_mult)
 
 lemma norm_setprod_le:
-  "norm (setprod f A) \<le> (\<Prod>a\<in>A. norm (f a :: 'a :: {real_normed_algebra_1, comm_monoid_mult}))"
-proof (induction A rule: infinite_finite_induct)
+  "norm (setprod f A) \<le> (\<Prod>a\<in>A. norm (f a :: 'a :: {real_normed_algebra_1,comm_monoid_mult}))"
+proof (induct A rule: infinite_finite_induct)
+  case empty
+  then show ?case by simp
+next
   case (insert a A)
   then have "norm (setprod f (insert a A)) \<le> norm (f a) * norm (setprod f A)"
     by (simp add: norm_mult_ineq)
@@ -1033,13 +989,19 @@
     by (rule insert)
   finally show ?case
     by (simp add: insert mult_left_mono)
-qed simp_all
+next
+  case infinite
+  then show ?case by simp
+qed
 
 lemma norm_setprod_diff:
   fixes z w :: "'i \<Rightarrow> 'a::{real_normed_algebra_1, comm_monoid_mult}"
   shows "(\<And>i. i \<in> I \<Longrightarrow> norm (z i) \<le> 1) \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> norm (w i) \<le> 1) \<Longrightarrow>
     norm ((\<Prod>i\<in>I. z i) - (\<Prod>i\<in>I. w i)) \<le> (\<Sum>i\<in>I. norm (z i - w i))"
 proof (induction I rule: infinite_finite_induct)
+  case empty
+  then show ?case by simp
+next
   case (insert i I)
   note insert.hyps[simp]
 
@@ -1047,7 +1009,7 @@
     norm ((\<Prod>i\<in>I. z i) * (z i - w i) + ((\<Prod>i\<in>I. z i) - (\<Prod>i\<in>I. w i)) * w i)"
     (is "_ = norm (?t1 + ?t2)")
     by (auto simp add: field_simps)
-  also have "... \<le> norm ?t1 + norm ?t2"
+  also have "\<dots> \<le> norm ?t1 + norm ?t2"
     by (rule norm_triangle_ineq)
   also have "norm ?t1 \<le> norm (\<Prod>i\<in>I. z i) * norm (z i - w i)"
     by (rule norm_mult_ineq)
@@ -1063,7 +1025,10 @@
     using insert by auto
   finally show ?case
     by (auto simp add: ac_simps mult_right_mono mult_left_mono)
-qed simp_all
+next
+  case infinite
+  then show ?case by simp
+qed
 
 lemma norm_power_diff:
   fixes z w :: "'a::{real_normed_algebra_1, comm_monoid_mult}"
@@ -1079,27 +1044,28 @@
   finally show ?thesis .
 qed
 
+
 subsection \<open>Metric spaces\<close>
 
 class metric_space = uniformity_dist + open_uniformity +
   assumes dist_eq_0_iff [simp]: "dist x y = 0 \<longleftrightarrow> x = y"
-  assumes dist_triangle2: "dist x y \<le> dist x z + dist y z"
+    and dist_triangle2: "dist x y \<le> dist x z + dist y z"
 begin
 
 lemma dist_self [simp]: "dist x x = 0"
-by simp
+  by simp
 
 lemma zero_le_dist [simp]: "0 \<le> dist x y"
-using dist_triangle2 [of x x y] by simp
+  using dist_triangle2 [of x x y] by simp
 
 lemma zero_less_dist_iff: "0 < dist x y \<longleftrightarrow> x \<noteq> y"
-by (simp add: less_le)
+  by (simp add: less_le)
 
 lemma dist_not_less_zero [simp]: "\<not> dist x y < 0"
-by (simp add: not_less)
+  by (simp add: not_less)
 
 lemma dist_le_zero_iff [simp]: "dist x y \<le> 0 \<longleftrightarrow> x = y"
-by (simp add: le_less)
+  by (simp add: le_less)
 
 lemma dist_commute: "dist x y = dist y x"
 proof (rule order_antisym)
@@ -1118,56 +1084,52 @@
 lemma dist_triangle3: "dist x y \<le> dist a x + dist a y"
   using dist_triangle2 [of x y a] by (simp add: dist_commute)
 
-lemma dist_pos_lt:
-  shows "x \<noteq> y ==> 0 < dist x y"
-by (simp add: zero_less_dist_iff)
+lemma dist_pos_lt: "x \<noteq> y \<Longrightarrow> 0 < dist x y"
+  by (simp add: zero_less_dist_iff)
 
-lemma dist_nz:
-  shows "x \<noteq> y \<longleftrightarrow> 0 < dist x y"
-by (simp add: zero_less_dist_iff)
+lemma dist_nz: "x \<noteq> y \<longleftrightarrow> 0 < dist x y"
+  by (simp add: zero_less_dist_iff)
 
 declare dist_nz [symmetric, simp]
 
-lemma dist_triangle_le:
-  shows "dist x z + dist y z <= e \<Longrightarrow> dist x y <= e"
-by (rule order_trans [OF dist_triangle2])
+lemma dist_triangle_le: "dist x z + dist y z \<le> e \<Longrightarrow> dist x y \<le> e"
+  by (rule order_trans [OF dist_triangle2])
 
-lemma dist_triangle_lt:
-  shows "dist x z + dist y z < e ==> dist x y < e"
-by (rule le_less_trans [OF dist_triangle2])
+lemma dist_triangle_lt: "dist x z + dist y z < e \<Longrightarrow> dist x y < e"
+  by (rule le_less_trans [OF dist_triangle2])
 
-lemma dist_triangle_less_add:
-   "\<lbrakk>dist x1 y < e1; dist x2 y < e2\<rbrakk> \<Longrightarrow> dist x1 x2 < e1 + e2"
-by (rule dist_triangle_lt [where z=y], simp)
+lemma dist_triangle_less_add: "dist x1 y < e1 \<Longrightarrow> dist x2 y < e2 \<Longrightarrow> dist x1 x2 < e1 + e2"
+  by (rule dist_triangle_lt [where z=y]) simp
 
-lemma dist_triangle_half_l:
-  shows "dist x1 y < e / 2 \<Longrightarrow> dist x2 y < e / 2 \<Longrightarrow> dist x1 x2 < e"
-by (rule dist_triangle_lt [where z=y], simp)
+lemma dist_triangle_half_l: "dist x1 y < e / 2 \<Longrightarrow> dist x2 y < e / 2 \<Longrightarrow> dist x1 x2 < e"
+  by (rule dist_triangle_lt [where z=y]) simp
 
-lemma dist_triangle_half_r:
-  shows "dist y x1 < e / 2 \<Longrightarrow> dist y x2 < e / 2 \<Longrightarrow> dist x1 x2 < e"
-by (rule dist_triangle_half_l, simp_all add: dist_commute)
+lemma dist_triangle_half_r: "dist y x1 < e / 2 \<Longrightarrow> dist y x2 < e / 2 \<Longrightarrow> dist x1 x2 < e"
+  by (rule dist_triangle_half_l) (simp_all add: dist_commute)
 
 subclass uniform_space
 proof
-  fix E x assume "eventually E uniformity"
+  fix E x
+  assume "eventually E uniformity"
   then obtain e where E: "0 < e" "\<And>x y. dist x y < e \<Longrightarrow> E (x, y)"
-    unfolding eventually_uniformity_metric by auto
+    by (auto simp: eventually_uniformity_metric)
   then show "E (x, x)" "\<forall>\<^sub>F (x, y) in uniformity. E (y, x)"
-    unfolding eventually_uniformity_metric by (auto simp: dist_commute)
-
+    by (auto simp: eventually_uniformity_metric dist_commute)
   show "\<exists>D. eventually D uniformity \<and> (\<forall>x y z. D (x, y) \<longrightarrow> D (y, z) \<longrightarrow> E (x, z))"
-    using E dist_triangle_half_l[where e=e] unfolding eventually_uniformity_metric
+    using E dist_triangle_half_l[where e=e]
+    unfolding eventually_uniformity_metric
     by (intro exI[of _ "\<lambda>(x, y). dist x y < e / 2"] exI[of _ "e/2"] conjI)
-       (auto simp: dist_commute)
+      (auto simp: dist_commute)
 qed
 
 lemma open_dist: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
-  unfolding open_uniformity eventually_uniformity_metric by (simp add: dist_commute)
+  by (simp add: dist_commute open_uniformity eventually_uniformity_metric)
 
 lemma open_ball: "open {y. dist x y < d}"
-proof (unfold open_dist, intro ballI)
-  fix y assume *: "y \<in> {y. dist x y < d}"
+  unfolding open_dist
+proof (intro ballI)
+  fix y
+  assume *: "y \<in> {y. dist x y < d}"
   then show "\<exists>e>0. \<forall>z. dist z y < e \<longrightarrow> z \<in> {y. dist x y < d}"
     by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt)
 qed
@@ -1177,7 +1139,8 @@
   fix x
   show "\<exists>A::nat \<Rightarrow> 'a set. (\<forall>i. x \<in> A i \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))"
   proof (safe intro!: exI[of _ "\<lambda>n. {y. dist x y < inverse (Suc n)}"])
-    fix S assume "open S" "x \<in> S"
+    fix S
+    assume "open S" "x \<in> S"
     then obtain e where e: "0 < e" and "{y. dist x y < e} \<subseteq> S"
       by (auto simp: open_dist subset_eq dist_commute)
     moreover
@@ -1198,34 +1161,33 @@
   assume xy: "x \<noteq> y"
   let ?U = "{y'. dist x y' < dist x y / 2}"
   let ?V = "{x'. dist y x' < dist x y / 2}"
-  have th0: "\<And>d x y z. (d x z :: real) \<le> d x y + d y z \<Longrightarrow> d y z = d z y
-               \<Longrightarrow> \<not>(d x y * 2 < d x z \<and> d z y * 2 < d x z)" by arith
+  have *: "d x z \<le> d x y + d y z \<Longrightarrow> d y z = d z y \<Longrightarrow> \<not> (d x y * 2 < d x z \<and> d z y * 2 < d x z)"
+    for d :: "'a \<Rightarrow> 'a \<Rightarrow> real" and x y z :: 'a
+    by arith
   have "open ?U \<and> open ?V \<and> x \<in> ?U \<and> y \<in> ?V \<and> ?U \<inter> ?V = {}"
-    using dist_pos_lt[OF xy] th0[of dist, OF dist_triangle dist_commute]
+    using dist_pos_lt[OF xy] *[of dist, OF dist_triangle dist_commute]
     using open_ball[of _ "dist x y / 2"] by auto
   then show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
     by blast
 qed
 
 text \<open>Every normed vector space is a metric space.\<close>
-
 instance real_normed_vector < metric_space
 proof
-  fix x y :: 'a show "dist x y = 0 \<longleftrightarrow> x = y"
-    unfolding dist_norm by simp
-next
-  fix x y z :: 'a show "dist x y \<le> dist x z + dist y z"
-    unfolding dist_norm
-    using norm_triangle_ineq4 [of "x - z" "y - z"] by simp
+  fix x y z :: 'a
+  show "dist x y = 0 \<longleftrightarrow> x = y"
+    by (simp add: dist_norm)
+  show "dist x y \<le> dist x z + dist y z"
+    using norm_triangle_ineq4 [of "x - z" "y - z"] by (simp add: dist_norm)
 qed
 
+
 subsection \<open>Class instances for real numbers\<close>
 
 instantiation real :: real_normed_field
 begin
 
-definition dist_real_def:
-  "dist x y = \<bar>x - y\<bar>"
+definition dist_real_def: "dist x y = \<bar>x - y\<bar>"
 
 definition uniformity_real_def [code del]:
   "(uniformity :: (real \<times> real) filter) = (INF e:{0 <..}. principal {(x, y). dist x y < e})"
@@ -1233,29 +1195,28 @@
 definition open_real_def [code del]:
   "open (U :: real set) \<longleftrightarrow> (\<forall>x\<in>U. eventually (\<lambda>(x', y). x' = x \<longrightarrow> y \<in> U) uniformity)"
 
-definition real_norm_def [simp]:
-  "norm r = \<bar>r\<bar>"
+definition real_norm_def [simp]: "norm r = \<bar>r\<bar>"
 
 instance
-apply (intro_classes, unfold real_norm_def real_scaleR_def)
-apply (rule dist_real_def)
-apply (simp add: sgn_real_def)
-apply (rule uniformity_real_def)
-apply (rule open_real_def)
-apply (rule abs_eq_0)
-apply (rule abs_triangle_ineq)
-apply (rule abs_mult)
-apply (rule abs_mult)
-done
+  apply intro_classes
+         apply (unfold real_norm_def real_scaleR_def)
+         apply (rule dist_real_def)
+        apply (simp add: sgn_real_def)
+       apply (rule uniformity_real_def)
+      apply (rule open_real_def)
+     apply (rule abs_eq_0)
+    apply (rule abs_triangle_ineq)
+   apply (rule abs_mult)
+  apply (rule abs_mult)
+  done
 
 end
 
 declare uniformity_Abort[where 'a=real, code]
 
-lemma dist_of_real [simp]:
-  fixes a :: "'a::real_normed_div_algebra"
-  shows "dist (of_real x :: 'a) (of_real y) = dist x y"
-by (metis dist_norm norm_of_real of_real_diff real_norm_def)
+lemma dist_of_real [simp]: "dist (of_real x :: 'a) (of_real y) = dist x y"
+  for a :: "'a::real_normed_div_algebra"
+  by (metis dist_norm norm_of_real of_real_diff real_norm_def)
 
 declare [[code abort: "open :: real set \<Rightarrow> bool"]]
 
@@ -1263,7 +1224,8 @@
 proof
   show "(open :: real set \<Rightarrow> bool) = generate_topology (range lessThan \<union> range greaterThan)"
   proof (rule ext, safe)
-    fix S :: "real set" assume "open S"
+    fix S :: "real set"
+    assume "open S"
     then obtain f where "\<forall>x\<in>S. 0 < f x \<and> (\<forall>y. dist y x < f x \<longrightarrow> y \<in> S)"
       unfolding open_dist bchoice_iff ..
     then have *: "S = (\<Union>x\<in>S. {x - f x <..} \<inter> {..< x + f x})"
@@ -1271,23 +1233,26 @@
     show "generate_topology (range lessThan \<union> range greaterThan) S"
       apply (subst *)
       apply (intro generate_topology_Union generate_topology.Int)
-      apply (auto intro: generate_topology.Basis)
+       apply (auto intro: generate_topology.Basis)
       done
   next
-    fix S :: "real set" assume "generate_topology (range lessThan \<union> range greaterThan) S"
+    fix S :: "real set"
+    assume "generate_topology (range lessThan \<union> range greaterThan) S"
     moreover have "\<And>a::real. open {..<a}"
       unfolding open_dist dist_real_def
     proof clarify
-      fix x a :: real assume "x < a"
-      hence "0 < a - x \<and> (\<forall>y. \<bar>y - x\<bar> < a - x \<longrightarrow> y \<in> {..<a})" by auto
-      thus "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {..<a}" ..
+      fix x a :: real
+      assume "x < a"
+      then have "0 < a - x \<and> (\<forall>y. \<bar>y - x\<bar> < a - x \<longrightarrow> y \<in> {..<a})" by auto
+      then show "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {..<a}" ..
     qed
     moreover have "\<And>a::real. open {a <..}"
       unfolding open_dist dist_real_def
     proof clarify
-      fix x a :: real assume "a < x"
-      hence "0 < x - a \<and> (\<forall>y. \<bar>y - x\<bar> < x - a \<longrightarrow> y \<in> {a<..})" by auto
-      thus "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {a<..}" ..
+      fix x a :: real
+      assume "a < x"
+      then have "0 < x - a \<and> (\<forall>y. \<bar>y - x\<bar> < x - a \<longrightarrow> y \<in> {a<..})" by auto
+      then show "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {a<..}" ..
     qed
     ultimately show "open S"
       by induct auto
@@ -1303,66 +1268,67 @@
 lemmas closed_real_atLeast = closed_atLeast[where 'a=real]
 lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real]
 
+
 subsection \<open>Extra type constraints\<close>
 
 text \<open>Only allow @{term "open"} in class \<open>topological_space\<close>.\<close>
-
 setup \<open>Sign.add_const_constraint
   (@{const_name "open"}, SOME @{typ "'a::topological_space set \<Rightarrow> bool"})\<close>
 
 text \<open>Only allow @{term "uniformity"} in class \<open>uniform_space\<close>.\<close>
-
 setup \<open>Sign.add_const_constraint
   (@{const_name "uniformity"}, SOME @{typ "('a::uniformity \<times> 'a) filter"})\<close>
 
 text \<open>Only allow @{term dist} in class \<open>metric_space\<close>.\<close>
-
 setup \<open>Sign.add_const_constraint
   (@{const_name dist}, SOME @{typ "'a::metric_space \<Rightarrow> 'a \<Rightarrow> real"})\<close>
 
 text \<open>Only allow @{term norm} in class \<open>real_normed_vector\<close>.\<close>
-
 setup \<open>Sign.add_const_constraint
   (@{const_name norm}, SOME @{typ "'a::real_normed_vector \<Rightarrow> real"})\<close>
 
+
 subsection \<open>Sign function\<close>
 
-lemma norm_sgn:
-  "norm (sgn(x::'a::real_normed_vector)) = (if x = 0 then 0 else 1)"
-by (simp add: sgn_div_norm)
+lemma norm_sgn: "norm (sgn x) = (if x = 0 then 0 else 1)"
+  for x :: "'a::real_normed_vector"
+  by (simp add: sgn_div_norm)
 
-lemma sgn_zero [simp]: "sgn(0::'a::real_normed_vector) = 0"
-by (simp add: sgn_div_norm)
+lemma sgn_zero [simp]: "sgn (0::'a::real_normed_vector) = 0"
+  by (simp add: sgn_div_norm)
 
-lemma sgn_zero_iff: "(sgn(x::'a::real_normed_vector) = 0) = (x = 0)"
-by (simp add: sgn_div_norm)
+lemma sgn_zero_iff: "sgn x = 0 \<longleftrightarrow> x = 0"
+  for x :: "'a::real_normed_vector"
+  by (simp add: sgn_div_norm)
 
-lemma sgn_minus: "sgn (- x) = - sgn(x::'a::real_normed_vector)"
-by (simp add: sgn_div_norm)
+lemma sgn_minus: "sgn (- x) = - sgn x"
+  for x :: "'a::real_normed_vector"
+  by (simp add: sgn_div_norm)
 
-lemma sgn_scaleR:
-  "sgn (scaleR r x) = scaleR (sgn r) (sgn(x::'a::real_normed_vector))"
-by (simp add: sgn_div_norm ac_simps)
+lemma sgn_scaleR: "sgn (scaleR r x) = scaleR (sgn r) (sgn x)"
+  for x :: "'a::real_normed_vector"
+  by (simp add: sgn_div_norm ac_simps)
 
 lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1"
-by (simp add: sgn_div_norm)
+  by (simp add: sgn_div_norm)
 
-lemma sgn_of_real:
-  "sgn (of_real r::'a::real_normed_algebra_1) = of_real (sgn r)"
-unfolding of_real_def by (simp only: sgn_scaleR sgn_one)
+lemma sgn_of_real: "sgn (of_real r :: 'a::real_normed_algebra_1) = of_real (sgn r)"
+  unfolding of_real_def by (simp only: sgn_scaleR sgn_one)
 
-lemma sgn_mult:
-  fixes x y :: "'a::real_normed_div_algebra"
-  shows "sgn (x * y) = sgn x * sgn y"
-by (simp add: sgn_div_norm norm_mult mult.commute)
+lemma sgn_mult: "sgn (x * y) = sgn x * sgn y"
+  for x y :: "'a::real_normed_div_algebra"
+  by (simp add: sgn_div_norm norm_mult mult.commute)
 
-lemma real_sgn_eq: "sgn (x::real) = x / \<bar>x\<bar>"
+lemma real_sgn_eq: "sgn x = x / \<bar>x\<bar>"
+  for x :: real
   by (simp add: sgn_div_norm divide_inverse)
 
-lemma zero_le_sgn_iff [simp]: "0 \<le> sgn x \<longleftrightarrow> 0 \<le> (x::real)"
+lemma zero_le_sgn_iff [simp]: "0 \<le> sgn x \<longleftrightarrow> 0 \<le> x"
+  for x :: real
   by (cases "0::real" x rule: linorder_cases) simp_all
 
-lemma sgn_le_0_iff [simp]: "sgn x \<le> 0 \<longleftrightarrow> (x::real) \<le> 0"
+lemma sgn_le_0_iff [simp]: "sgn x \<le> 0 \<longleftrightarrow> x \<le> 0"
+  for x :: real
   by (cases "0::real" x rule: linorder_cases) simp_all
 
 lemma norm_conv_dist: "norm x = dist x 0"
@@ -1370,10 +1336,9 @@
 
 declare norm_conv_dist [symmetric, simp]
 
-lemma dist_0_norm [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "dist 0 x = norm x"
-unfolding dist_norm by simp
+lemma dist_0_norm [simp]: "dist 0 x = norm x"
+  for x :: "'a::real_normed_vector"
+  by (simp add: dist_norm)
 
 lemma dist_diff [simp]: "dist a (a - b) = norm b"  "dist (a - b) a = norm b"
   by (simp_all add: dist_norm)
@@ -1390,23 +1355,25 @@
   "dist (of_nat m) (of_nat n :: 'a :: real_normed_algebra_1) = of_int \<bar>int m - int n\<bar>"
   by (subst (1 2) of_int_of_nat_eq [symmetric]) (rule dist_of_int)
 
+
 subsection \<open>Bounded Linear and Bilinear Operators\<close>
 
 locale linear = additive f for f :: "'a::real_vector \<Rightarrow> 'b::real_vector" +
   assumes scaleR: "f (scaleR r x) = scaleR r (f x)"
 
 lemma linear_imp_scaleR:
-  assumes "linear D" obtains d where "D = (\<lambda>x. x *\<^sub>R d)"
+  assumes "linear D"
+  obtains d where "D = (\<lambda>x. x *\<^sub>R d)"
   by (metis assms linear.scaleR mult.commute mult.left_neutral real_scaleR_def)
 
 corollary real_linearD:
   fixes f :: "real \<Rightarrow> real"
   assumes "linear f" obtains c where "f = op* c"
-by (rule linear_imp_scaleR [OF assms]) (force simp: scaleR_conv_of_real)
+  by (rule linear_imp_scaleR [OF assms]) (force simp: scaleR_conv_of_real)
 
 lemma linearI:
   assumes "\<And>x y. f (x + y) = f x + f y"
-  assumes "\<And>c x. f (c *\<^sub>R x) = c *\<^sub>R f x"
+    and "\<And>c x. f (c *\<^sub>R x) = c *\<^sub>R f x"
   shows "linear f"
   by standard (rule assms)+
 
@@ -1414,8 +1381,7 @@
   assumes bounded: "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
 begin
 
-lemma pos_bounded:
-  "\<exists>K>0. \<forall>x. norm (f x) \<le> norm x * K"
+lemma pos_bounded: "\<exists>K>0. \<forall>x. norm (f x) \<le> norm x * K"
 proof -
   obtain K where K: "\<And>x. norm (f x) \<le> norm x * K"
     using bounded by blast
@@ -1432,118 +1398,107 @@
   qed
 qed
 
-lemma nonneg_bounded:
-  "\<exists>K\<ge>0. \<forall>x. norm (f x) \<le> norm x * K"
-proof -
-  from pos_bounded
-  show ?thesis by (auto intro: order_less_imp_le)
-qed
+lemma nonneg_bounded: "\<exists>K\<ge>0. \<forall>x. norm (f x) \<le> norm x * K"
+  using pos_bounded by (auto intro: order_less_imp_le)
 
-lemma linear: "linear f" 
+lemma linear: "linear f"
   by (fact local.linear_axioms)
 
 end
 
 lemma bounded_linear_intro:
   assumes "\<And>x y. f (x + y) = f x + f y"
-  assumes "\<And>r x. f (scaleR r x) = scaleR r (f x)"
-  assumes "\<And>x. norm (f x) \<le> norm x * K"
+    and "\<And>r x. f (scaleR r x) = scaleR r (f x)"
+    and "\<And>x. norm (f x) \<le> norm x * K"
   shows "bounded_linear f"
   by standard (blast intro: assms)+
 
 locale bounded_bilinear =
-  fixes prod :: "['a::real_normed_vector, 'b::real_normed_vector]
-                 \<Rightarrow> 'c::real_normed_vector"
+  fixes prod :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector \<Rightarrow> 'c::real_normed_vector"
     (infixl "**" 70)
   assumes add_left: "prod (a + a') b = prod a b + prod a' b"
-  assumes add_right: "prod a (b + b') = prod a b + prod a b'"
-  assumes scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)"
-  assumes scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)"
-  assumes bounded: "\<exists>K. \<forall>a b. norm (prod a b) \<le> norm a * norm b * K"
+    and add_right: "prod a (b + b') = prod a b + prod a b'"
+    and scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)"
+    and scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)"
+    and bounded: "\<exists>K. \<forall>a b. norm (prod a b) \<le> norm a * norm b * K"
 begin
 
-lemma pos_bounded:
-  "\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
-apply (cut_tac bounded, erule exE)
-apply (rule_tac x="max 1 K" in exI, safe)
-apply (rule order_less_le_trans [OF zero_less_one max.cobounded1])
-apply (drule spec, drule spec, erule order_trans)
-apply (rule mult_left_mono [OF max.cobounded2])
-apply (intro mult_nonneg_nonneg norm_ge_zero)
-done
+lemma pos_bounded: "\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
+  apply (insert bounded)
+  apply (erule exE)
+  apply (rule_tac x="max 1 K" in exI)
+  apply safe
+   apply (rule order_less_le_trans [OF zero_less_one max.cobounded1])
+  apply (drule spec)
+  apply (drule spec)
+  apply (erule order_trans)
+  apply (rule mult_left_mono [OF max.cobounded2])
+  apply (intro mult_nonneg_nonneg norm_ge_zero)
+  done
 
-lemma nonneg_bounded:
-  "\<exists>K\<ge>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
-proof -
-  from pos_bounded
-  show ?thesis by (auto intro: order_less_imp_le)
-qed
+lemma nonneg_bounded: "\<exists>K\<ge>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
+  using pos_bounded by (auto intro: order_less_imp_le)
 
 lemma additive_right: "additive (\<lambda>b. prod a b)"
-by (rule additive.intro, rule add_right)
+  by (rule additive.intro, rule add_right)
 
 lemma additive_left: "additive (\<lambda>a. prod a b)"
-by (rule additive.intro, rule add_left)
+  by (rule additive.intro, rule add_left)
 
 lemma zero_left: "prod 0 b = 0"
-by (rule additive.zero [OF additive_left])
+  by (rule additive.zero [OF additive_left])
 
 lemma zero_right: "prod a 0 = 0"
-by (rule additive.zero [OF additive_right])
+  by (rule additive.zero [OF additive_right])
 
 lemma minus_left: "prod (- a) b = - prod a b"
-by (rule additive.minus [OF additive_left])
+  by (rule additive.minus [OF additive_left])
 
 lemma minus_right: "prod a (- b) = - prod a b"
-by (rule additive.minus [OF additive_right])
+  by (rule additive.minus [OF additive_right])
 
-lemma diff_left:
-  "prod (a - a') b = prod a b - prod a' b"
-by (rule additive.diff [OF additive_left])
+lemma diff_left: "prod (a - a') b = prod a b - prod a' b"
+  by (rule additive.diff [OF additive_left])
 
-lemma diff_right:
-  "prod a (b - b') = prod a b - prod a b'"
-by (rule additive.diff [OF additive_right])
+lemma diff_right: "prod a (b - b') = prod a b - prod a b'"
+  by (rule additive.diff [OF additive_right])
 
-lemma setsum_left:
-  "prod (setsum g S) x = setsum ((\<lambda>i. prod (g i) x)) S"
-by (rule additive.setsum [OF additive_left])
+lemma setsum_left: "prod (setsum g S) x = setsum ((\<lambda>i. prod (g i) x)) S"
+  by (rule additive.setsum [OF additive_left])
 
-lemma setsum_right:
-  "prod x (setsum g S) = setsum ((\<lambda>i. (prod x (g i)))) S"
-by (rule additive.setsum [OF additive_right])
+lemma setsum_right: "prod x (setsum g S) = setsum ((\<lambda>i. (prod x (g i)))) S"
+  by (rule additive.setsum [OF additive_right])
 
 
-lemma bounded_linear_left:
-  "bounded_linear (\<lambda>a. a ** b)"
-apply (cut_tac bounded, safe)
-apply (rule_tac K="norm b * K" in bounded_linear_intro)
-apply (rule add_left)
-apply (rule scaleR_left)
-apply (simp add: ac_simps)
-done
+lemma bounded_linear_left: "bounded_linear (\<lambda>a. a ** b)"
+  apply (insert bounded)
+  apply safe
+  apply (rule_tac K="norm b * K" in bounded_linear_intro)
+    apply (rule add_left)
+   apply (rule scaleR_left)
+  apply (simp add: ac_simps)
+  done
 
-lemma bounded_linear_right:
-  "bounded_linear (\<lambda>b. a ** b)"
-apply (cut_tac bounded, safe)
-apply (rule_tac K="norm a * K" in bounded_linear_intro)
-apply (rule add_right)
-apply (rule scaleR_right)
-apply (simp add: ac_simps)
-done
+lemma bounded_linear_right: "bounded_linear (\<lambda>b. a ** b)"
+  apply (insert bounded)
+  apply safe
+  apply (rule_tac K="norm a * K" in bounded_linear_intro)
+    apply (rule add_right)
+   apply (rule scaleR_right)
+  apply (simp add: ac_simps)
+  done
 
-lemma prod_diff_prod:
-  "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)"
-by (simp add: diff_left diff_right)
+lemma prod_diff_prod: "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)"
+  by (simp add: diff_left diff_right)
 
 lemma flip: "bounded_bilinear (\<lambda>x y. y ** x)"
   apply standard
-  apply (rule add_right)
-  apply (rule add_left)
-  apply (rule scaleR_right)
-  apply (rule scaleR_left)
+      apply (rule add_right)
+     apply (rule add_left)
+    apply (rule scaleR_right)
+   apply (rule scaleR_left)
   apply (subst mult.commute)
-  using bounded
+  apply (insert bounded)
   apply blast
   done
 
@@ -1557,11 +1512,10 @@
     "\<And>r a b. g (r *\<^sub>R a) ** b = r *\<^sub>R (g a ** b)"
     "\<And>a r b. g a ** (r *\<^sub>R b) = r *\<^sub>R (g a ** b)"
     by (auto simp: g.add add_left add_right g.scaleR scaleR_left scaleR_right)
-  from g.nonneg_bounded nonneg_bounded
-  obtain K L
-  where nn: "0 \<le> K" "0 \<le> L"
-    and K: "\<And>x. norm (g x) \<le> norm x * K"
-    and L: "\<And>a b. norm (a ** b) \<le> norm a * norm b * L"
+  from g.nonneg_bounded nonneg_bounded obtain K L
+    where nn: "0 \<le> K" "0 \<le> L"
+      and K: "\<And>x. norm (g x) \<le> norm x * K"
+      and L: "\<And>a b. norm (a ** b) \<le> norm a * norm b * L"
     by auto
   have "norm (g a ** b) \<le> norm a * K * norm b * L" for a b
     by (auto intro!:  order_trans[OF K] order_trans[OF L] mult_mono simp: nn)
@@ -1569,8 +1523,7 @@
     by (auto intro!: exI[where x="K * L"] simp: ac_simps)
 qed
 
-lemma comp:
-  "bounded_linear f \<Longrightarrow> bounded_linear g \<Longrightarrow> bounded_bilinear (\<lambda>x y. f x ** g y)"
+lemma comp: "bounded_linear f \<Longrightarrow> bounded_linear g \<Longrightarrow> bounded_bilinear (\<lambda>x y. f x ** g y)"
   by (rule bounded_bilinear.flip[OF bounded_bilinear.comp1[OF bounded_bilinear.flip[OF comp1]]])
 
 end
@@ -1583,15 +1536,17 @@
 
 lemma bounded_linear_add:
   assumes "bounded_linear f"
-  assumes "bounded_linear g"
+    and "bounded_linear g"
   shows "bounded_linear (\<lambda>x. f x + g x)"
 proof -
   interpret f: bounded_linear f by fact
   interpret g: bounded_linear g by fact
   show ?thesis
   proof
-    from f.bounded obtain Kf where Kf: "\<And>x. norm (f x) \<le> norm x * Kf" by blast
-    from g.bounded obtain Kg where Kg: "\<And>x. norm (g x) \<le> norm x * Kg" by blast
+    from f.bounded obtain Kf where Kf: "norm (f x) \<le> norm x * Kf" for x
+      by blast
+    from g.bounded obtain Kg where Kg: "norm (g x) \<le> norm x * Kg" for x
+      by blast
     show "\<exists>K. \<forall>x. norm (f x + g x) \<le> norm x * K"
       using add_mono[OF Kf Kg]
       by (intro exI[of _ "Kf + Kg"]) (auto simp: field_simps intro: norm_triangle_ineq order_trans)
@@ -1603,9 +1558,10 @@
   shows "bounded_linear (\<lambda>x. - f x)"
 proof -
   interpret f: bounded_linear f by fact
-  show ?thesis apply (unfold_locales)
-    apply (simp add: f.add)
-    apply (simp add: f.scaleR)
+  show ?thesis
+    apply unfold_locales
+      apply (simp add: f.add)
+     apply (simp add: f.scaleR)
     apply (simp add: f.bounded)
     done
 qed
@@ -1618,31 +1574,32 @@
   fixes f :: "'i \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   assumes "\<And>i. i \<in> I \<Longrightarrow> bounded_linear (f i)"
   shows "bounded_linear (\<lambda>x. \<Sum>i\<in>I. f i x)"
-proof cases
-  assume "finite I"
-  from this show ?thesis
-    using assms
-    by (induct I) (auto intro!: bounded_linear_add)
-qed simp
+proof (cases "finite I")
+  case True
+  then show ?thesis
+    using assms by induct (auto intro!: bounded_linear_add)
+next
+  case False
+  then show ?thesis by simp
+qed
 
 lemma bounded_linear_compose:
   assumes "bounded_linear f"
-  assumes "bounded_linear g"
+    and "bounded_linear g"
   shows "bounded_linear (\<lambda>x. f (g x))"
 proof -
   interpret f: bounded_linear f by fact
   interpret g: bounded_linear g by fact
-  show ?thesis proof (unfold_locales)
-    fix x y show "f (g (x + y)) = f (g x) + f (g y)"
+  show ?thesis
+  proof unfold_locales
+    show "f (g (x + y)) = f (g x) + f (g y)" for x y
       by (simp only: f.add g.add)
-  next
-    fix r x show "f (g (scaleR r x)) = scaleR r (f (g x))"
+    show "f (g (scaleR r x)) = scaleR r (f (g x))" for r x
       by (simp only: f.scaleR g.scaleR)
-  next
-    from f.pos_bounded
-    obtain Kf where f: "\<And>x. norm (f x) \<le> norm x * Kf" and Kf: "0 < Kf" by blast
-    from g.pos_bounded
-    obtain Kg where g: "\<And>x. norm (g x) \<le> norm x * Kg" by blast
+    from f.pos_bounded obtain Kf where f: "\<And>x. norm (f x) \<le> norm x * Kf" and Kf: "0 < Kf"
+      by blast
+    from g.pos_bounded obtain Kg where g: "\<And>x. norm (g x) \<le> norm x * Kg"
+      by blast
     show "\<exists>K. \<forall>x. norm (f (g x)) \<le> norm x * K"
     proof (intro exI allI)
       fix x
@@ -1657,24 +1614,21 @@
   qed
 qed
 
-lemma bounded_bilinear_mult:
-  "bounded_bilinear (op * :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra)"
-apply (rule bounded_bilinear.intro)
-apply (rule distrib_right)
-apply (rule distrib_left)
-apply (rule mult_scaleR_left)
-apply (rule mult_scaleR_right)
-apply (rule_tac x="1" in exI)
-apply (simp add: norm_mult_ineq)
-done
+lemma bounded_bilinear_mult: "bounded_bilinear (op * :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra)"
+  apply (rule bounded_bilinear.intro)
+      apply (rule distrib_right)
+     apply (rule distrib_left)
+    apply (rule mult_scaleR_left)
+   apply (rule mult_scaleR_right)
+  apply (rule_tac x="1" in exI)
+  apply (simp add: norm_mult_ineq)
+  done
 
-lemma bounded_linear_mult_left:
-  "bounded_linear (\<lambda>x::'a::real_normed_algebra. x * y)"
+lemma bounded_linear_mult_left: "bounded_linear (\<lambda>x::'a::real_normed_algebra. x * y)"
   using bounded_bilinear_mult
   by (rule bounded_bilinear.bounded_linear_left)
 
-lemma bounded_linear_mult_right:
-  "bounded_linear (\<lambda>y::'a::real_normed_algebra. x * y)"
+lemma bounded_linear_mult_right: "bounded_linear (\<lambda>y::'a::real_normed_algebra. x * y)"
   using bounded_bilinear_mult
   by (rule bounded_bilinear.bounded_linear_right)
 
@@ -1684,18 +1638,19 @@
 lemmas bounded_linear_const_mult =
   bounded_linear_mult_right [THEN bounded_linear_compose]
 
-lemma bounded_linear_divide:
-  "bounded_linear (\<lambda>x::'a::real_normed_field. x / y)"
+lemma bounded_linear_divide: "bounded_linear (\<lambda>x. x / y)"
+  for y :: "'a::real_normed_field"
   unfolding divide_inverse by (rule bounded_linear_mult_left)
 
 lemma bounded_bilinear_scaleR: "bounded_bilinear scaleR"
-apply (rule bounded_bilinear.intro)
-apply (rule scaleR_left_distrib)
-apply (rule scaleR_right_distrib)
-apply simp
-apply (rule scaleR_left_commute)
-apply (rule_tac x="1" in exI, simp)
-done
+  apply (rule bounded_bilinear.intro)
+      apply (rule scaleR_left_distrib)
+     apply (rule scaleR_right_distrib)
+    apply simp
+   apply (rule scaleR_left_commute)
+  apply (rule_tac x="1" in exI)
+  apply simp
+  done
 
 lemma bounded_linear_scaleR_left: "bounded_linear (\<lambda>r. scaleR r x)"
   using bounded_bilinear_scaleR
@@ -1714,48 +1669,53 @@
 lemma bounded_linear_of_real: "bounded_linear (\<lambda>r. of_real r)"
   unfolding of_real_def by (rule bounded_linear_scaleR_left)
 
-lemma real_bounded_linear:
-  fixes f :: "real \<Rightarrow> real"
-  shows "bounded_linear f \<longleftrightarrow> (\<exists>c::real. f = (\<lambda>x. x * c))"
+lemma real_bounded_linear: "bounded_linear f \<longleftrightarrow> (\<exists>c::real. f = (\<lambda>x. x * c))"
+  for f :: "real \<Rightarrow> real"
 proof -
-  { fix x assume "bounded_linear f"
+  {
+    fix x
+    assume "bounded_linear f"
     then interpret bounded_linear f .
     from scaleR[of x 1] have "f x = x * f 1"
-      by simp }
+      by simp
+  }
   then show ?thesis
     by (auto intro: exI[of _ "f 1"] bounded_linear_mult_left)
 qed
 
-lemma bij_linear_imp_inv_linear:
-  assumes "linear f" "bij f" shows "linear (inv f)"
-  using assms unfolding linear_def linear_axioms_def additive_def
-  by (auto simp: bij_is_surj bij_is_inj surj_f_inv_f intro!:  Hilbert_Choice.inv_f_eq)
+lemma bij_linear_imp_inv_linear: "linear f \<Longrightarrow> bij f \<Longrightarrow> linear (inv f)"
+  by (auto simp: linear_def linear_axioms_def additive_def bij_is_surj bij_is_inj surj_f_inv_f
+      intro!:  Hilbert_Choice.inv_f_eq)
 
 instance real_normed_algebra_1 \<subseteq> perfect_space
 proof
-  fix x::'a
-  show "\<not> open {x}"
-    unfolding open_dist dist_norm
-    by (clarsimp, rule_tac x="x + of_real (e/2)" in exI, simp)
+  show "\<not> open {x}" for x :: 'a
+    apply (simp only: open_dist dist_norm)
+    apply clarsimp
+    apply (rule_tac x = "x + of_real (e/2)" in exI)
+    apply simp
+    done
 qed
 
+
 subsection \<open>Filters and Limits on Metric Space\<close>
 
 lemma (in metric_space) nhds_metric: "nhds x = (INF e:{0 <..}. principal {y. dist y x < e})"
   unfolding nhds_def
 proof (safe intro!: INF_eq)
-  fix S assume "open S" "x \<in> S"
+  fix S
+  assume "open S" "x \<in> S"
   then obtain e where "{y. dist y x < e} \<subseteq> S" "0 < e"
     by (auto simp: open_dist subset_eq)
   then show "\<exists>e\<in>{0<..}. principal {y. dist y x < e} \<le> principal S"
     by auto
 qed (auto intro!: exI[of _ "{y. dist x y < e}" for e] open_ball simp: dist_commute)
 
-lemma (in metric_space) tendsto_iff:
-  "(f \<longlongrightarrow> l) F \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) F)"
+lemma (in metric_space) tendsto_iff: "(f \<longlongrightarrow> l) F \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) F)"
   unfolding nhds_metric filterlim_INF filterlim_principal by auto
 
-lemma (in metric_space) tendstoI [intro?]: "(\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F) \<Longrightarrow> (f \<longlongrightarrow> l) F"
+lemma (in metric_space) tendstoI [intro?]:
+  "(\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F) \<Longrightarrow> (f \<longlongrightarrow> l) F"
   by (auto simp: tendsto_iff)
 
 lemma (in metric_space) tendstoD: "(f \<longlongrightarrow> l) F \<Longrightarrow> 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F"
@@ -1767,15 +1727,13 @@
   by (subst eventually_INF_base)
      (auto simp: eventually_principal Bex_def subset_eq intro: exI[of _ "min a b" for a b])
 
-lemma eventually_at:
-  fixes a :: "'a :: metric_space"
-  shows "eventually P (at a within S) \<longleftrightarrow> (\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<and> dist x a < d \<longrightarrow> P x)"
-  unfolding eventually_at_filter eventually_nhds_metric by auto
+lemma eventually_at: "eventually P (at a within S) \<longleftrightarrow> (\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<and> dist x a < d \<longrightarrow> P x)"
+  for a :: "'a :: metric_space"
+  by (auto simp: eventually_at_filter eventually_nhds_metric)
 
-lemma eventually_at_le:
-  fixes a :: "'a::metric_space"
-  shows "eventually P (at a within S) \<longleftrightarrow> (\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<and> dist x a \<le> d \<longrightarrow> P x)"
-  unfolding eventually_at_filter eventually_nhds_metric
+lemma eventually_at_le: "eventually P (at a within S) \<longleftrightarrow> (\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<and> dist x a \<le> d \<longrightarrow> P x)"
+  for a :: "'a::metric_space"
+  apply (simp only: eventually_at_filter eventually_nhds_metric)
   apply auto
   apply (rule_tac x="d / 2" in exI)
   apply auto
@@ -1788,19 +1746,21 @@
   by (subst eventually_at, rule exI[of _ "b - a"]) (force simp: dist_real_def)
 
 lemma metric_tendsto_imp_tendsto:
-  fixes a :: "'a :: metric_space" and b :: "'b :: metric_space"
+  fixes a :: "'a :: metric_space"
+    and b :: "'b :: metric_space"
   assumes f: "(f \<longlongrightarrow> a) F"
-  assumes le: "eventually (\<lambda>x. dist (g x) b \<le> dist (f x) a) F"
+    and le: "eventually (\<lambda>x. dist (g x) b \<le> dist (f x) a) F"
   shows "(g \<longlongrightarrow> b) F"
 proof (rule tendstoI)
-  fix e :: real assume "0 < e"
+  fix e :: real
+  assume "0 < e"
   with f have "eventually (\<lambda>x. dist (f x) a < e) F" by (rule tendstoD)
   with le show "eventually (\<lambda>x. dist (g x) b < e) F"
     using le_less_trans by (rule eventually_elim2)
 qed
 
 lemma filterlim_real_sequentially: "LIM x sequentially. real x :> at_top"
-  unfolding filterlim_at_top
+  apply (simp only: filterlim_at_top)
   apply (intro allI)
   apply (rule_tac c="nat \<lceil>Z + 1\<rceil>" in eventually_sequentiallyI)
   apply linarith
@@ -1809,73 +1769,79 @@
 
 subsubsection \<open>Limits of Sequences\<close>
 
-lemma lim_sequentially: "X \<longlonglongrightarrow> (L::'a::metric_space) \<longleftrightarrow> (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. dist (X n) L < r)"
+lemma lim_sequentially: "X \<longlonglongrightarrow> L \<longleftrightarrow> (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. dist (X n) L < r)"
+  for L :: "'a::metric_space"
   unfolding tendsto_iff eventually_sequentially ..
 
 lemmas LIMSEQ_def = lim_sequentially  (*legacy binding*)
 
-lemma LIMSEQ_iff_nz: "X \<longlonglongrightarrow> (L::'a::metric_space) = (\<forall>r>0. \<exists>no>0. \<forall>n\<ge>no. dist (X n) L < r)"
+lemma LIMSEQ_iff_nz: "X \<longlonglongrightarrow> L \<longleftrightarrow> (\<forall>r>0. \<exists>no>0. \<forall>n\<ge>no. dist (X n) L < r)"
+  for L :: "'a::metric_space"
   unfolding lim_sequentially by (metis Suc_leD zero_less_Suc)
 
-lemma metric_LIMSEQ_I:
-  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r) \<Longrightarrow> X \<longlonglongrightarrow> (L::'a::metric_space)"
-by (simp add: lim_sequentially)
+lemma metric_LIMSEQ_I: "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r) \<Longrightarrow> X \<longlonglongrightarrow> L"
+  for L :: "'a::metric_space"
+  by (simp add: lim_sequentially)
 
-lemma metric_LIMSEQ_D:
-  "\<lbrakk>X \<longlonglongrightarrow> (L::'a::metric_space); 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r"
-by (simp add: lim_sequentially)
+lemma metric_LIMSEQ_D: "X \<longlonglongrightarrow> L \<Longrightarrow> 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r"
+  for L :: "'a::metric_space"
+  by (simp add: lim_sequentially)
 
 
 subsubsection \<open>Limits of Functions\<close>
 
-lemma LIM_def: "f \<midarrow>(a::'a::metric_space)\<rightarrow> (L::'b::metric_space) =
-     (\<forall>r > 0. \<exists>s > 0. \<forall>x. x \<noteq> a & dist x a < s
-        --> dist (f x) L < r)"
+lemma LIM_def: "f \<midarrow>a\<rightarrow> L \<longleftrightarrow> (\<forall>r > 0. \<exists>s > 0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r)"
+  for a :: "'a::metric_space" and L :: "'b::metric_space"
   unfolding tendsto_iff eventually_at by simp
 
 lemma metric_LIM_I:
-  "(\<And>r. 0 < r \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r)
-    \<Longrightarrow> f \<midarrow>(a::'a::metric_space)\<rightarrow> (L::'b::metric_space)"
-by (simp add: LIM_def)
+  "(\<And>r. 0 < r \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r) \<Longrightarrow> f \<midarrow>a\<rightarrow> L"
+  for a :: "'a::metric_space" and L :: "'b::metric_space"
+  by (simp add: LIM_def)
 
-lemma metric_LIM_D:
-  "\<lbrakk>f \<midarrow>(a::'a::metric_space)\<rightarrow> (L::'b::metric_space); 0 < r\<rbrakk>
-    \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r"
-by (simp add: LIM_def)
+lemma metric_LIM_D: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> 0 < r \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r"
+  for a :: "'a::metric_space" and L :: "'b::metric_space"
+  by (simp add: LIM_def)
 
 lemma metric_LIM_imp_LIM:
-  assumes f: "f \<midarrow>a\<rightarrow> (l::'a::metric_space)"
-  assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> dist (g x) m \<le> dist (f x) l"
-  shows "g \<midarrow>a\<rightarrow> (m::'b::metric_space)"
+  fixes l :: "'a::metric_space"
+    and m :: "'b::metric_space"
+  assumes f: "f \<midarrow>a\<rightarrow> l"
+    and le: "\<And>x. x \<noteq> a \<Longrightarrow> dist (g x) m \<le> dist (f x) l"
+  shows "g \<midarrow>a\<rightarrow> m"
   by (rule metric_tendsto_imp_tendsto [OF f]) (auto simp add: eventually_at_topological le)
 
 lemma metric_LIM_equal2:
-  assumes 1: "0 < R"
-  assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; dist x a < R\<rbrakk> \<Longrightarrow> f x = g x"
-  shows "g \<midarrow>a\<rightarrow> l \<Longrightarrow> f \<midarrow>(a::'a::metric_space)\<rightarrow> l"
-apply (rule topological_tendstoI)
-apply (drule (2) topological_tendstoD)
-apply (simp add: eventually_at, safe)
-apply (rule_tac x="min d R" in exI, safe)
-apply (simp add: 1)
-apply (simp add: 2)
-done
+  fixes a :: "'a::metric_space"
+  assumes "0 < R"
+    and "\<And>x. x \<noteq> a \<Longrightarrow> dist x a < R \<Longrightarrow> f x = g x"
+  shows "g \<midarrow>a\<rightarrow> l \<Longrightarrow> f \<midarrow>a\<rightarrow> l"
+  apply (rule topological_tendstoI)
+  apply (drule (2) topological_tendstoD)
+  apply (simp add: eventually_at)
+  apply safe
+  apply (rule_tac x="min d R" in exI)
+  apply safe
+   apply (simp add: assms(1))
+  apply (simp add: assms(2))
+  done
 
 lemma metric_LIM_compose2:
-  assumes f: "f \<midarrow>(a::'a::metric_space)\<rightarrow> b"
-  assumes g: "g \<midarrow>b\<rightarrow> c"
-  assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> b"
+  fixes a :: "'a::metric_space"
+  assumes f: "f \<midarrow>a\<rightarrow> b"
+    and g: "g \<midarrow>b\<rightarrow> c"
+    and inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> b"
   shows "(\<lambda>x. g (f x)) \<midarrow>a\<rightarrow> c"
-  using inj
-  by (intro tendsto_compose_eventually[OF g f]) (auto simp: eventually_at)
+  using inj by (intro tendsto_compose_eventually[OF g f]) (auto simp: eventually_at)
 
 lemma metric_isCont_LIM_compose2:
   fixes f :: "'a :: metric_space \<Rightarrow> _"
   assumes f [unfolded isCont_def]: "isCont f a"
-  assumes g: "g \<midarrow>f a\<rightarrow> l"
-  assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> f a"
+    and g: "g \<midarrow>f a\<rightarrow> l"
+    and inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> f a"
   shows "(\<lambda>x. g (f x)) \<midarrow>a\<rightarrow> l"
-by (rule metric_LIM_compose2 [OF f g inj])
+  by (rule metric_LIM_compose2 [OF f g inj])
+
 
 subsection \<open>Complete metric spaces\<close>
 
@@ -1883,12 +1849,14 @@
 
 lemma (in metric_space) Cauchy_def: "Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e)"
 proof -
-  have *: "eventually P (INF M. principal {(X m, X n) | n m. m \<ge> M \<and> n \<ge> M}) =
+  have *: "eventually P (INF M. principal {(X m, X n) | n m. m \<ge> M \<and> n \<ge> M}) \<longleftrightarrow>
     (\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. P (X m, X n))" for P
-  proof (subst eventually_INF_base, goal_cases)
-    case (2 a b) then show ?case
+    apply (subst eventually_INF_base)
+    subgoal by simp
+    subgoal for a b
       by (intro bexI[of _ "max a b"]) (auto simp: eventually_principal subset_eq)
-  qed (auto simp: eventually_principal, blast)
+    subgoal by (auto simp: eventually_principal, blast)
+    done
   have "Cauchy X \<longleftrightarrow> (INF M. principal {(X m, X n) | n m. m \<ge> M \<and> n \<ge> M}) \<le> uniformity"
     unfolding Cauchy_uniform_iff le_filter_def * ..
   also have "\<dots> = (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e)"
@@ -1896,26 +1864,31 @@
   finally show ?thesis .
 qed
 
-lemma (in metric_space) Cauchy_altdef:
-  "Cauchy f = (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n>m. dist (f m) (f n) < e)"
+lemma (in metric_space) Cauchy_altdef: "Cauchy f \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n>m. dist (f m) (f n) < e)"
+  (is "?lhs \<longleftrightarrow> ?rhs")
 proof
-  assume A: "\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n>m. dist (f m) (f n) < e"
-  show "Cauchy f" unfolding Cauchy_def
+  assume ?rhs
+  show ?lhs
+    unfolding Cauchy_def
   proof (intro allI impI)
     fix e :: real assume e: "e > 0"
-    with A obtain M where M: "\<And>m n. m \<ge> M \<Longrightarrow> n > m \<Longrightarrow> dist (f m) (f n) < e" by blast
+    with \<open>?rhs\<close> obtain M where M: "m \<ge> M \<Longrightarrow> n > m \<Longrightarrow> dist (f m) (f n) < e" for m n
+      by blast
     have "dist (f m) (f n) < e" if "m \<ge> M" "n \<ge> M" for m n
       using M[of m n] M[of n m] e that by (cases m n rule: linorder_cases) (auto simp: dist_commute)
-    thus "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (f m) (f n) < e" by blast
+    then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (f m) (f n) < e"
+      by blast
   qed
 next
-  assume "Cauchy f"
-  show "\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n>m. dist (f m) (f n) < e"
+  assume ?lhs
+  show ?rhs
   proof (intro allI impI)
-    fix e :: real assume e: "e > 0"
+    fix e :: real
+    assume e: "e > 0"
     with \<open>Cauchy f\<close> obtain M where "\<And>m n. m \<ge> M \<Longrightarrow> n \<ge> M \<Longrightarrow> dist (f m) (f n) < e"
       unfolding Cauchy_def by blast
-    thus "\<exists>M. \<forall>m\<ge>M. \<forall>n>m. dist (f m) (f n) < e" by (intro exI[of _ M]) force
+    then show "\<exists>M. \<forall>m\<ge>M. \<forall>n>m. dist (f m) (f n) < e"
+      by (intro exI[of _ M]) force
   qed
 qed
 
@@ -1923,7 +1896,8 @@
   "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e) \<Longrightarrow> Cauchy X"
   by (simp add: Cauchy_def)
 
-lemma (in metric_space) CauchyI': "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n>m. dist (X m) (X n) < e) \<Longrightarrow> Cauchy X"
+lemma (in metric_space) CauchyI':
+  "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n>m. dist (X m) (X n) < e) \<Longrightarrow> Cauchy X"
   unfolding Cauchy_altdef by blast
 
 lemma (in metric_space) metric_CauchyD:
@@ -1932,53 +1906,63 @@
 
 lemma (in metric_space) metric_Cauchy_iff2:
   "Cauchy X = (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. dist (X m) (X n) < inverse(real (Suc j))))"
-apply (simp add: Cauchy_def, auto)
-apply (drule reals_Archimedean, safe)
-apply (drule_tac x = n in spec, auto)
-apply (rule_tac x = M in exI, auto)
-apply (drule_tac x = m in spec, simp)
-apply (drule_tac x = na in spec, auto)
-done
+  apply (simp add: Cauchy_def)
+  apply auto
+  apply (drule reals_Archimedean)
+  apply safe
+  apply (drule_tac x = n in spec)
+  apply auto
+  apply (rule_tac x = M in exI)
+  apply auto
+  apply (drule_tac x = m in spec)
+  apply simp
+  apply (drule_tac x = na in spec)
+  apply auto
+  done
 
-lemma Cauchy_iff2:
-  "Cauchy X = (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. \<bar>X m - X n\<bar> < inverse(real (Suc j))))"
-  unfolding metric_Cauchy_iff2 dist_real_def ..
+lemma Cauchy_iff2: "Cauchy X \<longleftrightarrow> (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. \<bar>X m - X n\<bar> < inverse (real (Suc j))))"
+  by (simp only: metric_Cauchy_iff2 dist_real_def)
 
 lemma lim_1_over_n: "((\<lambda>n. 1 / of_nat n) \<longlongrightarrow> (0::'a::real_normed_field)) sequentially"
 proof (subst lim_sequentially, intro allI impI exI)
-  fix e :: real assume e: "e > 0"
-  fix n :: nat assume n: "n \<ge> nat \<lceil>inverse e + 1\<rceil>"
+  fix e :: real
+  assume e: "e > 0"
+  fix n :: nat
+  assume n: "n \<ge> nat \<lceil>inverse e + 1\<rceil>"
   have "inverse e < of_nat (nat \<lceil>inverse e + 1\<rceil>)" by linarith
   also note n
-  finally show "dist (1 / of_nat n :: 'a) 0 < e" using e
-    by (simp add: divide_simps mult.commute norm_divide)
+  finally show "dist (1 / of_nat n :: 'a) 0 < e"
+    using e by (simp add: divide_simps mult.commute norm_divide)
 qed
 
 lemma (in metric_space) complete_def:
   shows "complete S = (\<forall>f. (\<forall>n. f n \<in> S) \<and> Cauchy f \<longrightarrow> (\<exists>l\<in>S. f \<longlonglongrightarrow> l))"
   unfolding complete_uniform
 proof safe
-  fix f :: "nat \<Rightarrow> 'a" assume f: "\<forall>n. f n \<in> S" "Cauchy f"
+  fix f :: "nat \<Rightarrow> 'a"
+  assume f: "\<forall>n. f n \<in> S" "Cauchy f"
     and *: "\<forall>F\<le>principal S. F \<noteq> bot \<longrightarrow> cauchy_filter F \<longrightarrow> (\<exists>x\<in>S. F \<le> nhds x)"
   then show "\<exists>l\<in>S. f \<longlonglongrightarrow> l"
     unfolding filterlim_def using f
     by (intro *[rule_format])
        (auto simp: filtermap_sequentually_ne_bot le_principal eventually_filtermap Cauchy_uniform)
 next
-  fix F :: "'a filter" assume "F \<le> principal S" "F \<noteq> bot" "cauchy_filter F"
+  fix F :: "'a filter"
+  assume "F \<le> principal S" "F \<noteq> bot" "cauchy_filter F"
   assume seq: "\<forall>f. (\<forall>n. f n \<in> S) \<and> Cauchy f \<longrightarrow> (\<exists>l\<in>S. f \<longlonglongrightarrow> l)"
 
-  from \<open>F \<le> principal S\<close> \<open>cauchy_filter F\<close> have FF_le: "F \<times>\<^sub>F F \<le> uniformity_on S"
+  from \<open>F \<le> principal S\<close> \<open>cauchy_filter F\<close>
+  have FF_le: "F \<times>\<^sub>F F \<le> uniformity_on S"
     by (simp add: cauchy_filter_def principal_prod_principal[symmetric] prod_filter_mono)
 
   let ?P = "\<lambda>P e. eventually P F \<and> (\<forall>x. P x \<longrightarrow> x \<in> S) \<and> (\<forall>x y. P x \<longrightarrow> P y \<longrightarrow> dist x y < e)"
-
-  { fix \<epsilon> :: real assume "0 < \<epsilon>"
-    then have "eventually (\<lambda>(x, y). x \<in> S \<and> y \<in> S \<and> dist x y < \<epsilon>) (uniformity_on S)"
-      unfolding eventually_inf_principal eventually_uniformity_metric by auto
-    from filter_leD[OF FF_le this] have "\<exists>P. ?P P \<epsilon>"
-      unfolding eventually_prod_same by auto }
-  note P = this
+  have P: "\<exists>P. ?P P \<epsilon>" if "0 < \<epsilon>" for \<epsilon> :: real
+  proof -
+    from that have "eventually (\<lambda>(x, y). x \<in> S \<and> y \<in> S \<and> dist x y < \<epsilon>) (uniformity_on S)"
+      by (auto simp: eventually_inf_principal eventually_uniformity_metric)
+    from filter_leD[OF FF_le this] show ?thesis
+      by (auto simp: eventually_prod_same)
+  qed
 
   have "\<exists>P. \<forall>n. ?P (P n) (1 / Suc n) \<and> P (Suc n) \<le> P n"
   proof (rule dependent_nat_choice)
@@ -1991,18 +1975,20 @@
     ultimately show "\<exists>Q. ?P Q (1 / Suc (Suc n)) \<and> Q \<le> P"
       by (intro exI[of _ "\<lambda>x. P x \<and> Q x"]) (auto simp: eventually_conj_iff)
   qed
-  then obtain P where P: "\<And>n. eventually (P n) F" "\<And>n x. P n x \<Longrightarrow> x \<in> S"
-    "\<And>n x y. P n x \<Longrightarrow> P n y \<Longrightarrow> dist x y < 1 / Suc n" "\<And>n. P (Suc n) \<le> P n"
+  then obtain P where P: "eventually (P n) F" "P n x \<Longrightarrow> x \<in> S"
+    "P n x \<Longrightarrow> P n y \<Longrightarrow> dist x y < 1 / Suc n" "P (Suc n) \<le> P n"
+    for n x y
     by metis
   have "antimono P"
     using P(4) unfolding decseq_Suc_iff le_fun_def by blast
 
-  obtain X where X: "\<And>n. P n (X n)"
+  obtain X where X: "P n (X n)" for n
     using P(1)[THEN eventually_happens'[OF \<open>F \<noteq> bot\<close>]] by metis
   have "Cauchy X"
     unfolding metric_Cauchy_iff2 inverse_eq_divide
   proof (intro exI allI impI)
-    fix j m n :: nat assume "j \<le> m" "j \<le> n"
+    fix j m n :: nat
+    assume "j \<le> m" "j \<le> n"
     with \<open>antimono P\<close> X have "P j (X m)" "P j (X n)"
       by (auto simp: antimono_def)
     then show "dist (X m) (X n) < 1 / Suc j"
@@ -2015,23 +2001,27 @@
 
   show "\<exists>x\<in>S. F \<le> nhds x"
   proof (rule bexI)
-    { fix e :: real assume "0 < e"
-      then have "(\<lambda>n. 1 / Suc n :: real) \<longlonglongrightarrow> 0 \<and> 0 < e / 2"
+    have "eventually (\<lambda>y. dist y x < e) F" if "0 < e" for e :: real
+    proof -
+      from that have "(\<lambda>n. 1 / Suc n :: real) \<longlonglongrightarrow> 0 \<and> 0 < e / 2"
         by (subst LIMSEQ_Suc_iff) (auto intro!: lim_1_over_n)
       then have "\<forall>\<^sub>F n in sequentially. dist (X n) x < e / 2 \<and> 1 / Suc n < e / 2"
-        using \<open>X \<longlonglongrightarrow> x\<close> unfolding tendsto_iff order_tendsto_iff[where 'a=real] eventually_conj_iff by blast
+        using \<open>X \<longlonglongrightarrow> x\<close>
+        unfolding tendsto_iff order_tendsto_iff[where 'a=real] eventually_conj_iff
+        by blast
       then obtain n where "dist x (X n) < e / 2" "1 / Suc n < e / 2"
         by (auto simp: eventually_sequentially dist_commute)
-      have "eventually (\<lambda>y. dist y x < e) F"
+      show ?thesis
         using \<open>eventually (P n) F\<close>
       proof eventually_elim
-        fix y assume "P n y"
+        case (elim y)
         then have "dist y (X n) < 1 / Suc n"
           by (intro X P)
         also have "\<dots> < e / 2" by fact
         finally show "dist y x < e"
           by (rule dist_triangle_half_l) fact
-      qed }
+      qed
+    qed
     then show "F \<le> nhds x"
       unfolding nhds_metric le_INF_iff le_principal by auto
   qed fact
@@ -2039,7 +2029,7 @@
 
 lemma (in metric_space) totally_bounded_metric:
   "totally_bounded S \<longleftrightarrow> (\<forall>e>0. \<exists>k. finite k \<and> S \<subseteq> (\<Union>x\<in>k. {y. dist x y < e}))"
-  unfolding totally_bounded_def eventually_uniformity_metric imp_ex
+  apply (simp only: totally_bounded_def eventually_uniformity_metric imp_ex)
   apply (subst all_comm)
   apply (intro arg_cong[where f=All] ext)
   apply safe
@@ -2053,45 +2043,47 @@
     done
   done
 
+
 subsubsection \<open>Cauchy Sequences are Convergent\<close>
 
 (* TODO: update to uniform_space *)
 class complete_space = metric_space +
   assumes Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X"
 
-lemma Cauchy_convergent_iff:
-  fixes X :: "nat \<Rightarrow> 'a::complete_space"
-  shows "Cauchy X = convergent X"
-by (blast intro: Cauchy_convergent convergent_Cauchy)
+lemma Cauchy_convergent_iff: "Cauchy X \<longleftrightarrow> convergent X"
+  for X :: "nat \<Rightarrow> 'a::complete_space"
+  by (blast intro: Cauchy_convergent convergent_Cauchy)
+
 
 subsection \<open>The set of real numbers is a complete metric space\<close>
 
 text \<open>
-Proof that Cauchy sequences converge based on the one from
-@{url "http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html"}
+  Proof that Cauchy sequences converge based on the one from
+  @{url "http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html"}
 \<close>
 
 text \<open>
   If sequence @{term "X"} is Cauchy, then its limit is the lub of
   @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
 \<close>
-
 lemma increasing_LIMSEQ:
   fixes f :: "nat \<Rightarrow> real"
   assumes inc: "\<And>n. f n \<le> f (Suc n)"
-      and bdd: "\<And>n. f n \<le> l"
-      and en: "\<And>e. 0 < e \<Longrightarrow> \<exists>n. l \<le> f n + e"
+    and bdd: "\<And>n. f n \<le> l"
+    and en: "\<And>e. 0 < e \<Longrightarrow> \<exists>n. l \<le> f n + e"
   shows "f \<longlonglongrightarrow> l"
 proof (rule increasing_tendsto)
-  fix x assume "x < l"
+  fix x
+  assume "x < l"
   with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x"
     by auto
   from en[OF \<open>0 < e\<close>] obtain n where "l - e \<le> f n"
     by (auto simp: field_simps)
-  with \<open>e < l - x\<close> \<open>0 < e\<close> have "x < f n" by simp
+  with \<open>e < l - x\<close> \<open>0 < e\<close> have "x < f n"
+    by simp
   with incseq_SucI[of f, OF inc] show "eventually (\<lambda>n. x < f n) sequentially"
     by (auto simp: eventually_sequentially incseq_def intro: less_le_trans)
-qed (insert bdd, auto)
+qed (use bdd in auto)
 
 lemma real_Cauchy_convergent:
   fixes X :: "nat \<Rightarrow> real"
@@ -2099,63 +2091,66 @@
   shows "convergent X"
 proof -
   define S :: "real set" where "S = {x. \<exists>N. \<forall>n\<ge>N. x < X n}"
-  then have mem_S: "\<And>N x. \<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S" by auto
+  then have mem_S: "\<And>N x. \<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S"
+    by auto
 
-  { fix N x assume N: "\<forall>n\<ge>N. X n < x"
-  fix y::real assume "y \<in> S"
-  hence "\<exists>M. \<forall>n\<ge>M. y < X n"
-    by (simp add: S_def)
-  then obtain M where "\<forall>n\<ge>M. y < X n" ..
-  hence "y < X (max M N)" by simp
-  also have "\<dots> < x" using N by simp
-  finally have "y \<le> x"
-    by (rule order_less_imp_le) }
-  note bound_isUb = this
+  have bound_isUb: "y \<le> x" if N: "\<forall>n\<ge>N. X n < x" and "y \<in> S" for N and x y :: real
+  proof -
+    from that have "\<exists>M. \<forall>n\<ge>M. y < X n"
+      by (simp add: S_def)
+    then obtain M where "\<forall>n\<ge>M. y < X n" ..
+    then have "y < X (max M N)" by simp
+    also have "\<dots> < x" using N by simp
+    finally show ?thesis by (rule order_less_imp_le)
+  qed
 
   obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m) (X n) < 1"
     using X[THEN metric_CauchyD, OF zero_less_one] by auto
-  hence N: "\<forall>n\<ge>N. dist (X n) (X N) < 1" by simp
+  then have N: "\<forall>n\<ge>N. dist (X n) (X N) < 1" by simp
   have [simp]: "S \<noteq> {}"
   proof (intro exI ex_in_conv[THEN iffD1])
     from N have "\<forall>n\<ge>N. X N - 1 < X n"
       by (simp add: abs_diff_less_iff dist_real_def)
-    thus "X N - 1 \<in> S" by (rule mem_S)
+    then show "X N - 1 \<in> S" by (rule mem_S)
   qed
   have [simp]: "bdd_above S"
   proof
     from N have "\<forall>n\<ge>N. X n < X N + 1"
       by (simp add: abs_diff_less_iff dist_real_def)
-    thus "\<And>s. s \<in> S \<Longrightarrow>  s \<le> X N + 1"
+    then show "\<And>s. s \<in> S \<Longrightarrow>  s \<le> X N + 1"
       by (rule bound_isUb)
   qed
   have "X \<longlonglongrightarrow> Sup S"
   proof (rule metric_LIMSEQ_I)
-  fix r::real assume "0 < r"
-  hence r: "0 < r/2" by simp
-  obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. dist (X n) (X m) < r/2"
-    using metric_CauchyD [OF X r] by auto
-  hence "\<forall>n\<ge>N. dist (X n) (X N) < r/2" by simp
-  hence N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2"
-    by (simp only: dist_real_def abs_diff_less_iff)
+    fix r :: real
+    assume "0 < r"
+    then have r: "0 < r/2" by simp
+    obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. dist (X n) (X m) < r/2"
+      using metric_CauchyD [OF X r] by auto
+    then have "\<forall>n\<ge>N. dist (X n) (X N) < r/2" by simp
+    then have N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2"
+      by (simp only: dist_real_def abs_diff_less_iff)
 
-  from N have "\<forall>n\<ge>N. X N - r/2 < X n" by blast
-  hence "X N - r/2 \<in> S" by (rule mem_S)
-  hence 1: "X N - r/2 \<le> Sup S" by (simp add: cSup_upper)
+    from N have "\<forall>n\<ge>N. X N - r/2 < X n" by blast
+    then have "X N - r/2 \<in> S" by (rule mem_S)
+    then have 1: "X N - r/2 \<le> Sup S" by (simp add: cSup_upper)
 
-  from N have "\<forall>n\<ge>N. X n < X N + r/2" by blast
-  from bound_isUb[OF this]
-  have 2: "Sup S \<le> X N + r/2"
-    by (intro cSup_least) simp_all
+    from N have "\<forall>n\<ge>N. X n < X N + r/2" by blast
+    from bound_isUb[OF this]
+    have 2: "Sup S \<le> X N + r/2"
+      by (intro cSup_least) simp_all
 
-  show "\<exists>N. \<forall>n\<ge>N. dist (X n) (Sup S) < r"
-  proof (intro exI allI impI)
-    fix n assume n: "N \<le> n"
-    from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
-    thus "dist (X n) (Sup S) < r" using 1 2
-      by (simp add: abs_diff_less_iff dist_real_def)
+    show "\<exists>N. \<forall>n\<ge>N. dist (X n) (Sup S) < r"
+    proof (intro exI allI impI)
+      fix n
+      assume n: "N \<le> n"
+      from N n have "X n < X N + r/2" and "X N - r/2 < X n"
+        by simp_all
+      then show "dist (X n) (Sup S) < r" using 1 2
+        by (simp add: abs_diff_less_iff dist_real_def)
+    qed
   qed
-  qed
-  then show ?thesis unfolding convergent_def by auto
+  then show ?thesis by (auto simp: convergent_def)
 qed
 
 instance real :: complete_space
@@ -2170,7 +2165,8 @@
   assumes *: "\<And>X. filterlim X at_top sequentially \<Longrightarrow> (\<lambda>n. f (X n)) \<longlonglongrightarrow> y"
   shows "(f \<longlongrightarrow> y) at_top"
 proof -
-  from nhds_countable[of y] guess A . note A = this
+  obtain A where A: "decseq A" "open (A n)" "y \<in> A n" "nhds y = (INF n. principal (A n))" for n
+    by (rule nhds_countable[of y]) (rule that)
 
   have "\<forall>m. \<exists>k. \<forall>x\<ge>k. f x \<in> A m"
   proof (rule ccontr)
@@ -2181,45 +2177,46 @@
       by (intro dependent_nat_choice) (auto simp del: max.bounded_iff)
     then obtain X where X: "\<And>n. f (X n) \<notin> A m" "\<And>n. max n (X n) + 1 \<le> X (Suc n)"
       by auto
-    { fix n have "1 \<le> n \<longrightarrow> real n \<le> X n"
-        using X[of "n - 1"] by auto }
+    have "1 \<le> n \<Longrightarrow> real n \<le> X n" for n
+      using X[of "n - 1"] by auto
     then have "filterlim X at_top sequentially"
       by (force intro!: filterlim_at_top_mono[OF filterlim_real_sequentially]
-                simp: eventually_sequentially)
+          simp: eventually_sequentially)
     from topological_tendstoD[OF *[OF this] A(2, 3), of m] X(1) show False
       by auto
   qed
-  then obtain k where "\<And>m x. k m \<le> x \<Longrightarrow> f x \<in> A m"
+  then obtain k where "k m \<le> x \<Longrightarrow> f x \<in> A m" for m x
     by metis
   then show ?thesis
-    unfolding at_top_def A
-    by (intro filterlim_base[where i=k]) auto
+    unfolding at_top_def A by (intro filterlim_base[where i=k]) auto
 qed
 
 lemma tendsto_at_topI_sequentially_real:
   fixes f :: "real \<Rightarrow> real"
   assumes mono: "mono f"
-  assumes limseq: "(\<lambda>n. f (real n)) \<longlonglongrightarrow> y"
+    and limseq: "(\<lambda>n. f (real n)) \<longlonglongrightarrow> y"
   shows "(f \<longlongrightarrow> y) at_top"
 proof (rule tendstoI)
-  fix e :: real assume "0 < e"
-  with limseq obtain N :: nat where N: "\<And>n. N \<le> n \<Longrightarrow> \<bar>f (real n) - y\<bar> < e"
+  fix e :: real
+  assume "0 < e"
+  with limseq obtain N :: nat where N: "N \<le> n \<Longrightarrow> \<bar>f (real n) - y\<bar> < e" for n
     by (auto simp: lim_sequentially dist_real_def)
-  { fix x :: real
+  have le: "f x \<le> y" for x :: real
+  proof -
     obtain n where "x \<le> real_of_nat n"
       using real_arch_simple[of x] ..
     note monoD[OF mono this]
     also have "f (real_of_nat n) \<le> y"
       by (rule LIMSEQ_le_const[OF limseq]) (auto intro!: exI[of _ n] monoD[OF mono])
-    finally have "f x \<le> y" . }
-  note le = this
+    finally show ?thesis .
+  qed
   have "eventually (\<lambda>x. real N \<le> x) at_top"
     by (rule eventually_ge_at_top)
   then show "eventually (\<lambda>x. dist (f x) y < e) at_top"
   proof eventually_elim
-    fix x assume N': "real N \<le> x"
+    case (elim x)
     with N[of N] le have "y - f (real N) < e" by auto
-    moreover note monoD[OF mono N']
+    moreover note monoD[OF mono elim]
     ultimately show "dist (f x) y < e"
       using le[of x] by (auto simp: dist_real_def field_simps)
   qed
--- a/src/HOL/Set_Interval.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Set_Interval.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -578,9 +578,9 @@
   obtain y where "Max {a <..} < y"
     using gt_ex by auto
 
-  obtain x where "a < x"
+  obtain x where x: "a < x"
     using gt_ex by auto
-  also then have "x \<le> Max {a <..}"
+  also from x have "x \<le> Max {a <..}"
     by fact
   also note \<open>Max {a <..} < y\<close>
   finally have "y \<le> Max { a <..}"
--- a/src/HOL/Transcendental.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/Transcendental.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -1617,14 +1617,15 @@
 
 lemma isCont_ln:
   fixes x::real assumes "x \<noteq> 0" shows "isCont ln x"
-proof cases
-  assume "0 < x"
-  moreover then have "isCont ln (exp (ln x))"
+proof (cases "0 < x")
+  case True
+  then have "isCont ln (exp (ln x))"
     by (intro isCont_inv_fun[where d="\<bar>x\<bar>" and f=exp]) auto
-  ultimately show ?thesis
+  with True show ?thesis
     by simp
 next
-  assume "\<not> 0 < x" with \<open>x \<noteq> 0\<close> show "isCont ln x"
+  case False
+  with \<open>x \<noteq> 0\<close> show "isCont ln x"
     unfolding isCont_def
     by (subst filterlim_cong[OF _ refl, of _ "nhds (ln 0)" _ "\<lambda>_. ln 0"])
        (auto simp: ln_neg_is_const not_less eventually_at dist_real_def
@@ -4902,11 +4903,11 @@
   have x1: "x \<le> 1"
     using assms
     by (metis le_add_same_cancel1 power2_le_imp_le power_one zero_le_power2)
-  moreover with assms have ax: "0 \<le> arccos x" "cos(arccos x) = x"
+  with assms have ax: "0 \<le> arccos x" "cos (arccos x) = x"
     by (auto simp: arccos)
-  moreover have "y = sqrt (1 - x\<^sup>2)" using assms
+  from assms have "y = sqrt (1 - x\<^sup>2)"
     by (metis abs_of_nonneg add.commute add_diff_cancel real_sqrt_abs)
-  ultimately show ?thesis using assms arccos_le_pi2 [of x]
+  with x1 ax assms arccos_le_pi2 [of x] show ?thesis
     by (rule_tac x="arccos x" in exI) (auto simp: sin_arccos)
 qed
 
@@ -5836,26 +5837,25 @@
     assume ?lhs
     with * have "(\<forall>i\<le>n. c i = (if i=0 then k else 0))"
       by (simp add: polyfun_eq_coeffs [symmetric])
-    then show ?rhs
-      by simp
+    then show ?rhs by simp
   next
-    assume ?rhs then show ?lhs
-      by (induct n) auto
+    assume ?rhs
+    then show ?lhs by (induct n) auto
   qed
 qed
 
 lemma root_polyfun:
-  fixes z:: "'a::idom"
+  fixes z :: "'a::idom"
   assumes "1 \<le> n"
-    shows "z^n = a \<longleftrightarrow> (\<Sum>i\<le>n. (if i = 0 then -a else if i=n then 1 else 0) * z^i) = 0"
+  shows "z^n = a \<longleftrightarrow> (\<Sum>i\<le>n. (if i = 0 then -a else if i=n then 1 else 0) * z^i) = 0"
   using assms
-  by (cases n; simp add: setsum_head_Suc atLeast0AtMost [symmetric])
+  by (cases n) (simp_all add: setsum_head_Suc atLeast0AtMost [symmetric])
 
 lemma
-    fixes zz :: "'a::{idom,real_normed_div_algebra}"
+  fixes zz :: "'a::{idom,real_normed_div_algebra}"
   assumes "1 \<le> n"
-    shows finite_roots_unity: "finite {z::'a. z^n = 1}"
-      and card_roots_unity:   "card {z::'a. z^n = 1} \<le> n"
+  shows finite_roots_unity: "finite {z::'a. z^n = 1}"
+    and card_roots_unity: "card {z::'a. z^n = 1} \<le> n"
   using polyfun_rootbound [of "\<lambda>i. if i = 0 then -1 else if i=n then 1 else 0" n n] assms
   by (auto simp add: root_polyfun [OF assms])
 
--- a/src/HOL/ex/Ballot.thy	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/HOL/ex/Ballot.thy	Mon Jul 25 11:30:31 2016 +0200
@@ -180,10 +180,10 @@
       by auto
     show "(\<lambda>V. V - {?l}) \<in> ?V (\<lambda>V. ?l \<in> V) \<rightarrow> {V \<in> Pow {0..<a + Suc b}. card V = a \<and> ?Q V (a + Suc b)}"
       by (auto simp: Ico_subset_finite *)
-    { fix V assume "V \<subseteq> {0..<?l}"
-      moreover then have "finite V" "?l \<notin> V" "{0..<Suc ?l} \<inter> V = V"
+    { fix V assume V: "V \<subseteq> {0..<?l}"
+      then have "finite V" "?l \<notin> V" "{0..<Suc ?l} \<inter> V = V"
         by (auto dest: finite_subset)
-      ultimately have "card (insert ?l V) = Suc (card V)"
+      with V have "card (insert ?l V) = Suc (card V)"
         "card ({0..<m} \<inter> insert ?l V) = (if m = Suc ?l then Suc (card V) else card ({0..<m} \<inter> V))"
         if "m \<le> Suc ?l" for m
         using that by auto }
--- a/src/Pure/Isar/calculation.ML	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/Pure/Isar/calculation.ML	Mon Jul 25 11:30:31 2016 +0200
@@ -7,7 +7,7 @@
 signature CALCULATION =
 sig
   val print_rules: Proof.context -> unit
-  val get_calculation: Proof.state -> thm list option
+  val check: Proof.state -> thm list option
   val trans_add: attribute
   val trans_del: attribute
   val sym_add: attribute
@@ -28,9 +28,11 @@
 
 (** calculation data **)
 
+type calculation = {result: thm list, level: int, serial: serial, pos: Position.T};
+
 structure Data = Generic_Data
 (
-  type T = (thm Item_Net.T * thm list) * (thm list * int) option;
+  type T = (thm Item_Net.T * thm list) * calculation option;
   val empty = ((Thm.elim_rules, []), NONE);
   val extend = I;
   fun merge (((trans1, sym1), _), ((trans2, sym2), _)) =
@@ -38,6 +40,7 @@
 );
 
 val get_rules = #1 o Data.get o Context.Proof;
+val get_calculation = #2 o Data.get o Context.Proof;
 
 fun print_rules ctxt =
   let
@@ -51,17 +54,43 @@
 
 (* access calculation *)
 
-fun get_calculation state =
-  (case #2 (Data.get (Context.Proof (Proof.context_of state))) of
+fun check_calculation state =
+  (case get_calculation (Proof.context_of state) of
     NONE => NONE
-  | SOME (thms, lev) => if lev = Proof.level state then SOME thms else NONE);
+  | SOME calculation =>
+      if #level calculation = Proof.level state then SOME calculation else NONE);
+
+val check = Option.map #result o check_calculation;
 
 val calculationN = "calculation";
 
-fun put_calculation calc =
-  `Proof.level #-> (fn lev => Proof.map_context (Context.proof_map
-     (Data.map (apsnd (K (Option.map (rpair lev) calc))))))
-  #> Proof.map_context (Proof_Context.put_thms false (calculationN, calc));
+fun update_calculation calc state =
+  let
+    fun report def serial pos =
+      Context_Position.report (Proof.context_of state)
+        (Position.thread_data ())
+          (Markup.entity calculationN ""
+            |> Markup.properties (Position.entity_properties_of def serial pos));
+    val calculation =
+      (case calc of
+        NONE => NONE
+      | SOME result =>
+          (case check_calculation state of
+            NONE =>
+              let
+                val level = Proof.level state;
+                val serial = serial ();
+                val pos = Position.thread_data ();
+                val _ = report true serial pos;
+              in SOME {result = result, level = level, serial = serial, pos = pos} end
+          | SOME {level, serial, pos, ...} =>
+              (report false serial pos;
+                SOME {result = result, level = level, serial = serial, pos = pos})));
+  in
+    state
+    |> (Proof.map_context o Context.proof_map o Data.map o apsnd) (K calculation)
+    |> Proof.map_context (Proof_Context.put_thms false (calculationN, calc))
+  end;
 
 
 
@@ -122,7 +151,9 @@
 
 fun maintain_calculation int final calc state =
   let
-    val state' = put_calculation (SOME calc) state;
+    val state' = state
+      |> update_calculation (SOME calc)
+      |> Proof.improper_reset_facts;
     val ctxt' = Proof.context_of state';
     val _ =
       if int then
@@ -130,7 +161,7 @@
           (Proof_Context.full_name ctxt' (Binding.name calculationN), calc)
         |> Pretty.string_of |> writeln
       else ();
-  in state' |> final ? (put_calculation NONE #> Proof.chain_facts calc) end;
+  in state' |> final ? (update_calculation NONE #> Proof.chain_facts calc) end;
 
 
 (* also and finally *)
@@ -172,7 +203,7 @@
 
     val facts = Proof.the_facts (assert_sane final state);
     val (initial, calculations) =
-      (case get_calculation state of
+      (case check state of
         NONE => (true, Seq.single (Seq.Result facts))
       | SOME calc => (false, combine (calc @ facts)));
 
@@ -195,7 +226,7 @@
   let
     val facts = Proof.the_facts (assert_sane final state);
     val (initial, thms) =
-      (case get_calculation state of
+      (case check state of
         NONE => (true, [])
       | SOME thms => (false, thms));
     val calc = thms @ facts;
--- a/src/Pure/Isar/proof.ML	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/Pure/Isar/proof.ML	Mon Jul 25 11:30:31 2016 +0200
@@ -19,10 +19,12 @@
   val map_context_result : (context -> 'a * context) -> state -> 'a * state
   val map_contexts: (context -> context) -> state -> state
   val propagate_ml_env: state -> state
+  val report_improper: state -> unit
   val the_facts: state -> thm list
   val the_fact: state -> thm
   val set_facts: thm list -> state -> state
   val reset_facts: state -> state
+  val improper_reset_facts: state -> state
   val assert_forward: state -> state
   val assert_chain: state -> state
   val assert_forward_or_chain: state -> state
@@ -163,7 +165,7 @@
 and node =
   Node of
    {context: context,
-    facts: thm list option,
+    facts: (thm list * bool) option,
     mode: mode,
     goal: goal option}
 and goal =
@@ -242,11 +244,14 @@
 
 (* facts *)
 
+fun report_improper state =
+  Context_Position.report (context_of state) (Position.thread_data ()) Markup.improper;
+
 val get_facts = #facts o top;
 
 fun the_facts state =
   (case get_facts state of
-    SOME facts => facts
+    SOME (facts, proper) => (if proper then () else report_improper state; facts)
   | NONE => error "No current facts available");
 
 fun the_fact state =
@@ -256,10 +261,15 @@
 
 fun put_facts index facts =
   map_top (fn (ctxt, _, mode, goal) => (ctxt, facts, mode, goal)) #>
-  map_context (Proof_Context.put_thms index (Auto_Bind.thisN, facts));
+  map_context (Proof_Context.put_thms index (Auto_Bind.thisN, Option.map #1 facts));
+
+fun set_facts thms = put_facts false (SOME (thms, true));
+val reset_facts = put_facts false NONE;
 
-val set_facts = put_facts false o SOME;
-val reset_facts = put_facts false NONE;
+fun improper_reset_facts state =
+  (case get_facts state of
+    SOME (thms, _) => put_facts false (SOME (thms, false)) state
+  | NONE => state);
 
 fun these_factss more_facts (named_factss, state) =
   (named_factss, state |> set_facts (maps snd named_factss @ more_facts));
@@ -267,10 +277,9 @@
 fun export_facts inner outer =
   (case get_facts inner of
     NONE => reset_facts outer
-  | SOME thms =>
-      thms
-      |> Proof_Context.export (context_of inner) (context_of outer)
-      |> (fn ths => put_facts true (SOME ths) outer));
+  | SOME (thms, proper) =>
+      let val thms' = Proof_Context.export (context_of inner) (context_of outer) thms
+      in put_facts true (SOME (thms', proper)) outer end);
 
 
 (* mode *)
@@ -403,8 +412,8 @@
       [Pretty.mark_str (position_markup, "proof"), Pretty.str (" (" ^ mode_name mode ^ ")")]] @
     (if null prt_ctxt then [] else prt_ctxt @ [Pretty.str ""]) @
     (if verbose orelse mode = Forward then
-       pretty_sep (pretty_facts ctxt "" facts) (prt_goal (try find_goal state))
-     else if mode = Chain then pretty_facts ctxt "picking" facts
+       pretty_sep (pretty_facts ctxt "" (Option.map #1 facts)) (prt_goal (try find_goal state))
+     else if mode = Chain then pretty_facts ctxt "picking" (Option.map #1 facts)
      else prt_goal (try find_goal state))
   end;
 
--- a/src/Pure/Isar/proof_context.ML	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/Pure/Isar/proof_context.ML	Mon Jul 25 11:30:31 2016 +0200
@@ -1006,7 +1006,7 @@
           | dups => err (distinct (op =) (map #2 dups)) "Ambiguous specification of literal fact");
 
         val markup =
-          (Markup.entity Markup.literal_factN "")
+          Markup.entity Markup.literal_factN ""
           |> Markup.properties (Position.def_properties_of thm_pos);
         val _ = Context_Position.report_generic context pos markup;
       in pick true ("", thm_pos) [thm] end
--- a/src/Tools/jEdit/src/rendering.scala	Sat Jul 23 13:25:44 2016 +0200
+++ b/src/Tools/jEdit/src/rendering.scala	Mon Jul 25 11:30:31 2016 +0200
@@ -413,40 +413,30 @@
 
   /* caret focus */
 
+  def entity_focus(range: Text.Range,
+    check: (Boolean, Long) => Boolean = (is_def: Boolean, i: Long) => is_def): Set[Long] =
+  {
+    val results =
+      snapshot.cumulate[Set[Long]](range, Set.empty, Rendering.caret_focus_elements, _ =>
+          {
+            case (serials, Text.Info(_, XML.Elem(Markup(Markup.ENTITY, props), _))) =>
+              props match {
+                case Markup.Entity.Def(i) if check(true, i) => Some(serials + i)
+                case Markup.Entity.Ref(i) if check(false, i) => Some(serials + i)
+                case _ => None
+              }
+            case _ => None
+          })
+    (Set.empty[Long] /: results){ case (s1, Text.Info(_, s2)) => s1 ++ s2 }
+  }
+
   def caret_focus(caret_range: Text.Range, visible_range: Text.Range): Set[Long] =
   {
-    def focus(is_def: Boolean): Set[Long] =
-    {
-      val results =
-        snapshot.cumulate[Set[Long]](caret_range, Set.empty, Rendering.caret_focus_elements, _ =>
-            {
-              case (serials, Text.Info(_, XML.Elem(Markup(Markup.ENTITY, props), _))) =>
-                props match {
-                  case Markup.Entity.Def(i) if is_def => Some(serials + i)
-                  case Markup.Entity.Ref(i) if !is_def => Some(serials + i)
-                  case _ => None
-                }
-              case _ => None
-            })
-      (Set.empty[Long] /: results){ case (s1, Text.Info(_, s2)) => s1 ++ s2 }
-    }
-
-    val focus_defs = focus(true)
-
+    val focus_defs = entity_focus(caret_range)
     if (focus_defs.nonEmpty) focus_defs
     else {
-      val focus_refs = focus(false)
-      def visible_def: Boolean =
-        snapshot.cumulate[Boolean](visible_range, false, Rendering.caret_focus_elements, _ =>
-          {
-            case (_, Text.Info(_, XML.Elem(Markup(Markup.ENTITY, props), _))) =>
-              props match {
-                case Markup.Entity.Def(i) if focus_refs(i) => Some(true)
-                case _ => None
-              }
-          }).exists(info => info.info)
-      if (focus_refs.nonEmpty && visible_def) focus_refs
-      else Set.empty[Long]
+      val visible_defs = entity_focus(visible_range)
+      entity_focus(caret_range, (is_def: Boolean, i: Long) => !is_def && visible_defs.contains(i))
     }
   }