misc tuning and modernization;
authorwenzelm
Tue, 01 Nov 2016 00:44:24 +0100
changeset 64435 c93b0e6131c3
parent 64434 af5235830c16
child 64436 254c9411fc48
misc tuning and modernization;
src/HOL/Nonstandard_Analysis/Free_Ultrafilter.thy
src/HOL/Nonstandard_Analysis/HDeriv.thy
src/HOL/Nonstandard_Analysis/HLim.thy
src/HOL/Nonstandard_Analysis/HyperDef.thy
src/HOL/Nonstandard_Analysis/HyperNat.thy
src/HOL/Nonstandard_Analysis/NSA.thy
src/HOL/Nonstandard_Analysis/NSComplex.thy
src/HOL/Nonstandard_Analysis/NatStar.thy
src/HOL/Nonstandard_Analysis/Star.thy
src/HOL/Nonstandard_Analysis/StarDef.thy
src/HOL/Nonstandard_Analysis/document/root.tex
src/HOL/Nonstandard_Analysis/transfer.ML
--- a/src/HOL/Nonstandard_Analysis/Free_Ultrafilter.thy	Mon Oct 31 16:26:36 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/Free_Ultrafilter.thy	Tue Nov 01 00:44:24 2016 +0100
@@ -2,14 +2,15 @@
     Author:     Jacques D. Fleuriot, University of Cambridge
     Author:     Lawrence C Paulson
     Author:     Brian Huffman
-*) 
+*)
 
 section \<open>Filters and Ultrafilters\<close>
 
 theory Free_Ultrafilter
-imports "~~/src/HOL/Library/Infinite_Set"
+  imports "~~/src/HOL/Library/Infinite_Set"
 begin
 
+
 subsection \<open>Definitions and basic properties\<close>
 
 subsubsection \<open>Ultrafilters\<close>
@@ -43,28 +44,29 @@
 
 end
 
+
 subsection \<open>Maximal filter = Ultrafilter\<close>
 
 text \<open>
-   A filter F is an ultrafilter iff it is a maximal filter,
-   i.e. whenever G is a filter and @{term "F \<subseteq> G"} then @{term "F = G"}
+   A filter \<open>F\<close> is an ultrafilter iff it is a maximal filter,
+   i.e. whenever \<open>G\<close> is a filter and @{prop "F \<subseteq> G"} then @{prop "F = G"}
 \<close>
+
 text \<open>
-  Lemmas that shows existence of an extension to what was assumed to
+  Lemma that shows existence of an extension to what was assumed to
   be a maximal filter. Will be used to derive contradiction in proof of
   property of ultrafilter.
 \<close>
 
-lemma extend_filter:
-  "frequently P F \<Longrightarrow> inf F (principal {x. P x}) \<noteq> bot"
-  unfolding trivial_limit_def eventually_inf_principal by (simp add: not_eventually)
+lemma extend_filter: "frequently P F \<Longrightarrow> inf F (principal {x. P x}) \<noteq> bot"
+  by (simp add: trivial_limit_def eventually_inf_principal not_eventually)
 
 lemma max_filter_ultrafilter:
-  assumes proper: "F \<noteq> bot"
+  assumes "F \<noteq> bot"
   assumes max: "\<And>G. G \<noteq> bot \<Longrightarrow> G \<le> F \<Longrightarrow> F = G"
   shows "ultrafilter F"
 proof
-  fix P show "eventually P F \<or> (\<forall>\<^sub>Fx in F. \<not> P x)"
+  show "eventually P F \<or> (\<forall>\<^sub>Fx in F. \<not> P x)" for P
   proof (rule disjCI)
     assume "\<not> (\<forall>\<^sub>Fx in F. \<not> P x)"
     then have "inf F (principal {x. P x}) \<noteq> bot"
@@ -84,7 +86,9 @@
   done
 
 lemma (in ultrafilter) max_filter:
-  assumes G: "G \<noteq> bot" and sub: "G \<le> F" shows "F = G"
+  assumes G: "G \<noteq> bot"
+    and sub: "G \<le> F"
+  shows "F = G"
 proof (rule antisym)
   show "F \<le> G"
     using sub
@@ -92,10 +96,9 @@
              intro!: eventually_frequently G proper)
 qed fact
 
+
 subsection \<open>Ultrafilter Theorem\<close>
 
-text "A local context makes proof of ultrafilter Theorem more modular"
-
 lemma ex_max_ultrafilter:
   fixes F :: "'a filter"
   assumes F: "F \<noteq> bot"
@@ -104,73 +107,77 @@
   let ?X = "{G. G \<noteq> bot \<and> G \<le> F}"
   let ?R = "{(b, a). a \<noteq> bot \<and> a \<le> b \<and> b \<le> F}"
 
-  have bot_notin_R: "\<And>c. c \<in> Chains ?R \<Longrightarrow> bot \<notin> c"
+  have bot_notin_R: "c \<in> Chains ?R \<Longrightarrow> bot \<notin> c" for c
     by (auto simp: Chains_def)
 
   have [simp]: "Field ?R = ?X"
     by (auto simp: Field_def bot_unique)
 
-  have "\<exists>m\<in>Field ?R. \<forall>a\<in>Field ?R. (m, a) \<in> ?R \<longrightarrow> a = m"
+  have "\<exists>m\<in>Field ?R. \<forall>a\<in>Field ?R. (m, a) \<in> ?R \<longrightarrow> a = m" (is "\<exists>m\<in>?A. ?B m")
   proof (rule Zorns_po_lemma)
     show "Partial_order ?R"
-      unfolding partial_order_on_def preorder_on_def
-      by (auto simp: antisym_def refl_on_def trans_def Field_def bot_unique)
+      by (auto simp: partial_order_on_def preorder_on_def
+          antisym_def refl_on_def trans_def Field_def bot_unique)
     show "\<forall>C\<in>Chains ?R. \<exists>u\<in>Field ?R. \<forall>a\<in>C. (a, u) \<in> ?R"
     proof (safe intro!: bexI del: notI)
-      fix c x assume c: "c \<in> Chains ?R"
+      fix c x
+      assume c: "c \<in> Chains ?R"
 
-      { assume "c \<noteq> {}"
-        with c have "Inf c = bot \<longleftrightarrow> (\<exists>x\<in>c. x = bot)"
+      have Inf_c: "Inf c \<noteq> bot" "Inf c \<le> F" if "c \<noteq> {}"
+      proof -
+        from c that have "Inf c = bot \<longleftrightarrow> (\<exists>x\<in>c. x = bot)"
           unfolding trivial_limit_def by (intro eventually_Inf_base) (auto simp: Chains_def)
-        with c have 1: "Inf c \<noteq> bot"
+        with c show "Inf c \<noteq> bot"
           by (simp add: bot_notin_R)
-        from \<open>c \<noteq> {}\<close> obtain x where "x \<in> c" by auto
-        with c have 2: "Inf c \<le> F"
+        from that obtain x where "x \<in> c" by auto
+        with c show "Inf c \<le> F"
           by (auto intro!: Inf_lower2[of x] simp: Chains_def)
-        note 1 2 }
-      note Inf_c = this
+      qed
       then have [simp]: "inf F (Inf c) = (if c = {} then F else Inf c)"
         using c by (auto simp add: inf_absorb2)
 
-      show "inf F (Inf c) \<noteq> bot"
-        using c by (simp add: F Inf_c)
+      from c show "inf F (Inf c) \<noteq> bot"
+        by (simp add: F Inf_c)
+      from c show "inf F (Inf c) \<in> Field ?R"
+        by (simp add: Chains_def Inf_c F)
 
-      show "inf F (Inf c) \<in> Field ?R"
-        using c by (simp add: Chains_def Inf_c F)
-
-      assume x: "x \<in> c"
+      assume "x \<in> c"
       with c show "inf F (Inf c) \<le> x" "x \<le> F"
         by (auto intro: Inf_lower simp: Chains_def)
     qed
   qed
-  then guess U ..
-  then show ?thesis
-    by (intro exI[of _ U] conjI max_filter_ultrafilter) auto
+  then obtain U where U: "U \<in> ?A" "?B U" ..
+  show ?thesis
+  proof
+    from U show "U \<le> F \<and> ultrafilter U"
+      by (auto intro!: max_filter_ultrafilter)
+  qed
 qed
 
+
 subsubsection \<open>Free Ultrafilters\<close>
 
-text \<open>There exists a free ultrafilter on any infinite set\<close>
+text \<open>There exists a free ultrafilter on any infinite set.\<close>
 
 locale freeultrafilter = ultrafilter +
   assumes infinite: "eventually P F \<Longrightarrow> infinite {x. P x}"
 begin
 
 lemma finite: "finite {x. P x} \<Longrightarrow> \<not> eventually P F"
-  by (erule contrapos_pn, erule infinite)
+  by (erule contrapos_pn) (erule infinite)
 
 lemma finite': "finite {x. \<not> P x} \<Longrightarrow> eventually P F"
   by (drule finite) (simp add: not_eventually frequently_eq_eventually)
 
 lemma le_cofinite: "F \<le> cofinite"
   by (intro filter_leI)
-     (auto simp add: eventually_cofinite not_eventually frequently_eq_eventually dest!: finite)
+    (auto simp add: eventually_cofinite not_eventually frequently_eq_eventually dest!: finite)
 
 lemma singleton: "\<not> eventually (\<lambda>x. x = a) F"
-by (rule finite, simp)
+  by (rule finite) simp
 
 lemma singleton': "\<not> eventually (op = a) F"
-by (rule finite, simp)
+  by (rule finite) simp
 
 lemma ultrafilter: "ultrafilter F" ..
 
@@ -186,7 +193,8 @@
   interpret ultrafilter U by fact
   have "freeultrafilter U"
   proof
-    fix P assume "eventually P U"
+    fix P
+    assume "eventually P U"
     with proper have "frequently P U"
       by (rule eventually_frequently)
     then have "frequently P cofinite"
--- a/src/HOL/Nonstandard_Analysis/HDeriv.thy	Mon Oct 31 16:26:36 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/HDeriv.thy	Tue Nov 01 00:44:24 2016 +0100
@@ -4,284 +4,252 @@
     Conversion to Isar and new proofs by Lawrence C Paulson, 2004
 *)
 
-section\<open>Differentiation (Nonstandard)\<close>
+section \<open>Differentiation (Nonstandard)\<close>
 
 theory HDeriv
-imports HLim
+  imports HLim
 begin
 
-text\<open>Nonstandard Definitions\<close>
+text \<open>Nonstandard Definitions.\<close>
 
-definition
-  nsderiv :: "['a::real_normed_field \<Rightarrow> 'a, 'a, 'a] \<Rightarrow> bool"
-          ("(NSDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where
-  "NSDERIV f x :> D = (\<forall>h \<in> Infinitesimal - {0}.
-      (( *f* f)(star_of x + h)
-       - star_of (f x))/h \<approx> star_of D)"
+definition nsderiv :: "['a::real_normed_field \<Rightarrow> 'a, 'a, 'a] \<Rightarrow> bool"
+    ("(NSDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60)
+  where "NSDERIV f x :> D \<longleftrightarrow>
+    (\<forall>h \<in> Infinitesimal - {0}. (( *f* f)(star_of x + h) - star_of (f x)) / h \<approx> star_of D)"
 
-definition
-  NSdifferentiable :: "['a::real_normed_field \<Rightarrow> 'a, 'a] \<Rightarrow> bool"
-    (infixl "NSdifferentiable" 60) where
-  "f NSdifferentiable x = (\<exists>D. NSDERIV f x :> D)"
+definition NSdifferentiable :: "['a::real_normed_field \<Rightarrow> 'a, 'a] \<Rightarrow> bool"
+    (infixl "NSdifferentiable" 60)
+  where "f NSdifferentiable x \<longleftrightarrow> (\<exists>D. NSDERIV f x :> D)"
 
-definition
-  increment :: "[real=>real,real,hypreal] => hypreal" where
-  "increment f x h = (@inc. f NSdifferentiable x &
-           inc = ( *f* f)(hypreal_of_real x + h) - hypreal_of_real (f x))"
+definition increment :: "(real \<Rightarrow> real) \<Rightarrow> real \<Rightarrow> hypreal \<Rightarrow> hypreal"
+  where "increment f x h =
+    (SOME inc. f NSdifferentiable x \<and> inc = ( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x))"
 
 
 subsection \<open>Derivatives\<close>
 
-lemma DERIV_NS_iff:
-      "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D)"
-by (simp add: DERIV_def LIM_NSLIM_iff)
+lemma DERIV_NS_iff: "(DERIV f x :> D) \<longleftrightarrow> (\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D"
+  by (simp add: DERIV_def LIM_NSLIM_iff)
 
-lemma NS_DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D"
-by (simp add: DERIV_def LIM_NSLIM_iff)
+lemma NS_DERIV_D: "DERIV f x :> D \<Longrightarrow> (\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D"
+  by (simp add: DERIV_def LIM_NSLIM_iff)
 
-lemma hnorm_of_hypreal:
-  "\<And>r. hnorm (( *f* of_real) r::'a::real_normed_div_algebra star) = \<bar>r\<bar>"
-by transfer (rule norm_of_real)
+lemma hnorm_of_hypreal: "\<And>r. hnorm (( *f* of_real) r::'a::real_normed_div_algebra star) = \<bar>r\<bar>"
+  by transfer (rule norm_of_real)
 
 lemma Infinitesimal_of_hypreal:
-  "x \<in> Infinitesimal \<Longrightarrow>
-   (( *f* of_real) x::'a::real_normed_div_algebra star) \<in> Infinitesimal"
-apply (rule InfinitesimalI2)
-apply (drule (1) InfinitesimalD2)
-apply (simp add: hnorm_of_hypreal)
-done
+  "x \<in> Infinitesimal \<Longrightarrow> (( *f* of_real) x::'a::real_normed_div_algebra star) \<in> Infinitesimal"
+  apply (rule InfinitesimalI2)
+  apply (drule (1) InfinitesimalD2)
+  apply (simp add: hnorm_of_hypreal)
+  done
 
-lemma of_hypreal_eq_0_iff:
-  "\<And>x. (( *f* of_real) x = (0::'a::real_algebra_1 star)) = (x = 0)"
-by transfer (rule of_real_eq_0_iff)
+lemma of_hypreal_eq_0_iff: "\<And>x. (( *f* of_real) x = (0::'a::real_algebra_1 star)) = (x = 0)"
+  by transfer (rule of_real_eq_0_iff)
 
-lemma NSDeriv_unique:
-     "[| NSDERIV f x :> D; NSDERIV f x :> E |] ==> D = E"
-apply (subgoal_tac "( *f* of_real) \<epsilon> \<in> Infinitesimal - {0::'a star}")
-apply (simp only: nsderiv_def)
-apply (drule (1) bspec)+
-apply (drule (1) approx_trans3)
-apply simp
-apply (simp add: Infinitesimal_of_hypreal Infinitesimal_epsilon)
-apply (simp add: of_hypreal_eq_0_iff hypreal_epsilon_not_zero)
-done
+lemma NSDeriv_unique: "NSDERIV f x :> D \<Longrightarrow> NSDERIV f x :> E \<Longrightarrow> D = E"
+  apply (subgoal_tac "( *f* of_real) \<epsilon> \<in> Infinitesimal - {0::'a star}")
+   apply (simp only: nsderiv_def)
+   apply (drule (1) bspec)+
+   apply (drule (1) approx_trans3)
+   apply simp
+  apply (simp add: Infinitesimal_of_hypreal Infinitesimal_epsilon)
+  apply (simp add: of_hypreal_eq_0_iff hypreal_epsilon_not_zero)
+  done
 
-text \<open>First NSDERIV in terms of NSLIM\<close>
+text \<open>First \<open>NSDERIV\<close> in terms of \<open>NSLIM\<close>.\<close>
 
-text\<open>first equivalence\<close>
-lemma NSDERIV_NSLIM_iff:
-      "(NSDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D)"
-apply (simp add: nsderiv_def NSLIM_def, auto)
-apply (drule_tac x = xa in bspec)
-apply (rule_tac [3] ccontr)
-apply (drule_tac [3] x = h in spec)
-apply (auto simp add: mem_infmal_iff starfun_lambda_cancel)
-done
+text \<open>First equivalence.\<close>
+lemma NSDERIV_NSLIM_iff: "(NSDERIV f x :> D) \<longleftrightarrow> (\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D"
+  apply (auto simp add: nsderiv_def NSLIM_def)
+   apply (drule_tac x = xa in bspec)
+    apply (rule_tac [3] ccontr)
+    apply (drule_tac [3] x = h in spec)
+    apply (auto simp add: mem_infmal_iff starfun_lambda_cancel)
+  done
 
-text\<open>second equivalence\<close>
-lemma NSDERIV_NSLIM_iff2:
-     "(NSDERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S D)"
+text \<open>Second equivalence.\<close>
+lemma NSDERIV_NSLIM_iff2: "(NSDERIV f x :> D) \<longleftrightarrow> (\<lambda>z. (f z - f x) / (z - x)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S D"
   by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff LIM_NSLIM_iff [symmetric])
 
-(* while we're at it! *)
-
+text \<open>While we're at it!\<close>
 lemma NSDERIV_iff2:
-     "(NSDERIV f x :> D) =
-      (\<forall>w.
-        w \<noteq> star_of x & w \<approx> star_of x -->
-        ( *f* (%z. (f z - f x) / (z-x))) w \<approx> star_of D)"
-by (simp add: NSDERIV_NSLIM_iff2 NSLIM_def)
+  "(NSDERIV f x :> D) \<longleftrightarrow>
+    (\<forall>w. w \<noteq> star_of x & w \<approx> star_of x \<longrightarrow> ( *f* (%z. (f z - f x) / (z-x))) w \<approx> star_of D)"
+  by (simp add: NSDERIV_NSLIM_iff2 NSLIM_def)
 
-(*FIXME DELETE*)
-lemma hypreal_not_eq_minus_iff:
-  "(x \<noteq> a) = (x - a \<noteq> (0::'a::ab_group_add))"
-by auto
+(* FIXME delete *)
+lemma hypreal_not_eq_minus_iff: "x \<noteq> a \<longleftrightarrow> x - a \<noteq> (0::'a::ab_group_add)"
+  by auto
 
 lemma NSDERIVD5:
-  "(NSDERIV f x :> D) ==>
-   (\<forall>u. u \<approx> hypreal_of_real x -->
-     ( *f* (%z. f z - f x)) u \<approx> hypreal_of_real D * (u - hypreal_of_real x))"
-apply (auto simp add: NSDERIV_iff2)
-apply (case_tac "u = hypreal_of_real x", auto)
-apply (drule_tac x = u in spec, auto)
-apply (drule_tac c = "u - hypreal_of_real x" and b = "hypreal_of_real D" in approx_mult1)
-apply (drule_tac [!] hypreal_not_eq_minus_iff [THEN iffD1])
-apply (subgoal_tac [2] "( *f* (%z. z-x)) u \<noteq> (0::hypreal) ")
-apply (auto simp add:
-         approx_minus_iff [THEN iffD1, THEN mem_infmal_iff [THEN iffD2]]
-         Infinitesimal_subset_HFinite [THEN subsetD])
-done
+  "(NSDERIV f x :> D) \<Longrightarrow>
+   (\<forall>u. u \<approx> hypreal_of_real x \<longrightarrow>
+     ( *f* (\<lambda>z. f z - f x)) u \<approx> hypreal_of_real D * (u - hypreal_of_real x))"
+  apply (auto simp add: NSDERIV_iff2)
+  apply (case_tac "u = hypreal_of_real x", auto)
+  apply (drule_tac x = u in spec, auto)
+  apply (drule_tac c = "u - hypreal_of_real x" and b = "hypreal_of_real D" in approx_mult1)
+   apply (drule_tac [!] hypreal_not_eq_minus_iff [THEN iffD1])
+   apply (subgoal_tac [2] "( *f* (%z. z-x)) u \<noteq> (0::hypreal) ")
+    apply (auto simp: approx_minus_iff [THEN iffD1, THEN mem_infmal_iff [THEN iffD2]]
+      Infinitesimal_subset_HFinite [THEN subsetD])
+  done
 
 lemma NSDERIVD4:
-     "(NSDERIV f x :> D) ==>
-      (\<forall>h \<in> Infinitesimal.
-               (( *f* f)(hypreal_of_real x + h) -
-                 hypreal_of_real (f x))\<approx> (hypreal_of_real D) * h)"
-apply (auto simp add: nsderiv_def)
-apply (case_tac "h = (0::hypreal) ")
-apply auto
-apply (drule_tac x = h in bspec)
-apply (drule_tac [2] c = h in approx_mult1)
-apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD])
-done
+  "(NSDERIV f x :> D) \<Longrightarrow>
+    (\<forall>h \<in> Infinitesimal.
+      ( *f* f)(hypreal_of_real x + h) - hypreal_of_real (f x) \<approx> hypreal_of_real D * h)"
+  apply (auto simp add: nsderiv_def)
+  apply (case_tac "h = 0")
+   apply auto
+  apply (drule_tac x = h in bspec)
+   apply (drule_tac [2] c = h in approx_mult1)
+    apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD])
+  done
 
 lemma NSDERIVD3:
-     "(NSDERIV f x :> D) ==>
-      (\<forall>h \<in> Infinitesimal - {0}.
-               (( *f* f)(hypreal_of_real x + h) -
-                 hypreal_of_real (f x))\<approx> (hypreal_of_real D) * h)"
-apply (auto simp add: nsderiv_def)
-apply (rule ccontr, drule_tac x = h in bspec)
-apply (drule_tac [2] c = h in approx_mult1)
-apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mult.assoc)
-done
+  "(NSDERIV f x :> D) \<Longrightarrow>
+    \<forall>h \<in> Infinitesimal - {0}.
+      (( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x)) \<approx> hypreal_of_real D * h"
+  apply (auto simp add: nsderiv_def)
+  apply (rule ccontr, drule_tac x = h in bspec)
+   apply (drule_tac [2] c = h in approx_mult1)
+    apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mult.assoc)
+  done
 
-text\<open>Differentiability implies continuity
-         nice and simple "algebraic" proof\<close>
-lemma NSDERIV_isNSCont: "NSDERIV f x :> D ==> isNSCont f x"
-apply (auto simp add: nsderiv_def isNSCont_NSLIM_iff NSLIM_def)
-apply (drule approx_minus_iff [THEN iffD1])
-apply (drule hypreal_not_eq_minus_iff [THEN iffD1])
-apply (drule_tac x = "xa - star_of x" in bspec)
- prefer 2 apply (simp add: add.assoc [symmetric])
-apply (auto simp add: mem_infmal_iff [symmetric] add.commute)
-apply (drule_tac c = "xa - star_of x" in approx_mult1)
-apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
-            simp add: mult.assoc nonzero_mult_div_cancel_right)
-apply (drule_tac x3=D in
-           HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult,
-             THEN mem_infmal_iff [THEN iffD1]])
-apply (auto simp add: mult.commute
-            intro: approx_trans approx_minus_iff [THEN iffD2])
-done
+text \<open>Differentiability implies continuity nice and simple "algebraic" proof.\<close>
+lemma NSDERIV_isNSCont: "NSDERIV f x :> D \<Longrightarrow> isNSCont f x"
+  apply (auto simp add: nsderiv_def isNSCont_NSLIM_iff NSLIM_def)
+  apply (drule approx_minus_iff [THEN iffD1])
+  apply (drule hypreal_not_eq_minus_iff [THEN iffD1])
+  apply (drule_tac x = "xa - star_of x" in bspec)
+   prefer 2 apply (simp add: add.assoc [symmetric])
+   apply (auto simp add: mem_infmal_iff [symmetric] add.commute)
+  apply (drule_tac c = "xa - star_of x" in approx_mult1)
+   apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mult.assoc)
+  apply (drule_tac x3=D in
+      HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult, THEN mem_infmal_iff [THEN iffD1]])
+  apply (auto simp add: mult.commute intro: approx_trans approx_minus_iff [THEN iffD2])
+  done
 
-text\<open>Differentiation rules for combinations of functions
-      follow from clear, straightforard, algebraic
-      manipulations\<close>
-text\<open>Constant function\<close>
+text \<open>Differentiation rules for combinations of functions
+  follow from clear, straightforard, algebraic manipulations.\<close>
+
+text \<open>Constant function.\<close>
 
 (* use simple constant nslimit theorem *)
-lemma NSDERIV_const [simp]: "(NSDERIV (%x. k) x :> 0)"
-by (simp add: NSDERIV_NSLIM_iff)
-
-text\<open>Sum of functions- proved easily\<close>
+lemma NSDERIV_const [simp]: "NSDERIV (\<lambda>x. k) x :> 0"
+  by (simp add: NSDERIV_NSLIM_iff)
 
-lemma NSDERIV_add: "[| NSDERIV f x :> Da;  NSDERIV g x :> Db |]
-      ==> NSDERIV (%x. f x + g x) x :> Da + Db"
-apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def)
-apply (auto simp add: add_divide_distrib diff_divide_distrib dest!: spec)
-apply (drule_tac b = "star_of Da" and d = "star_of Db" in approx_add)
-apply (auto simp add: ac_simps algebra_simps)
-done
-
-text\<open>Product of functions - Proof is trivial but tedious
-  and long due to rearrangement of terms\<close>
+text \<open>Sum of functions- proved easily.\<close>
 
-lemma lemma_nsderiv1:
-  fixes a b c d :: "'a::comm_ring star"
-  shows "(a*b) - (c*d) = (b*(a - c)) + (c*(b - d))"
-by (simp add: right_diff_distrib ac_simps)
+lemma NSDERIV_add:
+  "NSDERIV f x :> Da \<Longrightarrow> NSDERIV g x :> Db \<Longrightarrow> NSDERIV (\<lambda>x. f x + g x) x :> Da + Db"
+  apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def)
+  apply (auto simp add: add_divide_distrib diff_divide_distrib dest!: spec)
+  apply (drule_tac b = "star_of Da" and d = "star_of Db" in approx_add)
+   apply (auto simp add: ac_simps algebra_simps)
+  done
 
-lemma lemma_nsderiv2:
-  fixes x y z :: "'a::real_normed_field star"
-  shows "[| (x - y) / z = star_of D + yb; z \<noteq> 0;
-         z \<in> Infinitesimal; yb \<in> Infinitesimal |]
-      ==> x - y \<approx> 0"
-apply (simp add: nonzero_divide_eq_eq)
-apply (auto intro!: Infinitesimal_HFinite_mult2 HFinite_add
-            simp add: mult.assoc mem_infmal_iff [symmetric])
-apply (erule Infinitesimal_subset_HFinite [THEN subsetD])
-done
+text \<open>Product of functions - Proof is trivial but tedious
+  and long due to rearrangement of terms.\<close>
+
+lemma lemma_nsderiv1: "(a * b) - (c * d) = (b * (a - c)) + (c * (b - d))"
+  for a b c d :: "'a::comm_ring star"
+  by (simp add: right_diff_distrib ac_simps)
 
-lemma NSDERIV_mult: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |]
-      ==> NSDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))"
-apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def)
-apply (auto dest!: spec
-      simp add: starfun_lambda_cancel lemma_nsderiv1)
-apply (simp (no_asm) add: add_divide_distrib diff_divide_distrib)
-apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+
-apply (auto simp add: times_divide_eq_right [symmetric]
-            simp del: times_divide_eq_right times_divide_eq_left)
-apply (drule_tac D = Db in lemma_nsderiv2, assumption+)
-apply (drule_tac
-     approx_minus_iff [THEN iffD2, THEN bex_Infinitesimal_iff2 [THEN iffD2]])
-apply (auto intro!: approx_add_mono1
-            simp add: distrib_right distrib_left mult.commute add.assoc)
-apply (rule_tac b1 = "star_of Db * star_of (f x)"
-         in add.commute [THEN subst])
-apply (auto intro!: Infinitesimal_add_approx_self2 [THEN approx_sym]
-                    Infinitesimal_add Infinitesimal_mult
-                    Infinitesimal_star_of_mult
-                    Infinitesimal_star_of_mult2
-          simp add: add.assoc [symmetric])
-done
+lemma lemma_nsderiv2: "(x - y) / z = star_of D + yb \<Longrightarrow> z \<noteq> 0 \<Longrightarrow>
+  z \<in> Infinitesimal \<Longrightarrow> yb \<in> Infinitesimal \<Longrightarrow> x - y \<approx> 0"
+  for x y z :: "'a::real_normed_field star"
+  apply (simp add: nonzero_divide_eq_eq)
+  apply (auto intro!: Infinitesimal_HFinite_mult2 HFinite_add
+      simp add: mult.assoc mem_infmal_iff [symmetric])
+  apply (erule Infinitesimal_subset_HFinite [THEN subsetD])
+  done
 
-text\<open>Multiplying by a constant\<close>
-lemma NSDERIV_cmult: "NSDERIV f x :> D
-      ==> NSDERIV (%x. c * f x) x :> c*D"
-apply (simp only: times_divide_eq_right [symmetric] NSDERIV_NSLIM_iff
-                  minus_mult_right right_diff_distrib [symmetric])
-apply (erule NSLIM_const [THEN NSLIM_mult])
-done
+lemma NSDERIV_mult:
+  "NSDERIV f x :> Da \<Longrightarrow> NSDERIV g x :> Db \<Longrightarrow>
+    NSDERIV (\<lambda>x. f x * g x) x :> (Da * g x) + (Db * f x)"
+  apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def)
+  apply (auto dest!: spec simp add: starfun_lambda_cancel lemma_nsderiv1)
+  apply (simp (no_asm) add: add_divide_distrib diff_divide_distrib)
+  apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+
+  apply (auto simp add: times_divide_eq_right [symmetric]
+      simp del: times_divide_eq_right times_divide_eq_left)
+  apply (drule_tac D = Db in lemma_nsderiv2, assumption+)
+  apply (drule_tac approx_minus_iff [THEN iffD2, THEN bex_Infinitesimal_iff2 [THEN iffD2]])
+  apply (auto intro!: approx_add_mono1 simp: distrib_right distrib_left mult.commute add.assoc)
+  apply (rule_tac b1 = "star_of Db * star_of (f x)" in add.commute [THEN subst])
+  apply (auto intro!: Infinitesimal_add_approx_self2 [THEN approx_sym]
+      Infinitesimal_add Infinitesimal_mult Infinitesimal_star_of_mult Infinitesimal_star_of_mult2
+      simp add: add.assoc [symmetric])
+  done
 
-text\<open>Negation of function\<close>
-lemma NSDERIV_minus: "NSDERIV f x :> D ==> NSDERIV (%x. -(f x)) x :> -D"
+text \<open>Multiplying by a constant.\<close>
+lemma NSDERIV_cmult: "NSDERIV f x :> D \<Longrightarrow> NSDERIV (\<lambda>x. c * f x) x :> c * D"
+  apply (simp only: times_divide_eq_right [symmetric] NSDERIV_NSLIM_iff
+      minus_mult_right right_diff_distrib [symmetric])
+  apply (erule NSLIM_const [THEN NSLIM_mult])
+  done
+
+text \<open>Negation of function.\<close>
+lemma NSDERIV_minus: "NSDERIV f x :> D \<Longrightarrow> NSDERIV (\<lambda>x. - f x) x :> - D"
 proof (simp add: NSDERIV_NSLIM_iff)
   assume "(\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D"
-  hence deriv: "(\<lambda>h. - ((f(x+h) - f x) / h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S - D"
+  then have deriv: "(\<lambda>h. - ((f(x+h) - f x) / h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S - D"
     by (rule NSLIM_minus)
   have "\<forall>h. - ((f (x + h) - f x) / h) = (- f (x + h) + f x) / h"
     by (simp add: minus_divide_left)
-  with deriv
-  have "(\<lambda>h. (- f (x + h) + f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S - D" by simp
-  then show "(\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D \<Longrightarrow>
-    (\<lambda>h. (f x - f (x + h)) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S - D" by simp
+  with deriv have "(\<lambda>h. (- f (x + h) + f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S - D"
+    by simp
+  then show "(\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D \<Longrightarrow> (\<lambda>h. (f x - f (x + h)) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S - D"
+    by simp
 qed
 
-text\<open>Subtraction\<close>
-lemma NSDERIV_add_minus: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x + -g x) x :> Da + -Db"
-by (blast dest: NSDERIV_add NSDERIV_minus)
+text \<open>Subtraction.\<close>
+lemma NSDERIV_add_minus:
+  "NSDERIV f x :> Da \<Longrightarrow> NSDERIV g x :> Db \<Longrightarrow> NSDERIV (\<lambda>x. f x + - g x) x :> Da + - Db"
+  by (blast dest: NSDERIV_add NSDERIV_minus)
 
 lemma NSDERIV_diff:
-  "NSDERIV f x :> Da \<Longrightarrow> NSDERIV g x :> Db \<Longrightarrow> NSDERIV (\<lambda>x. f x - g x) x :> Da-Db"
+  "NSDERIV f x :> Da \<Longrightarrow> NSDERIV g x :> Db \<Longrightarrow> NSDERIV (\<lambda>x. f x - g x) x :> Da - Db"
   using NSDERIV_add_minus [of f x Da g Db] by simp
 
-text\<open>Similarly to the above, the chain rule admits an entirely
-   straightforward derivation. Compare this with Harrison's
-   HOL proof of the chain rule, which proved to be trickier and
-   required an alternative characterisation of differentiability-
-   the so-called Carathedory derivative. Our main problem is
-   manipulation of terms.\<close>
+text \<open>Similarly to the above, the chain rule admits an entirely
+  straightforward derivation. Compare this with Harrison's
+  HOL proof of the chain rule, which proved to be trickier and
+  required an alternative characterisation of differentiability-
+  the so-called Carathedory derivative. Our main problem is
+  manipulation of terms.\<close>
 
-(* lemmas *)
+
+subsection \<open>Lemmas\<close>
 
 lemma NSDERIV_zero:
-      "[| NSDERIV g x :> D;
-               ( *f* g) (star_of x + xa) = star_of (g x);
-               xa \<in> Infinitesimal;
-               xa \<noteq> 0
-            |] ==> D = 0"
-apply (simp add: nsderiv_def)
-apply (drule bspec, auto)
-done
+  "NSDERIV g x :> D \<Longrightarrow> ( *f* g) (star_of x + xa) = star_of (g x) \<Longrightarrow>
+    xa \<in> Infinitesimal \<Longrightarrow> xa \<noteq> 0 \<Longrightarrow> D = 0"
+  apply (simp add: nsderiv_def)
+  apply (drule bspec)
+   apply auto
+  done
 
-(* can be proved differently using NSLIM_isCont_iff *)
+text \<open>Can be proved differently using \<open>NSLIM_isCont_iff\<close>.\<close>
 lemma NSDERIV_approx:
-     "[| NSDERIV f x :> D;  h \<in> Infinitesimal;  h \<noteq> 0 |]
-      ==> ( *f* f) (star_of x + h) - star_of (f x) \<approx> 0"
-apply (simp add: nsderiv_def)
-apply (simp add: mem_infmal_iff [symmetric])
-apply (rule Infinitesimal_ratio)
-apply (rule_tac [3] approx_star_of_HFinite, auto)
-done
+  "NSDERIV f x :> D \<Longrightarrow> h \<in> Infinitesimal \<Longrightarrow> h \<noteq> 0 \<Longrightarrow>
+    ( *f* f) (star_of x + h) - star_of (f x) \<approx> 0"
+  apply (simp add: nsderiv_def)
+  apply (simp add: mem_infmal_iff [symmetric])
+  apply (rule Infinitesimal_ratio)
+    apply (rule_tac [3] approx_star_of_HFinite, auto)
+  done
 
-(*---------------------------------------------------------------
-   from one version of differentiability
+text \<open>From one version of differentiability
 
-                f(x) - f(a)
-              --------------- \<approx> Db
-                  x - a
- ---------------------------------------------------------------*)
+        \<open>f x - f a\<close>
+      \<open>-------------- \<approx> Db\<close>
+          \<open>x - a\<close>
+\<close>
 
 lemma NSDERIVD1: "[| NSDERIV f (g x) :> Da;
          ( *f* g) (star_of(x) + xa) \<noteq> star_of (g x);
@@ -290,186 +258,185 @@
                    - star_of (f (g x)))
               / (( *f* g) (star_of(x) + xa) - star_of (g x))
              \<approx> star_of(Da)"
-by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def)
+  by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def)
 
-(*--------------------------------------------------------------
-   from other version of differentiability
+text \<open>From other version of differentiability
 
-                f(x + h) - f(x)
-               ----------------- \<approx> Db
-                       h
- --------------------------------------------------------------*)
+      \<open>f (x + h) - f x\<close>
+     \<open>------------------ \<approx> Db\<close>
+             \<open>h\<close>
+\<close>
 
 lemma NSDERIVD2: "[| NSDERIV g x :> Db; xa \<in> Infinitesimal; xa \<noteq> 0 |]
       ==> (( *f* g) (star_of(x) + xa) - star_of(g x)) / xa
           \<approx> star_of(Db)"
-by (auto simp add: NSDERIV_NSLIM_iff NSLIM_def mem_infmal_iff starfun_lambda_cancel)
+  by (auto simp add: NSDERIV_NSLIM_iff NSLIM_def mem_infmal_iff starfun_lambda_cancel)
 
-lemma lemma_chain: "(z::'a::real_normed_field star) \<noteq> 0 ==> x*y = (x*inverse(z))*(z*y)"
+lemma lemma_chain: "z \<noteq> 0 \<Longrightarrow> x * y = (x * inverse z) * (z * y)"
+  for x y z :: "'a::real_normed_field star"
 proof -
   assume z: "z \<noteq> 0"
   have "x * y = x * (inverse z * z) * y" by (simp add: z)
-  thus ?thesis by (simp add: mult.assoc)
+  then show ?thesis by (simp add: mult.assoc)
 qed
 
-text\<open>This proof uses both definitions of differentiability.\<close>
-lemma NSDERIV_chain: "[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |]
-      ==> NSDERIV (f o g) x :> Da * Db"
-apply (simp (no_asm_simp) add: NSDERIV_NSLIM_iff NSLIM_def
-                mem_infmal_iff [symmetric])
-apply clarify
-apply (frule_tac f = g in NSDERIV_approx)
-apply (auto simp add: starfun_lambda_cancel2 starfun_o [symmetric])
-apply (case_tac "( *f* g) (star_of (x) + xa) = star_of (g x) ")
-apply (drule_tac g = g in NSDERIV_zero)
-apply (auto simp add: divide_inverse)
-apply (rule_tac z1 = "( *f* g) (star_of (x) + xa) - star_of (g x) " and y1 = "inverse xa" in lemma_chain [THEN ssubst])
-apply (erule hypreal_not_eq_minus_iff [THEN iffD1])
-apply (rule approx_mult_star_of)
-apply (simp_all add: divide_inverse [symmetric])
-apply (blast intro: NSDERIVD1 approx_minus_iff [THEN iffD2])
-apply (blast intro: NSDERIVD2)
-done
+text \<open>This proof uses both definitions of differentiability.\<close>
+lemma NSDERIV_chain:
+  "NSDERIV f (g x) :> Da \<Longrightarrow> NSDERIV g x :> Db \<Longrightarrow> NSDERIV (f \<circ> g) x :> Da * Db"
+  apply (simp (no_asm_simp) add: NSDERIV_NSLIM_iff NSLIM_def mem_infmal_iff [symmetric])
+  apply clarify
+  apply (frule_tac f = g in NSDERIV_approx)
+    apply (auto simp add: starfun_lambda_cancel2 starfun_o [symmetric])
+  apply (case_tac "( *f* g) (star_of (x) + xa) = star_of (g x) ")
+   apply (drule_tac g = g in NSDERIV_zero)
+      apply (auto simp add: divide_inverse)
+  apply (rule_tac z1 = "( *f* g) (star_of (x) + xa) - star_of (g x) " and y1 = "inverse xa" in lemma_chain [THEN ssubst])
+   apply (erule hypreal_not_eq_minus_iff [THEN iffD1])
+  apply (rule approx_mult_star_of)
+   apply (simp_all add: divide_inverse [symmetric])
+   apply (blast intro: NSDERIVD1 approx_minus_iff [THEN iffD2])
+  apply (blast intro: NSDERIVD2)
+  done
 
-text\<open>Differentiation of natural number powers\<close>
-lemma NSDERIV_Id [simp]: "NSDERIV (%x. x) x :> 1"
-by (simp add: NSDERIV_NSLIM_iff NSLIM_def del: divide_self_if)
+text \<open>Differentiation of natural number powers.\<close>
+lemma NSDERIV_Id [simp]: "NSDERIV (\<lambda>x. x) x :> 1"
+  by (simp add: NSDERIV_NSLIM_iff NSLIM_def del: divide_self_if)
 
 lemma NSDERIV_cmult_Id [simp]: "NSDERIV (op * c) x :> c"
-by (cut_tac c = c and x = x in NSDERIV_Id [THEN NSDERIV_cmult], simp)
+  using NSDERIV_Id [THEN NSDERIV_cmult] by simp
 
 lemma NSDERIV_inverse:
   fixes x :: "'a::real_normed_field"
   assumes "x \<noteq> 0" \<comment> \<open>can't get rid of @{term "x \<noteq> 0"} because it isn't continuous at zero\<close>
   shows "NSDERIV (\<lambda>x. inverse x) x :> - (inverse x ^ Suc (Suc 0))"
 proof -
-  { fix h :: "'a star"
+  {
+    fix h :: "'a star"
     assume h_Inf: "h \<in> Infinitesimal"
-    from this assms have not_0: "star_of x + h \<noteq> 0" by (rule Infinitesimal_add_not_zero)
+    from this assms have not_0: "star_of x + h \<noteq> 0"
+      by (rule Infinitesimal_add_not_zero)
     assume "h \<noteq> 0"
-    from h_Inf have "h * star_of x \<in> Infinitesimal" by (rule Infinitesimal_HFinite_mult) simp
+    from h_Inf have "h * star_of x \<in> Infinitesimal"
+      by (rule Infinitesimal_HFinite_mult) simp
     with assms have "inverse (- (h * star_of x) + - (star_of x * star_of x)) \<approx>
       inverse (- (star_of x * star_of x))"
-      apply - apply (rule inverse_add_Infinitesimal_approx2)
-      apply (auto
-        dest!: hypreal_of_real_HFinite_diff_Infinitesimal
+      apply -
+      apply (rule inverse_add_Infinitesimal_approx2)
+      apply (auto dest!: hypreal_of_real_HFinite_diff_Infinitesimal
         simp add: inverse_minus_eq [symmetric] HFinite_minus_iff)
       done
     moreover from not_0 \<open>h \<noteq> 0\<close> assms
-      have "inverse (- (h * star_of x) + - (star_of x * star_of x)) =
-        (inverse (star_of x + h) - inverse (star_of x)) / h"
+    have "inverse (- (h * star_of x) + - (star_of x * star_of x)) =
+      (inverse (star_of x + h) - inverse (star_of x)) / h"
       apply (simp add: division_ring_inverse_diff nonzero_inverse_mult_distrib [symmetric]
-        nonzero_inverse_minus_eq [symmetric] ac_simps ring_distribs)
+          nonzero_inverse_minus_eq [symmetric] ac_simps ring_distribs)
       apply (subst nonzero_inverse_minus_eq [symmetric])
       using distrib_right [symmetric, of h "star_of x" "star_of x"] apply simp
       apply (simp add: field_simps)
       done
     ultimately have "(inverse (star_of x + h) - inverse (star_of x)) / h \<approx>
       - (inverse (star_of x) * inverse (star_of x))"
-      using assms by simp 
-  } then show ?thesis by (simp add: nsderiv_def)
+      using assms by simp
+  }
+  then show ?thesis by (simp add: nsderiv_def)
 qed
 
+
 subsubsection \<open>Equivalence of NS and Standard definitions\<close>
 
 lemma divideR_eq_divide: "x /\<^sub>R y = x / y"
-by (simp add: divide_inverse mult.commute)
+  by (simp add: divide_inverse mult.commute)
 
-text\<open>Now equivalence between NSDERIV and DERIV\<close>
-lemma NSDERIV_DERIV_iff: "(NSDERIV f x :> D) = (DERIV f x :> D)"
-by (simp add: DERIV_def NSDERIV_NSLIM_iff LIM_NSLIM_iff)
+text \<open>Now equivalence between \<open>NSDERIV\<close> and \<open>DERIV\<close>.\<close>
+lemma NSDERIV_DERIV_iff: "NSDERIV f x :> D \<longleftrightarrow> DERIV f x :> D"
+  by (simp add: DERIV_def NSDERIV_NSLIM_iff LIM_NSLIM_iff)
 
-(* NS version *)
-lemma NSDERIV_pow: "NSDERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))"
-by (simp add: NSDERIV_DERIV_iff DERIV_pow)
+text \<open>NS version.\<close>
+lemma NSDERIV_pow: "NSDERIV (\<lambda>x. x ^ n) x :> real n * (x ^ (n - Suc 0))"
+  by (simp add: NSDERIV_DERIV_iff DERIV_pow)
 
-text\<open>Derivative of inverse\<close>
-
+text \<open>Derivative of inverse.\<close>
 lemma NSDERIV_inverse_fun:
-  fixes x :: "'a::{real_normed_field}"
-  shows "[| NSDERIV f x :> d; f(x) \<noteq> 0 |]
-      ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))"
-by (simp add: NSDERIV_DERIV_iff DERIV_inverse_fun del: power_Suc)
+  "NSDERIV f x :> d \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow>
+    NSDERIV (\<lambda>x. inverse (f x)) x :> (- (d * inverse (f x ^ Suc (Suc 0))))"
+  for x :: "'a::{real_normed_field}"
+  by (simp add: NSDERIV_DERIV_iff DERIV_inverse_fun del: power_Suc)
 
-text\<open>Derivative of quotient\<close>
-
+text \<open>Derivative of quotient.\<close>
 lemma NSDERIV_quotient:
-  fixes x :: "'a::{real_normed_field}"
-  shows "[| NSDERIV f x :> d; NSDERIV g x :> e; g(x) \<noteq> 0 |]
-       ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x)
-                            - (e*f(x))) / (g(x) ^ Suc (Suc 0))"
-by (simp add: NSDERIV_DERIV_iff DERIV_quotient del: power_Suc)
+  fixes x :: "'a::real_normed_field"
+  shows "NSDERIV f x :> d \<Longrightarrow> NSDERIV g x :> e \<Longrightarrow> g x \<noteq> 0 \<Longrightarrow>
+    NSDERIV (\<lambda>y. f y / g y) x :> (d * g x - (e * f x)) / (g x ^ Suc (Suc 0))"
+  by (simp add: NSDERIV_DERIV_iff DERIV_quotient del: power_Suc)
 
-lemma CARAT_NSDERIV: "NSDERIV f x :> l ==>
-      \<exists>g. (\<forall>z. f z - f x = g z * (z-x)) & isNSCont g x & g x = l"
-by (auto simp add: NSDERIV_DERIV_iff isNSCont_isCont_iff CARAT_DERIV
-                   mult.commute)
+lemma CARAT_NSDERIV:
+  "NSDERIV f x :> l \<Longrightarrow> \<exists>g. (\<forall>z. f z - f x = g z * (z - x)) \<and> isNSCont g x \<and> g x = l"
+  by (auto simp add: NSDERIV_DERIV_iff isNSCont_isCont_iff CARAT_DERIV mult.commute)
 
-lemma hypreal_eq_minus_iff3: "(x = y + z) = (x + -z = (y::hypreal))"
-by auto
+lemma hypreal_eq_minus_iff3: "x = y + z \<longleftrightarrow> x + - z = y"
+  for x y z :: hypreal
+  by auto
 
 lemma CARAT_DERIVD:
-  assumes all: "\<forall>z. f z - f x = g z * (z-x)"
-      and nsc: "isNSCont g x"
+  assumes all: "\<forall>z. f z - f x = g z * (z - x)"
+    and nsc: "isNSCont g x"
   shows "NSDERIV f x :> g x"
 proof -
-  from nsc
-  have "\<forall>w. w \<noteq> star_of x \<and> w \<approx> star_of x \<longrightarrow>
-         ( *f* g) w * (w - star_of x) / (w - star_of x) \<approx>
-         star_of (g x)"
-    by (simp add: isNSCont_def nonzero_mult_div_cancel_right)
-  thus ?thesis using all
+  from nsc have "\<forall>w. w \<noteq> star_of x \<and> w \<approx> star_of x \<longrightarrow>
+       ( *f* g) w * (w - star_of x) / (w - star_of x) \<approx> star_of (g x)"
+    by (simp add: isNSCont_def)
+  with all show ?thesis
     by (simp add: NSDERIV_iff2 starfun_if_eq cong: if_cong)
 qed
 
+
 subsubsection \<open>Differentiability predicate\<close>
 
-lemma NSdifferentiableD: "f NSdifferentiable x ==> \<exists>D. NSDERIV f x :> D"
-by (simp add: NSdifferentiable_def)
+lemma NSdifferentiableD: "f NSdifferentiable x \<Longrightarrow> \<exists>D. NSDERIV f x :> D"
+  by (simp add: NSdifferentiable_def)
 
-lemma NSdifferentiableI: "NSDERIV f x :> D ==> f NSdifferentiable x"
-by (force simp add: NSdifferentiable_def)
+lemma NSdifferentiableI: "NSDERIV f x :> D \<Longrightarrow> f NSdifferentiable x"
+  by (force simp add: NSdifferentiable_def)
 
 
 subsection \<open>(NS) Increment\<close>
-lemma incrementI:
-      "f NSdifferentiable x ==>
-      increment f x h = ( *f* f) (hypreal_of_real(x) + h) -
-      hypreal_of_real (f x)"
-by (simp add: increment_def)
 
-lemma incrementI2: "NSDERIV f x :> D ==>
-     increment f x h = ( *f* f) (hypreal_of_real(x) + h) -
-     hypreal_of_real (f x)"
-apply (erule NSdifferentiableI [THEN incrementI])
-done
+lemma incrementI:
+  "f NSdifferentiable x \<Longrightarrow>
+    increment f x h = ( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x)"
+  by (simp add: increment_def)
+
+lemma incrementI2:
+  "NSDERIV f x :> D \<Longrightarrow>
+    increment f x h = ( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x)"
+  by (erule NSdifferentiableI [THEN incrementI])
 
-(* The Increment theorem -- Keisler p. 65 *)
-lemma increment_thm: "[| NSDERIV f x :> D; h \<in> Infinitesimal; h \<noteq> 0 |]
-      ==> \<exists>e \<in> Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h"
-apply (frule_tac h = h in incrementI2, simp add: nsderiv_def)
-apply (drule bspec, auto)
-apply (drule bex_Infinitesimal_iff2 [THEN iffD2], clarify)
-apply (frule_tac b1 = "hypreal_of_real (D) + y"
-        in mult_right_cancel [THEN iffD2])
-apply (erule_tac [2] V = "(( *f* f) (hypreal_of_real (x) + h) - hypreal_of_real (f x)) / h = hypreal_of_real (D) + y" in thin_rl)
-apply assumption
-apply (simp add: times_divide_eq_right [symmetric])
-apply (auto simp add: distrib_right)
-done
+text \<open>The Increment theorem -- Keisler p. 65.\<close>
+lemma increment_thm:
+  "NSDERIV f x :> D \<Longrightarrow> h \<in> Infinitesimal \<Longrightarrow> h \<noteq> 0 \<Longrightarrow>
+    \<exists>e \<in> Infinitesimal. increment f x h = hypreal_of_real D * h + e * h"
+  apply (frule_tac h = h in incrementI2, simp add: nsderiv_def)
+  apply (drule bspec, auto)
+  apply (drule bex_Infinitesimal_iff2 [THEN iffD2], clarify)
+  apply (frule_tac b1 = "hypreal_of_real D + y" in mult_right_cancel [THEN iffD2])
+   apply (erule_tac [2]
+      V = "(( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x)) / h = hypreal_of_real (D) + y"
+      in thin_rl)
+   apply assumption
+  apply (simp add: times_divide_eq_right [symmetric])
+  apply (auto simp add: distrib_right)
+  done
 
 lemma increment_thm2:
-     "[| NSDERIV f x :> D; h \<approx> 0; h \<noteq> 0 |]
-      ==> \<exists>e \<in> Infinitesimal. increment f x h =
-              hypreal_of_real(D)*h + e*h"
-by (blast dest!: mem_infmal_iff [THEN iffD2] intro!: increment_thm)
-
+  "NSDERIV f x :> D \<Longrightarrow> h \<approx> 0 \<Longrightarrow> h \<noteq> 0 \<Longrightarrow>
+    \<exists>e \<in> Infinitesimal. increment f x h = hypreal_of_real D * h + e * h"
+  by (blast dest!: mem_infmal_iff [THEN iffD2] intro!: increment_thm)
 
-lemma increment_approx_zero: "[| NSDERIV f x :> D; h \<approx> 0; h \<noteq> 0 |]
-      ==> increment f x h \<approx> 0"
-apply (drule increment_thm2,
-       auto intro!: Infinitesimal_HFinite_mult2 HFinite_add simp add: distrib_right [symmetric] mem_infmal_iff [symmetric])
-apply (erule Infinitesimal_subset_HFinite [THEN subsetD])
-done
+lemma increment_approx_zero: "NSDERIV f x :> D \<Longrightarrow> h \<approx> 0 \<Longrightarrow> h \<noteq> 0 \<Longrightarrow> increment f x h \<approx> 0"
+  apply (drule increment_thm2)
+    apply (auto intro!: Infinitesimal_HFinite_mult2 HFinite_add
+      simp add: distrib_right [symmetric] mem_infmal_iff [symmetric])
+  apply (erule Infinitesimal_subset_HFinite [THEN subsetD])
+  done
 
 end
--- a/src/HOL/Nonstandard_Analysis/HLim.thy	Mon Oct 31 16:26:36 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/HLim.thy	Tue Nov 01 00:44:24 2016 +0100
@@ -3,330 +3,308 @@
     Author:     Lawrence C Paulson
 *)
 
-section\<open>Limits and Continuity (Nonstandard)\<close>
+section \<open>Limits and Continuity (Nonstandard)\<close>
 
 theory HLim
   imports Star
   abbrevs "--->" = "\<midarrow>\<rightarrow>\<^sub>N\<^sub>S"
 begin
 
-text\<open>Nonstandard Definitions\<close>
+text \<open>Nonstandard Definitions.\<close>
 
-definition
-  NSLIM :: "['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool"
-            ("((_)/ \<midarrow>(_)/\<rightarrow>\<^sub>N\<^sub>S (_))" [60, 0, 60] 60) where
-  "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L =
-    (\<forall>x. (x \<noteq> star_of a & x \<approx> star_of a --> ( *f* f) x \<approx> star_of L))"
+definition NSLIM :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
+    ("((_)/ \<midarrow>(_)/\<rightarrow>\<^sub>N\<^sub>S (_))" [60, 0, 60] 60)
+  where "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<forall>x. x \<noteq> star_of a \<and> x \<approx> star_of a \<longrightarrow> ( *f* f) x \<approx> star_of L)"
 
-definition
-  isNSCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool" where
-    \<comment>\<open>NS definition dispenses with limit notions\<close>
-  "isNSCont f a = (\<forall>y. y \<approx> star_of a -->
-         ( *f* f) y \<approx> star_of (f a))"
+definition isNSCont :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool"
+  where  \<comment> \<open>NS definition dispenses with limit notions\<close>
+    "isNSCont f a \<longleftrightarrow> (\<forall>y. y \<approx> star_of a \<longrightarrow> ( *f* f) y \<approx> star_of (f a))"
 
-definition
-  isNSUCont :: "['a::real_normed_vector => 'b::real_normed_vector] => bool" where
-  "isNSUCont f = (\<forall>x y. x \<approx> y --> ( *f* f) x \<approx> ( *f* f) y)"
+definition isNSUCont :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> bool"
+  where "isNSUCont f \<longleftrightarrow> (\<forall>x y. x \<approx> y \<longrightarrow> ( *f* f) x \<approx> ( *f* f) y)"
 
 
 subsection \<open>Limits of Functions\<close>
 
-lemma NSLIM_I:
-  "(\<And>x. \<lbrakk>x \<noteq> star_of a; x \<approx> star_of a\<rbrakk> \<Longrightarrow> starfun f x \<approx> star_of L)
-   \<Longrightarrow> f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L"
-by (simp add: NSLIM_def)
+lemma NSLIM_I: "(\<And>x. x \<noteq> star_of a \<Longrightarrow> x \<approx> star_of a \<Longrightarrow> starfun f x \<approx> star_of L) \<Longrightarrow> f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L"
+  by (simp add: NSLIM_def)
+
+lemma NSLIM_D: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> x \<noteq> star_of a \<Longrightarrow> x \<approx> star_of a \<Longrightarrow> starfun f x \<approx> star_of L"
+  by (simp add: NSLIM_def)
 
-lemma NSLIM_D:
-  "\<lbrakk>f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L; x \<noteq> star_of a; x \<approx> star_of a\<rbrakk>
-   \<Longrightarrow> starfun f x \<approx> star_of L"
-by (simp add: NSLIM_def)
+text \<open>Proving properties of limits using nonstandard definition.
+  The properties hold for standard limits as well!\<close>
 
-text\<open>Proving properties of limits using nonstandard definition.
-      The properties hold for standard limits as well!\<close>
-
-lemma NSLIM_mult:
-  fixes l m :: "'a::real_normed_algebra"
-  shows "[| f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l; g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m |]
-      ==> (%x. f(x) * g(x)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l * m)"
-by (auto simp add: NSLIM_def intro!: approx_mult_HFinite)
+lemma NSLIM_mult: "f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m \<Longrightarrow> (\<lambda>x. f x * g x) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l * m)"
+  for l m :: "'a::real_normed_algebra"
+  by (auto simp add: NSLIM_def intro!: approx_mult_HFinite)
 
-lemma starfun_scaleR [simp]:
-  "starfun (\<lambda>x. f x *\<^sub>R g x) = (\<lambda>x. scaleHR (starfun f x) (starfun g x))"
-by transfer (rule refl)
+lemma starfun_scaleR [simp]: "starfun (\<lambda>x. f x *\<^sub>R g x) = (\<lambda>x. scaleHR (starfun f x) (starfun g x))"
+  by transfer (rule refl)
 
-lemma NSLIM_scaleR:
-  "[| f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l; g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m |]
-      ==> (%x. f(x) *\<^sub>R g(x)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l *\<^sub>R m)"
-by (auto simp add: NSLIM_def intro!: approx_scaleR_HFinite)
+lemma NSLIM_scaleR: "f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m \<Longrightarrow> (\<lambda>x. f x *\<^sub>R g x) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l *\<^sub>R m)"
+  by (auto simp add: NSLIM_def intro!: approx_scaleR_HFinite)
 
-lemma NSLIM_add:
-     "[| f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l; g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m |]
-      ==> (%x. f(x) + g(x)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l + m)"
-by (auto simp add: NSLIM_def intro!: approx_add)
+lemma NSLIM_add: "f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m \<Longrightarrow> (\<lambda>x. f x + g x) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l + m)"
+  by (auto simp add: NSLIM_def intro!: approx_add)
 
-lemma NSLIM_const [simp]: "(%x. k) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S k"
-by (simp add: NSLIM_def)
+lemma NSLIM_const [simp]: "(\<lambda>x. k) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S k"
+  by (simp add: NSLIM_def)
 
-lemma NSLIM_minus: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L ==> (%x. -f(x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S -L"
-by (simp add: NSLIM_def)
+lemma NSLIM_minus: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> (\<lambda>x. - f x) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S -L"
+  by (simp add: NSLIM_def)
 
-lemma NSLIM_diff:
-  "\<lbrakk>f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l; g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m\<rbrakk> \<Longrightarrow> (\<lambda>x. f x - g x) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l - m)"
+lemma NSLIM_diff: "f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m \<Longrightarrow> (\<lambda>x. f x - g x) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l - m)"
   by (simp only: NSLIM_add NSLIM_minus diff_conv_add_uminus)
 
-lemma NSLIM_add_minus: "[| f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l; g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m |] ==> (%x. f(x) + -g(x)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l + -m)"
-by (simp only: NSLIM_add NSLIM_minus)
+lemma NSLIM_add_minus: "f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m \<Longrightarrow> (\<lambda>x. f x + - g x) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l + -m)"
+  by (simp only: NSLIM_add NSLIM_minus)
 
-lemma NSLIM_inverse:
-  fixes L :: "'a::real_normed_div_algebra"
-  shows "[| f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L;  L \<noteq> 0 |]
-      ==> (%x. inverse(f(x))) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S (inverse L)"
-apply (simp add: NSLIM_def, clarify)
-apply (drule spec)
-apply (auto simp add: star_of_approx_inverse)
-done
+lemma NSLIM_inverse: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> L \<noteq> 0 \<Longrightarrow> (\<lambda>x. inverse (f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S (inverse L)"
+  for L :: "'a::real_normed_div_algebra"
+  apply (simp add: NSLIM_def, clarify)
+  apply (drule spec)
+  apply (auto simp add: star_of_approx_inverse)
+  done
 
 lemma NSLIM_zero:
-  assumes f: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S l" shows "(%x. f(x) - l) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S 0"
+  assumes f: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S l"
+  shows "(\<lambda>x. f(x) - l) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S 0"
 proof -
   have "(\<lambda>x. f x - l) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S l - l"
     by (rule NSLIM_diff [OF f NSLIM_const])
-  thus ?thesis by simp
+  then show ?thesis by simp
 qed
 
-lemma NSLIM_zero_cancel: "(%x. f(x) - l) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0 ==> f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l"
-apply (drule_tac g = "%x. l" and m = l in NSLIM_add)
-apply (auto simp add: add.assoc)
-done
+lemma NSLIM_zero_cancel: "(\<lambda>x. f x - l) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0 \<Longrightarrow> f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l"
+  apply (drule_tac g = "%x. l" and m = l in NSLIM_add)
+   apply (auto simp add: add.assoc)
+  done
 
-lemma NSLIM_const_not_eq:
-  fixes a :: "'a::real_normed_algebra_1"
-  shows "k \<noteq> L \<Longrightarrow> \<not> (\<lambda>x. k) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L"
-apply (simp add: NSLIM_def)
-apply (rule_tac x="star_of a + of_hypreal \<epsilon>" in exI)
-apply (simp add: hypreal_epsilon_not_zero approx_def)
-done
+lemma NSLIM_const_not_eq: "k \<noteq> L \<Longrightarrow> \<not> (\<lambda>x. k) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L"
+  for a :: "'a::real_normed_algebra_1"
+  apply (simp add: NSLIM_def)
+  apply (rule_tac x="star_of a + of_hypreal \<epsilon>" in exI)
+  apply (simp add: hypreal_epsilon_not_zero approx_def)
+  done
 
-lemma NSLIM_not_zero:
-  fixes a :: "'a::real_normed_algebra_1"
-  shows "k \<noteq> 0 \<Longrightarrow> \<not> (\<lambda>x. k) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S 0"
-by (rule NSLIM_const_not_eq)
+lemma NSLIM_not_zero: "k \<noteq> 0 \<Longrightarrow> \<not> (\<lambda>x. k) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S 0"
+  for a :: "'a::real_normed_algebra_1"
+  by (rule NSLIM_const_not_eq)
 
-lemma NSLIM_const_eq:
-  fixes a :: "'a::real_normed_algebra_1"
-  shows "(\<lambda>x. k) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> k = L"
-apply (rule ccontr)
-apply (blast dest: NSLIM_const_not_eq)
-done
+lemma NSLIM_const_eq: "(\<lambda>x. k) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> k = L"
+  for a :: "'a::real_normed_algebra_1"
+  by (rule ccontr) (blast dest: NSLIM_const_not_eq)
+
+lemma NSLIM_unique: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S M \<Longrightarrow> L = M"
+  for a :: "'a::real_normed_algebra_1"
+  by (drule (1) NSLIM_diff) (auto dest!: NSLIM_const_eq)
 
-lemma NSLIM_unique:
-  fixes a :: "'a::real_normed_algebra_1"
-  shows "\<lbrakk>f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L; f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S M\<rbrakk> \<Longrightarrow> L = M"
-apply (drule (1) NSLIM_diff)
-apply (auto dest!: NSLIM_const_eq)
-done
+lemma NSLIM_mult_zero: "f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0 \<Longrightarrow> g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0 \<Longrightarrow> (\<lambda>x. f x * g x) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0"
+  for f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra"
+  by (drule NSLIM_mult) auto
 
-lemma NSLIM_mult_zero:
-  fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra"
-  shows "[| f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0; g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0 |] ==> (%x. f(x)*g(x)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0"
-by (drule NSLIM_mult, auto)
+lemma NSLIM_self: "(\<lambda>x. x) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S a"
+  by (simp add: NSLIM_def)
 
-lemma NSLIM_self: "(%x. x) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S a"
-by (simp add: NSLIM_def)
 
 subsubsection \<open>Equivalence of @{term filterlim} and @{term NSLIM}\<close>
 
 lemma LIM_NSLIM:
-  assumes f: "f \<midarrow>a\<rightarrow> L" shows "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L"
+  assumes f: "f \<midarrow>a\<rightarrow> L"
+  shows "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L"
 proof (rule NSLIM_I)
   fix x
   assume neq: "x \<noteq> star_of a"
   assume approx: "x \<approx> star_of a"
   have "starfun f x - star_of L \<in> Infinitesimal"
   proof (rule InfinitesimalI2)
-    fix r::real assume r: "0 < r"
-    from LIM_D [OF f r]
-    obtain s where s: "0 < s" and
-      less_r: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < s\<rbrakk> \<Longrightarrow> norm (f x - L) < r"
+    fix r :: real
+    assume r: "0 < r"
+    from LIM_D [OF f r] obtain s
+      where s: "0 < s" and less_r: "\<And>x. x \<noteq> a \<Longrightarrow> norm (x - a) < s \<Longrightarrow> norm (f x - L) < r"
       by fast
     from less_r have less_r':
-       "\<And>x. \<lbrakk>x \<noteq> star_of a; hnorm (x - star_of a) < star_of s\<rbrakk>
-        \<Longrightarrow> hnorm (starfun f x - star_of L) < star_of r"
+      "\<And>x. x \<noteq> star_of a \<Longrightarrow> hnorm (x - star_of a) < star_of s \<Longrightarrow>
+        hnorm (starfun f x - star_of L) < star_of r"
       by transfer
     from approx have "x - star_of a \<in> Infinitesimal"
-      by (unfold approx_def)
-    hence "hnorm (x - star_of a) < star_of s"
+      by (simp only: approx_def)
+    then have "hnorm (x - star_of a) < star_of s"
       using s by (rule InfinitesimalD2)
     with neq show "hnorm (starfun f x - star_of L) < star_of r"
       by (rule less_r')
   qed
-  thus "starfun f x \<approx> star_of L"
+  then show "starfun f x \<approx> star_of L"
     by (unfold approx_def)
 qed
 
 lemma NSLIM_LIM:
-  assumes f: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L" shows "f \<midarrow>a\<rightarrow> L"
+  assumes f: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L"
+  shows "f \<midarrow>a\<rightarrow> L"
 proof (rule LIM_I)
-  fix r::real assume r: "0 < r"
-  have "\<exists>s>0. \<forall>x. x \<noteq> star_of a \<and> hnorm (x - star_of a) < s
-        \<longrightarrow> hnorm (starfun f x - star_of L) < star_of r"
+  fix r :: real
+  assume r: "0 < r"
+  have "\<exists>s>0. \<forall>x. x \<noteq> star_of a \<and> hnorm (x - star_of a) < s \<longrightarrow>
+    hnorm (starfun f x - star_of L) < star_of r"
   proof (rule exI, safe)
-    show "0 < \<epsilon>" by (rule hypreal_epsilon_gt_zero)
+    show "0 < \<epsilon>"
+      by (rule hypreal_epsilon_gt_zero)
   next
-    fix x assume neq: "x \<noteq> star_of a"
+    fix x
+    assume neq: "x \<noteq> star_of a"
     assume "hnorm (x - star_of a) < \<epsilon>"
-    with Infinitesimal_epsilon
-    have "x - star_of a \<in> Infinitesimal"
+    with Infinitesimal_epsilon have "x - star_of a \<in> Infinitesimal"
       by (rule hnorm_less_Infinitesimal)
-    hence "x \<approx> star_of a"
+    then have "x \<approx> star_of a"
       by (unfold approx_def)
     with f neq have "starfun f x \<approx> star_of L"
       by (rule NSLIM_D)
-    hence "starfun f x - star_of L \<in> Infinitesimal"
+    then have "starfun f x - star_of L \<in> Infinitesimal"
       by (unfold approx_def)
-    thus "hnorm (starfun f x - star_of L) < star_of r"
+    then show "hnorm (starfun f x - star_of L) < star_of r"
       using r by (rule InfinitesimalD2)
   qed
-  thus "\<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - L) < r"
+  then show "\<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - L) < r"
     by transfer
 qed
 
-theorem LIM_NSLIM_iff: "(f \<midarrow>x\<rightarrow> L) = (f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S L)"
-by (blast intro: LIM_NSLIM NSLIM_LIM)
+theorem LIM_NSLIM_iff: "f \<midarrow>x\<rightarrow> L \<longleftrightarrow> f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S L"
+  by (blast intro: LIM_NSLIM NSLIM_LIM)
 
 
 subsection \<open>Continuity\<close>
 
-lemma isNSContD:
-  "\<lbrakk>isNSCont f a; y \<approx> star_of a\<rbrakk> \<Longrightarrow> ( *f* f) y \<approx> star_of (f a)"
-by (simp add: isNSCont_def)
+lemma isNSContD: "isNSCont f a \<Longrightarrow> y \<approx> star_of a \<Longrightarrow> ( *f* f) y \<approx> star_of (f a)"
+  by (simp add: isNSCont_def)
+
+lemma isNSCont_NSLIM: "isNSCont f a \<Longrightarrow> f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S (f a)"
+  by (simp add: isNSCont_def NSLIM_def)
 
-lemma isNSCont_NSLIM: "isNSCont f a ==> f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S (f a) "
-by (simp add: isNSCont_def NSLIM_def)
+lemma NSLIM_isNSCont: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S (f a) \<Longrightarrow> isNSCont f a"
+  apply (auto simp add: isNSCont_def NSLIM_def)
+  apply (case_tac "y = star_of a")
+   apply auto
+  done
 
-lemma NSLIM_isNSCont: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S (f a) ==> isNSCont f a"
-apply (simp add: isNSCont_def NSLIM_def, auto)
-apply (case_tac "y = star_of a", auto)
-done
+text \<open>NS continuity can be defined using NS Limit in
+  similar fashion to standard definition of continuity.\<close>
+lemma isNSCont_NSLIM_iff: "isNSCont f a \<longleftrightarrow> f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S (f a)"
+  by (blast intro: isNSCont_NSLIM NSLIM_isNSCont)
 
-text\<open>NS continuity can be defined using NS Limit in
-    similar fashion to standard definition of continuity\<close>
-lemma isNSCont_NSLIM_iff: "(isNSCont f a) = (f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S (f a))"
-by (blast intro: isNSCont_NSLIM NSLIM_isNSCont)
+text \<open>Hence, NS continuity can be given in terms of standard limit.\<close>
+lemma isNSCont_LIM_iff: "(isNSCont f a) = (f \<midarrow>a\<rightarrow> (f a))"
+  by (simp add: LIM_NSLIM_iff isNSCont_NSLIM_iff)
 
-text\<open>Hence, NS continuity can be given
-  in terms of standard limit\<close>
-lemma isNSCont_LIM_iff: "(isNSCont f a) = (f \<midarrow>a\<rightarrow> (f a))"
-by (simp add: LIM_NSLIM_iff isNSCont_NSLIM_iff)
+text \<open>Moreover, it's trivial now that NS continuity
+  is equivalent to standard continuity.\<close>
+lemma isNSCont_isCont_iff: "isNSCont f a \<longleftrightarrow> isCont f a"
+  by (simp add: isCont_def) (rule isNSCont_LIM_iff)
 
-text\<open>Moreover, it's trivial now that NS continuity
-  is equivalent to standard continuity\<close>
-lemma isNSCont_isCont_iff: "(isNSCont f a) = (isCont f a)"
-apply (simp add: isCont_def)
-apply (rule isNSCont_LIM_iff)
-done
+text \<open>Standard continuity \<open>\<Longrightarrow>\<close> NS continuity.\<close>
+lemma isCont_isNSCont: "isCont f a \<Longrightarrow> isNSCont f a"
+  by (erule isNSCont_isCont_iff [THEN iffD2])
 
-text\<open>Standard continuity ==> NS continuity\<close>
-lemma isCont_isNSCont: "isCont f a ==> isNSCont f a"
-by (erule isNSCont_isCont_iff [THEN iffD2])
+text \<open>NS continuity \<open>==>\<close> Standard continuity.\<close>
+lemma isNSCont_isCont: "isNSCont f a ==> isCont f a"
+  by (erule isNSCont_isCont_iff [THEN iffD1])
 
-text\<open>NS continuity ==> Standard continuity\<close>
-lemma isNSCont_isCont: "isNSCont f a ==> isCont f a"
-by (erule isNSCont_isCont_iff [THEN iffD1])
+
+text \<open>Alternative definition of continuity.\<close>
 
-text\<open>Alternative definition of continuity\<close>
+text \<open>Prove equivalence between NS limits --
+  seems easier than using standard definition.\<close>
+lemma NSLIM_h_iff: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S L"
+  apply (simp add: NSLIM_def, auto)
+   apply (drule_tac x = "star_of a + x" in spec)
+   apply (drule_tac [2] x = "- star_of a + x" in spec, safe, simp)
+      apply (erule mem_infmal_iff [THEN iffD2, THEN Infinitesimal_add_approx_self [THEN approx_sym]])
+     apply (erule_tac [3] approx_minus_iff2 [THEN iffD1])
+    prefer 2 apply (simp add: add.commute)
+   apply (rule_tac x = x in star_cases)
+   apply (rule_tac [2] x = x in star_cases)
+   apply (auto simp add: starfun star_of_def star_n_minus star_n_add add.assoc star_n_zero_num)
+  done
 
-(* Prove equivalence between NS limits - *)
-(* seems easier than using standard definition  *)
-lemma NSLIM_h_iff: "(f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L) = ((%h. f(a + h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S L)"
-apply (simp add: NSLIM_def, auto)
-apply (drule_tac x = "star_of a + x" in spec)
-apply (drule_tac [2] x = "- star_of a + x" in spec, safe, simp)
-apply (erule mem_infmal_iff [THEN iffD2, THEN Infinitesimal_add_approx_self [THEN approx_sym]])
-apply (erule_tac [3] approx_minus_iff2 [THEN iffD1])
- prefer 2 apply (simp add: add.commute)
-apply (rule_tac x = x in star_cases)
-apply (rule_tac [2] x = x in star_cases)
-apply (auto simp add: starfun star_of_def star_n_minus star_n_add add.assoc star_n_zero_num)
-done
-
-lemma NSLIM_isCont_iff: "(f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S f a) = ((%h. f(a + h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S f a)"
+lemma NSLIM_isCont_iff: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S f a \<longleftrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S f a"
   by (fact NSLIM_h_iff)
 
-lemma isNSCont_minus: "isNSCont f a ==> isNSCont (%x. - f x) a"
-by (simp add: isNSCont_def)
+lemma isNSCont_minus: "isNSCont f a \<Longrightarrow> isNSCont (\<lambda>x. - f x) a"
+  by (simp add: isNSCont_def)
 
-lemma isNSCont_inverse:
-  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_div_algebra"
-  shows "[| isNSCont f x; f x \<noteq> 0 |] ==> isNSCont (%x. inverse (f x)) x"
-by (auto intro: isCont_inverse simp add: isNSCont_isCont_iff)
+lemma isNSCont_inverse: "isNSCont f x \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> isNSCont (\<lambda>x. inverse (f x)) x"
+  for f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_div_algebra"
+  by (auto intro: isCont_inverse simp add: isNSCont_isCont_iff)
 
-lemma isNSCont_const [simp]: "isNSCont (%x. k) a"
-by (simp add: isNSCont_def)
+lemma isNSCont_const [simp]: "isNSCont (\<lambda>x. k) a"
+  by (simp add: isNSCont_def)
 
-lemma isNSCont_abs [simp]: "isNSCont abs (a::real)"
-apply (simp add: isNSCont_def)
-apply (auto intro: approx_hrabs simp add: starfun_rabs_hrabs)
-done
+lemma isNSCont_abs [simp]: "isNSCont abs a"
+  for a :: real
+  by (auto simp: isNSCont_def intro: approx_hrabs simp: starfun_rabs_hrabs)
 
 
 subsection \<open>Uniform Continuity\<close>
 
-lemma isNSUContD: "[| isNSUCont f; x \<approx> y|] ==> ( *f* f) x \<approx> ( *f* f) y"
-by (simp add: isNSUCont_def)
+lemma isNSUContD: "isNSUCont f \<Longrightarrow> x \<approx> y \<Longrightarrow> ( *f* f) x \<approx> ( *f* f) y"
+  by (simp add: isNSUCont_def)
 
 lemma isUCont_isNSUCont:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
-  assumes f: "isUCont f" shows "isNSUCont f"
-proof (unfold isNSUCont_def, safe)
+  assumes f: "isUCont f"
+  shows "isNSUCont f"
+  unfolding isNSUCont_def
+proof safe
   fix x y :: "'a star"
   assume approx: "x \<approx> y"
   have "starfun f x - starfun f y \<in> Infinitesimal"
   proof (rule InfinitesimalI2)
-    fix r::real assume r: "0 < r"
-    with f obtain s where s: "0 < s" and
-      less_r: "\<And>x y. norm (x - y) < s \<Longrightarrow> norm (f x - f y) < r"
+    fix r :: real
+    assume r: "0 < r"
+    with f obtain s where s: "0 < s"
+      and less_r: "\<And>x y. norm (x - y) < s \<Longrightarrow> norm (f x - f y) < r"
       by (auto simp add: isUCont_def dist_norm)
     from less_r have less_r':
-       "\<And>x y. hnorm (x - y) < star_of s
-        \<Longrightarrow> hnorm (starfun f x - starfun f y) < star_of r"
+      "\<And>x y. hnorm (x - y) < star_of s \<Longrightarrow> hnorm (starfun f x - starfun f y) < star_of r"
       by transfer
     from approx have "x - y \<in> Infinitesimal"
       by (unfold approx_def)
-    hence "hnorm (x - y) < star_of s"
+    then have "hnorm (x - y) < star_of s"
       using s by (rule InfinitesimalD2)
-    thus "hnorm (starfun f x - starfun f y) < star_of r"
+    then show "hnorm (starfun f x - starfun f y) < star_of r"
       by (rule less_r')
   qed
-  thus "starfun f x \<approx> starfun f y"
+  then show "starfun f x \<approx> starfun f y"
     by (unfold approx_def)
 qed
 
 lemma isNSUCont_isUCont:
   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
-  assumes f: "isNSUCont f" shows "isUCont f"
-proof (unfold isUCont_def dist_norm, safe)
-  fix r::real assume r: "0 < r"
-  have "\<exists>s>0. \<forall>x y. hnorm (x - y) < s
-        \<longrightarrow> hnorm (starfun f x - starfun f y) < star_of r"
+  assumes f: "isNSUCont f"
+  shows "isUCont f"
+  unfolding isUCont_def dist_norm
+proof safe
+  fix r :: real
+  assume r: "0 < r"
+  have "\<exists>s>0. \<forall>x y. hnorm (x - y) < s \<longrightarrow> hnorm (starfun f x - starfun f y) < star_of r"
   proof (rule exI, safe)
-    show "0 < \<epsilon>" by (rule hypreal_epsilon_gt_zero)
+    show "0 < \<epsilon>"
+      by (rule hypreal_epsilon_gt_zero)
   next
     fix x y :: "'a star"
     assume "hnorm (x - y) < \<epsilon>"
-    with Infinitesimal_epsilon
-    have "x - y \<in> Infinitesimal"
+    with Infinitesimal_epsilon have "x - y \<in> Infinitesimal"
       by (rule hnorm_less_Infinitesimal)
-    hence "x \<approx> y"
+    then have "x \<approx> y"
       by (unfold approx_def)
     with f have "starfun f x \<approx> starfun f y"
       by (simp add: isNSUCont_def)
-    hence "starfun f x - starfun f y \<in> Infinitesimal"
+    then have "starfun f x - starfun f y \<in> Infinitesimal"
       by (unfold approx_def)
-    thus "hnorm (starfun f x - starfun f y) < star_of r"
+    then show "hnorm (starfun f x - starfun f y) < star_of r"
       using r by (rule InfinitesimalD2)
   qed
-  thus "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r"
+  then show "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r"
     by transfer
 qed
 
--- a/src/HOL/Nonstandard_Analysis/HyperDef.thy	Mon Oct 31 16:26:36 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/HyperDef.thy	Tue Nov 01 00:44:24 2016 +0100
@@ -4,50 +4,42 @@
     Conversion to Isar and new proofs by Lawrence C Paulson, 2004
 *)
 
-section\<open>Construction of Hyperreals Using Ultrafilters\<close>
+section \<open>Construction of Hyperreals Using Ultrafilters\<close>
 
 theory HyperDef
-imports Complex_Main HyperNat
+  imports Complex_Main HyperNat
 begin
 
 type_synonym hypreal = "real star"
 
-abbreviation
-  hypreal_of_real :: "real => real star" where
-  "hypreal_of_real == star_of"
+abbreviation hypreal_of_real :: "real \<Rightarrow> real star"
+  where "hypreal_of_real \<equiv> star_of"
 
-abbreviation
-  hypreal_of_hypnat :: "hypnat \<Rightarrow> hypreal" where
-  "hypreal_of_hypnat \<equiv> of_hypnat"
+abbreviation hypreal_of_hypnat :: "hypnat \<Rightarrow> hypreal"
+  where "hypreal_of_hypnat \<equiv> of_hypnat"
 
-definition
-  omega :: hypreal  ("\<omega>") where
-   \<comment> \<open>an infinite number \<open>= [<1,2,3,...>]\<close>\<close>
-  "\<omega> = star_n (\<lambda>n. real (Suc n))"
+definition omega :: hypreal  ("\<omega>")
+  where "\<omega> = star_n (\<lambda>n. real (Suc n))"
+    \<comment> \<open>an infinite number \<open>= [<1, 2, 3, \<dots>>]\<close>\<close>
 
-definition
-  epsilon :: hypreal  ("\<epsilon>") where
-   \<comment> \<open>an infinitesimal number \<open>= [<1,1/2,1/3,...>]\<close>\<close>
-  "\<epsilon> = star_n (\<lambda>n. inverse (real (Suc n)))"
+definition epsilon :: hypreal  ("\<epsilon>")
+  where "\<epsilon> = star_n (\<lambda>n. inverse (real (Suc n)))"
+    \<comment> \<open>an infinitesimal number \<open>= [<1, 1/2, 1/3, \<dots>>]\<close>\<close>
 
 
 subsection \<open>Real vector class instances\<close>
 
 instantiation star :: (scaleR) scaleR
 begin
-
-definition
-  star_scaleR_def [transfer_unfold]: "scaleR r \<equiv> *f* (scaleR r)"
-
-instance ..
-
+  definition star_scaleR_def [transfer_unfold]: "scaleR r \<equiv> *f* (scaleR r)"
+  instance ..
 end
 
 lemma Standard_scaleR [simp]: "x \<in> Standard \<Longrightarrow> scaleR r x \<in> Standard"
-by (simp add: star_scaleR_def)
+  by (simp add: star_scaleR_def)
 
 lemma star_of_scaleR [simp]: "star_of (scaleR r x) = scaleR r (star_of x)"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 instance star :: (real_vector) real_vector
 proof
@@ -80,230 +72,213 @@
 instance star :: (real_field) real_field ..
 
 lemma star_of_real_def [transfer_unfold]: "of_real r = star_of (of_real r)"
-by (unfold of_real_def, transfer, rule refl)
+  by (unfold of_real_def, transfer, rule refl)
 
 lemma Standard_of_real [simp]: "of_real r \<in> Standard"
-by (simp add: star_of_real_def)
+  by (simp add: star_of_real_def)
 
 lemma star_of_of_real [simp]: "star_of (of_real r) = of_real r"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma of_real_eq_star_of [simp]: "of_real = star_of"
 proof
-  fix r :: real
-  show "of_real r = star_of r"
+  show "of_real r = star_of r" for r :: real
     by transfer simp
 qed
 
 lemma Reals_eq_Standard: "(\<real> :: hypreal set) = Standard"
-by (simp add: Reals_def Standard_def)
+  by (simp add: Reals_def Standard_def)
 
 
 subsection \<open>Injection from @{typ hypreal}\<close>
 
-definition
-  of_hypreal :: "hypreal \<Rightarrow> 'a::real_algebra_1 star" where
-  [transfer_unfold]: "of_hypreal = *f* of_real"
+definition of_hypreal :: "hypreal \<Rightarrow> 'a::real_algebra_1 star"
+  where [transfer_unfold]: "of_hypreal = *f* of_real"
 
-lemma Standard_of_hypreal [simp]:
-  "r \<in> Standard \<Longrightarrow> of_hypreal r \<in> Standard"
-by (simp add: of_hypreal_def)
+lemma Standard_of_hypreal [simp]: "r \<in> Standard \<Longrightarrow> of_hypreal r \<in> Standard"
+  by (simp add: of_hypreal_def)
 
 lemma of_hypreal_0 [simp]: "of_hypreal 0 = 0"
-by transfer (rule of_real_0)
+  by transfer (rule of_real_0)
 
 lemma of_hypreal_1 [simp]: "of_hypreal 1 = 1"
-by transfer (rule of_real_1)
+  by transfer (rule of_real_1)
 
-lemma of_hypreal_add [simp]:
-  "\<And>x y. of_hypreal (x + y) = of_hypreal x + of_hypreal y"
-by transfer (rule of_real_add)
+lemma of_hypreal_add [simp]: "\<And>x y. of_hypreal (x + y) = of_hypreal x + of_hypreal y"
+  by transfer (rule of_real_add)
 
 lemma of_hypreal_minus [simp]: "\<And>x. of_hypreal (- x) = - of_hypreal x"
-by transfer (rule of_real_minus)
+  by transfer (rule of_real_minus)
 
-lemma of_hypreal_diff [simp]:
-  "\<And>x y. of_hypreal (x - y) = of_hypreal x - of_hypreal y"
-by transfer (rule of_real_diff)
+lemma of_hypreal_diff [simp]: "\<And>x y. of_hypreal (x - y) = of_hypreal x - of_hypreal y"
+  by transfer (rule of_real_diff)
 
-lemma of_hypreal_mult [simp]:
-  "\<And>x y. of_hypreal (x * y) = of_hypreal x * of_hypreal y"
-by transfer (rule of_real_mult)
+lemma of_hypreal_mult [simp]: "\<And>x y. of_hypreal (x * y) = of_hypreal x * of_hypreal y"
+  by transfer (rule of_real_mult)
 
 lemma of_hypreal_inverse [simp]:
   "\<And>x. of_hypreal (inverse x) =
-   inverse (of_hypreal x :: 'a::{real_div_algebra, division_ring} star)"
-by transfer (rule of_real_inverse)
+    inverse (of_hypreal x :: 'a::{real_div_algebra, division_ring} star)"
+  by transfer (rule of_real_inverse)
 
 lemma of_hypreal_divide [simp]:
   "\<And>x y. of_hypreal (x / y) =
-   (of_hypreal x / of_hypreal y :: 'a::{real_field, field} star)"
-by transfer (rule of_real_divide)
+    (of_hypreal x / of_hypreal y :: 'a::{real_field, field} star)"
+  by transfer (rule of_real_divide)
 
-lemma of_hypreal_eq_iff [simp]:
-  "\<And>x y. (of_hypreal x = of_hypreal y) = (x = y)"
-by transfer (rule of_real_eq_iff)
+lemma of_hypreal_eq_iff [simp]: "\<And>x y. (of_hypreal x = of_hypreal y) = (x = y)"
+  by transfer (rule of_real_eq_iff)
 
-lemma of_hypreal_eq_0_iff [simp]:
-  "\<And>x. (of_hypreal x = 0) = (x = 0)"
-by transfer (rule of_real_eq_0_iff)
+lemma of_hypreal_eq_0_iff [simp]: "\<And>x. (of_hypreal x = 0) = (x = 0)"
+  by transfer (rule of_real_eq_0_iff)
 
 
-subsection\<open>Properties of @{term starrel}\<close>
+subsection \<open>Properties of @{term starrel}\<close>
 
 lemma lemma_starrel_refl [simp]: "x \<in> starrel `` {x}"
-by (simp add: starrel_def)
+  by (simp add: starrel_def)
 
 lemma starrel_in_hypreal [simp]: "starrel``{x}:star"
-by (simp add: star_def starrel_def quotient_def, blast)
+  by (simp add: star_def starrel_def quotient_def, blast)
 
 declare Abs_star_inject [simp] Abs_star_inverse [simp]
 declare equiv_starrel [THEN eq_equiv_class_iff, simp]
 
-subsection\<open>@{term hypreal_of_real}:
-            the Injection from @{typ real} to @{typ hypreal}\<close>
+
+subsection \<open>@{term hypreal_of_real}: the Injection from @{typ real} to @{typ hypreal}\<close>
 
 lemma inj_star_of: "inj star_of"
-by (rule inj_onI, simp)
+  by (rule inj_onI) simp
 
-lemma mem_Rep_star_iff: "(X \<in> Rep_star x) = (x = star_n X)"
-by (cases x, simp add: star_n_def)
+lemma mem_Rep_star_iff: "X \<in> Rep_star x \<longleftrightarrow> x = star_n X"
+  by (cases x) (simp add: star_n_def)
 
-lemma Rep_star_star_n_iff [simp]:
-  "(X \<in> Rep_star (star_n Y)) = (eventually (\<lambda>n. Y n = X n) \<U>)"
-by (simp add: star_n_def)
+lemma Rep_star_star_n_iff [simp]: "X \<in> Rep_star (star_n Y) \<longleftrightarrow> eventually (\<lambda>n. Y n = X n) \<U>"
+  by (simp add: star_n_def)
 
 lemma Rep_star_star_n: "X \<in> Rep_star (star_n X)"
-by simp
+  by simp
 
-subsection\<open>Properties of @{term star_n}\<close>
 
-lemma star_n_add:
-  "star_n X + star_n Y = star_n (%n. X n + Y n)"
-by (simp only: star_add_def starfun2_star_n)
+subsection \<open>Properties of @{term star_n}\<close>
+
+lemma star_n_add: "star_n X + star_n Y = star_n (\<lambda>n. X n + Y n)"
+  by (simp only: star_add_def starfun2_star_n)
 
-lemma star_n_minus:
-   "- star_n X = star_n (%n. -(X n))"
-by (simp only: star_minus_def starfun_star_n)
+lemma star_n_minus: "- star_n X = star_n (\<lambda>n. -(X n))"
+  by (simp only: star_minus_def starfun_star_n)
 
-lemma star_n_diff:
-     "star_n X - star_n Y = star_n (%n. X n - Y n)"
-by (simp only: star_diff_def starfun2_star_n)
+lemma star_n_diff: "star_n X - star_n Y = star_n (\<lambda>n. X n - Y n)"
+  by (simp only: star_diff_def starfun2_star_n)
 
-lemma star_n_mult:
-  "star_n X * star_n Y = star_n (%n. X n * Y n)"
-by (simp only: star_mult_def starfun2_star_n)
+lemma star_n_mult: "star_n X * star_n Y = star_n (\<lambda>n. X n * Y n)"
+  by (simp only: star_mult_def starfun2_star_n)
 
-lemma star_n_inverse:
-      "inverse (star_n X) = star_n (%n. inverse(X n))"
-by (simp only: star_inverse_def starfun_star_n)
+lemma star_n_inverse: "inverse (star_n X) = star_n (\<lambda>n. inverse (X n))"
+  by (simp only: star_inverse_def starfun_star_n)
 
-lemma star_n_le:
-      "star_n X \<le> star_n Y = (eventually (\<lambda>n. X n \<le> Y n) FreeUltrafilterNat)"
-by (simp only: star_le_def starP2_star_n)
+lemma star_n_le: "star_n X \<le> star_n Y = eventually (\<lambda>n. X n \<le> Y n) FreeUltrafilterNat"
+  by (simp only: star_le_def starP2_star_n)
+
+lemma star_n_less: "star_n X < star_n Y = eventually (\<lambda>n. X n < Y n) FreeUltrafilterNat"
+  by (simp only: star_less_def starP2_star_n)
 
-lemma star_n_less:
-      "star_n X < star_n Y = (eventually (\<lambda>n. X n < Y n) FreeUltrafilterNat)"
-by (simp only: star_less_def starP2_star_n)
+lemma star_n_zero_num: "0 = star_n (\<lambda>n. 0)"
+  by (simp only: star_zero_def star_of_def)
 
-lemma star_n_zero_num: "0 = star_n (%n. 0)"
-by (simp only: star_zero_def star_of_def)
+lemma star_n_one_num: "1 = star_n (\<lambda>n. 1)"
+  by (simp only: star_one_def star_of_def)
 
-lemma star_n_one_num: "1 = star_n (%n. 1)"
-by (simp only: star_one_def star_of_def)
-
-lemma star_n_abs: "\<bar>star_n X\<bar> = star_n (%n. \<bar>X n\<bar>)"
-by (simp only: star_abs_def starfun_star_n)
+lemma star_n_abs: "\<bar>star_n X\<bar> = star_n (\<lambda>n. \<bar>X n\<bar>)"
+  by (simp only: star_abs_def starfun_star_n)
 
 lemma hypreal_omega_gt_zero [simp]: "0 < \<omega>"
-by (simp add: omega_def star_n_zero_num star_n_less)
-
-subsection\<open>Existence of Infinite Hyperreal Number\<close>
-
-text\<open>Existence of infinite number not corresponding to any real number.
-Use assumption that member @{term FreeUltrafilterNat} is not finite.\<close>
+  by (simp add: omega_def star_n_zero_num star_n_less)
 
 
-text\<open>A few lemmas first\<close>
+subsection \<open>Existence of Infinite Hyperreal Number\<close>
+
+text \<open>Existence of infinite number not corresponding to any real number.
+  Use assumption that member @{term FreeUltrafilterNat} is not finite.\<close>
+
+text \<open>A few lemmas first.\<close>
 
 lemma lemma_omega_empty_singleton_disj:
   "{n::nat. x = real n} = {} \<or> (\<exists>y. {n::nat. x = real n} = {y})"
-by force
+  by force
 
 lemma lemma_finite_omega_set: "finite {n::nat. x = real n}"
   using lemma_omega_empty_singleton_disj [of x] by auto
 
-lemma not_ex_hypreal_of_real_eq_omega:
-      "~ (\<exists>x. hypreal_of_real x = \<omega>)"
-apply (simp add: omega_def star_of_def star_n_eq_iff)
-apply clarify
-apply (rule_tac x2="x-1" in lemma_finite_omega_set [THEN FreeUltrafilterNat.finite, THEN notE])
-apply (erule eventually_mono)
-apply auto
-done
+lemma not_ex_hypreal_of_real_eq_omega: "\<nexists>x. hypreal_of_real x = \<omega>"
+  apply (simp add: omega_def star_of_def star_n_eq_iff)
+  apply clarify
+  apply (rule_tac x2="x-1" in lemma_finite_omega_set [THEN FreeUltrafilterNat.finite, THEN notE])
+  apply (erule eventually_mono)
+  apply auto
+  done
 
 lemma hypreal_of_real_not_eq_omega: "hypreal_of_real x \<noteq> \<omega>"
-by (insert not_ex_hypreal_of_real_eq_omega, auto)
+  using not_ex_hypreal_of_real_eq_omega by auto
 
-text\<open>Existence of infinitesimal number also not corresponding to any
- real number\<close>
+text \<open>Existence of infinitesimal number also not corresponding to any real number.\<close>
 
 lemma lemma_epsilon_empty_singleton_disj:
-     "{n::nat. x = inverse(real(Suc n))} = {} |
-      (\<exists>y. {n::nat. x = inverse(real(Suc n))} = {y})"
-by auto
+  "{n::nat. x = inverse(real(Suc n))} = {} \<or> (\<exists>y. {n::nat. x = inverse(real(Suc n))} = {y})"
+  by auto
 
-lemma lemma_finite_epsilon_set: "finite {n. x = inverse(real(Suc n))}"
-by (cut_tac x = x in lemma_epsilon_empty_singleton_disj, auto)
+lemma lemma_finite_epsilon_set: "finite {n. x = inverse (real (Suc n))}"
+  using lemma_epsilon_empty_singleton_disj [of x] by auto
 
-lemma not_ex_hypreal_of_real_eq_epsilon: "~ (\<exists>x. hypreal_of_real x = \<epsilon>)"
-by (auto simp add: epsilon_def star_of_def star_n_eq_iff
-                   lemma_finite_epsilon_set [THEN FreeUltrafilterNat.finite] simp del: of_nat_Suc)
+lemma not_ex_hypreal_of_real_eq_epsilon: "\<nexists>x. hypreal_of_real x = \<epsilon>"
+  by (auto simp: epsilon_def star_of_def star_n_eq_iff
+      lemma_finite_epsilon_set [THEN FreeUltrafilterNat.finite] simp del: of_nat_Suc)
 
 lemma hypreal_of_real_not_eq_epsilon: "hypreal_of_real x \<noteq> \<epsilon>"
-by (insert not_ex_hypreal_of_real_eq_epsilon, auto)
+  using not_ex_hypreal_of_real_eq_epsilon by auto
 
 lemma hypreal_epsilon_not_zero: "\<epsilon> \<noteq> 0"
-by (simp add: epsilon_def star_zero_def star_of_def star_n_eq_iff FreeUltrafilterNat.proper
-         del: star_of_zero)
+  by (simp add: epsilon_def star_zero_def star_of_def star_n_eq_iff FreeUltrafilterNat.proper
+      del: star_of_zero)
 
 lemma hypreal_epsilon_inverse_omega: "\<epsilon> = inverse \<omega>"
-by (simp add: epsilon_def omega_def star_n_inverse)
+  by (simp add: epsilon_def omega_def star_n_inverse)
 
 lemma hypreal_epsilon_gt_zero: "0 < \<epsilon>"
-by (simp add: hypreal_epsilon_inverse_omega)
-
-subsection\<open>Absolute Value Function for the Hyperreals\<close>
-
-lemma hrabs_add_less: "[| \<bar>x\<bar> < r; \<bar>y\<bar> < s |] ==> \<bar>x + y\<bar> < r + (s::hypreal)"
-by (simp add: abs_if split: if_split_asm)
-
-lemma hrabs_less_gt_zero: "\<bar>x\<bar> < r ==> (0::hypreal) < r"
-by (blast intro!: order_le_less_trans abs_ge_zero)
-
-lemma hrabs_disj: "\<bar>x\<bar> = (x::'a::abs_if) \<or> \<bar>x\<bar> = -x"
-by (simp add: abs_if)
-
-lemma hrabs_add_lemma_disj: "(y::hypreal) + - x + (y + - z) = \<bar>x + - z\<bar> ==> y = z | x = y"
-by (simp add: abs_if split: if_split_asm)
+  by (simp add: hypreal_epsilon_inverse_omega)
 
 
-subsection\<open>Embedding the Naturals into the Hyperreals\<close>
+subsection \<open>Absolute Value Function for the Hyperreals\<close>
+
+lemma hrabs_add_less: "\<bar>x\<bar> < r \<Longrightarrow> \<bar>y\<bar> < s \<Longrightarrow> \<bar>x + y\<bar> < r + s"
+  for x y r s :: hypreal
+  by (simp add: abs_if split: if_split_asm)
+
+lemma hrabs_less_gt_zero: "\<bar>x\<bar> < r \<Longrightarrow> 0 < r"
+  for x r :: hypreal
+  by (blast intro!: order_le_less_trans abs_ge_zero)
 
-abbreviation
-  hypreal_of_nat :: "nat => hypreal" where
-  "hypreal_of_nat == of_nat"
+lemma hrabs_disj: "\<bar>x\<bar> = x \<or> \<bar>x\<bar> = -x"
+  for x :: "'a::abs_if"
+  by (simp add: abs_if)
+
+lemma hrabs_add_lemma_disj: "y + - x + (y + - z) = \<bar>x + - z\<bar> \<Longrightarrow> y = z \<or> x = y"
+  for x y z :: hypreal
+  by (simp add: abs_if split: if_split_asm)
+
+
+subsection \<open>Embedding the Naturals into the Hyperreals\<close>
+
+abbreviation hypreal_of_nat :: "nat \<Rightarrow> hypreal"
+  where "hypreal_of_nat \<equiv> of_nat"
 
 lemma SNat_eq: "Nats = {n. \<exists>N. n = hypreal_of_nat N}"
-by (simp add: Nats_def image_def)
+  by (simp add: Nats_def image_def)
 
-(*------------------------------------------------------------*)
-(* naturals embedded in hyperreals                            *)
-(* is a hyperreal c.f. NS extension                           *)
-(*------------------------------------------------------------*)
+text \<open>Naturals embedded in hyperreals: is a hyperreal c.f. NS extension.\<close>
 
-lemma hypreal_of_nat: "hypreal_of_nat m = star_n (%n. real m)"
-by (simp add: star_of_def [symmetric])
+lemma hypreal_of_nat: "hypreal_of_nat m = star_n (\<lambda>n. real m)"
+  by (simp add: star_of_def [symmetric])
 
 declaration \<open>
   K (Lin_Arith.add_inj_thms [@{thm star_of_le} RS iffD2,
@@ -314,75 +289,77 @@
   #> Lin_Arith.add_inj_const (@{const_name "StarDef.star_of"}, @{typ "real \<Rightarrow> hypreal"}))
 \<close>
 
-simproc_setup fast_arith_hypreal ("(m::hypreal) < n" | "(m::hypreal) <= n" | "(m::hypreal) = n") =
+simproc_setup fast_arith_hypreal ("(m::hypreal) < n" | "(m::hypreal) \<le> n" | "(m::hypreal) = n") =
   \<open>K Lin_Arith.simproc\<close>
 
 
 subsection \<open>Exponentials on the Hyperreals\<close>
 
-lemma hpowr_0 [simp]:   "r ^ 0       = (1::hypreal)"
-by (rule power_0)
-
-lemma hpowr_Suc [simp]: "r ^ (Suc n) = (r::hypreal) * (r ^ n)"
-by (rule power_Suc)
+lemma hpowr_0 [simp]: "r ^ 0 = (1::hypreal)"
+  for r :: hypreal
+  by (rule power_0)
 
-lemma hrealpow_two: "(r::hypreal) ^ Suc (Suc 0) = r * r"
-by simp
+lemma hpowr_Suc [simp]: "r ^ (Suc n) = r * (r ^ n)"
+  for r :: hypreal
+  by (rule power_Suc)
 
-lemma hrealpow_two_le [simp]: "(0::hypreal) \<le> r ^ Suc (Suc 0)"
-by (auto simp add: zero_le_mult_iff)
+lemma hrealpow_two: "r ^ Suc (Suc 0) = r * r"
+  for r :: hypreal
+  by simp
 
-lemma hrealpow_two_le_add_order [simp]:
-     "(0::hypreal) \<le> u ^ Suc (Suc 0) + v ^ Suc (Suc 0)"
-by (simp only: hrealpow_two_le add_nonneg_nonneg)
+lemma hrealpow_two_le [simp]: "0 \<le> r ^ Suc (Suc 0)"
+  for r :: hypreal
+  by (auto simp add: zero_le_mult_iff)
+
+lemma hrealpow_two_le_add_order [simp]: "0 \<le> u ^ Suc (Suc 0) + v ^ Suc (Suc 0)"
+  for u v :: hypreal
+  by (simp only: hrealpow_two_le add_nonneg_nonneg)
 
-lemma hrealpow_two_le_add_order2 [simp]:
-     "(0::hypreal) \<le> u ^ Suc (Suc 0) + v ^ Suc (Suc 0) + w ^ Suc (Suc 0)"
-by (simp only: hrealpow_two_le add_nonneg_nonneg)
+lemma hrealpow_two_le_add_order2 [simp]: "0 \<le> u ^ Suc (Suc 0) + v ^ Suc (Suc 0) + w ^ Suc (Suc 0)"
+  for u v w :: hypreal
+  by (simp only: hrealpow_two_le add_nonneg_nonneg)
 
-lemma hypreal_add_nonneg_eq_0_iff:
-     "[| 0 \<le> x; 0 \<le> y |] ==> (x+y = 0) = (x = 0 & y = (0::hypreal))"
-by arith
+lemma hypreal_add_nonneg_eq_0_iff: "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> x + y = 0 \<longleftrightarrow> x = 0 \<and> y = 0"
+  for x y :: hypreal
+  by arith
 
 
-text\<open>FIXME: DELETE THESE\<close>
-lemma hypreal_three_squares_add_zero_iff:
-     "(x*x + y*y + z*z = 0) = (x = 0 & y = 0 & z = (0::hypreal))"
-apply (simp only: zero_le_square add_nonneg_nonneg hypreal_add_nonneg_eq_0_iff, auto)
-done
+(* FIXME: DELETE THESE *)
+lemma hypreal_three_squares_add_zero_iff: "x * x + y * y + z * z = 0 \<longleftrightarrow> x = 0 \<and> y = 0 \<and> z = 0"
+  for x y z :: hypreal
+  by (simp only: zero_le_square add_nonneg_nonneg hypreal_add_nonneg_eq_0_iff) auto
 
 lemma hrealpow_three_squares_add_zero_iff [simp]:
-     "(x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + z ^ Suc (Suc 0) = (0::hypreal)) =
-      (x = 0 & y = 0 & z = 0)"
-by (simp only: hypreal_three_squares_add_zero_iff hrealpow_two)
+  "x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + z ^ Suc (Suc 0) = 0 \<longleftrightarrow> x = 0 \<and> y = 0 \<and> z = 0"
+  for x y z :: hypreal
+  by (simp only: hypreal_three_squares_add_zero_iff hrealpow_two)
 
 (*FIXME: This and RealPow.abs_realpow_two should be replaced by an abstract
   result proved in Rings or Fields*)
-lemma hrabs_hrealpow_two [simp]: "\<bar>x ^ Suc (Suc 0)\<bar> = (x::hypreal) ^ Suc (Suc 0)"
-by (simp add: abs_mult)
+lemma hrabs_hrealpow_two [simp]: "\<bar>x ^ Suc (Suc 0)\<bar> = x ^ Suc (Suc 0)"
+  for x :: hypreal
+  by (simp add: abs_mult)
 
 lemma two_hrealpow_ge_one [simp]: "(1::hypreal) \<le> 2 ^ n"
-by (insert power_increasing [of 0 n "2::hypreal"], simp)
+  using power_increasing [of 0 n "2::hypreal"] by simp
 
-lemma hrealpow:
-    "star_n X ^ m = star_n (%n. (X n::real) ^ m)"
-apply (induct_tac "m")
-apply (auto simp add: star_n_one_num star_n_mult power_0)
-done
+lemma hrealpow: "star_n X ^ m = star_n (\<lambda>n. (X n::real) ^ m)"
+  by (induct m) (auto simp: star_n_one_num star_n_mult)
 
 lemma hrealpow_sum_square_expand:
-     "(x + (y::hypreal)) ^ Suc (Suc 0) =
-      x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + (hypreal_of_nat (Suc (Suc 0)))*x*y"
-by (simp add: distrib_left distrib_right)
+  "(x + y) ^ Suc (Suc 0) =
+    x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + (hypreal_of_nat (Suc (Suc 0))) * x * y"
+  for x y :: hypreal
+  by (simp add: distrib_left distrib_right)
 
 lemma power_hypreal_of_real_numeral:
-     "(numeral v :: hypreal) ^ n = hypreal_of_real ((numeral v) ^ n)"
-by simp
+  "(numeral v :: hypreal) ^ n = hypreal_of_real ((numeral v) ^ n)"
+  by simp
 declare power_hypreal_of_real_numeral [of _ "numeral w", simp] for w
 
 lemma power_hypreal_of_real_neg_numeral:
-     "(- numeral v :: hypreal) ^ n = hypreal_of_real ((- numeral v) ^ n)"
-by simp
+  "(- numeral v :: hypreal) ^ n = hypreal_of_real ((- numeral v) ^ n)"
+  by simp
 declare power_hypreal_of_real_neg_numeral [of _ "numeral w", simp] for w
 (*
 lemma hrealpow_HFinite:
@@ -393,148 +370,119 @@
 done
 *)
 
-subsection\<open>Powers with Hypernatural Exponents\<close>
 
-definition pow :: "['a::power star, nat star] \<Rightarrow> 'a star" (infixr "pow" 80) where
-  hyperpow_def [transfer_unfold]: "R pow N = ( *f2* op ^) R N"
-  (* hypernatural powers of hyperreals *)
+subsection \<open>Powers with Hypernatural Exponents\<close>
 
-lemma Standard_hyperpow [simp]:
-  "\<lbrakk>r \<in> Standard; n \<in> Standard\<rbrakk> \<Longrightarrow> r pow n \<in> Standard"
-unfolding hyperpow_def by simp
+text \<open>Hypernatural powers of hyperreals.\<close>
+definition pow :: "'a::power star \<Rightarrow> nat star \<Rightarrow> 'a star"  (infixr "pow" 80)
+  where hyperpow_def [transfer_unfold]: "R pow N = ( *f2* op ^) R N"
 
-lemma hyperpow: "star_n X pow star_n Y = star_n (%n. X n ^ Y n)"
-by (simp add: hyperpow_def starfun2_star_n)
-
-lemma hyperpow_zero [simp]:
-  "\<And>n. (0::'a::{power,semiring_0} star) pow (n + (1::hypnat)) = 0"
-by transfer simp
+lemma Standard_hyperpow [simp]: "r \<in> Standard \<Longrightarrow> n \<in> Standard \<Longrightarrow> r pow n \<in> Standard"
+  by (simp add: hyperpow_def)
 
-lemma hyperpow_not_zero:
-  "\<And>r n. r \<noteq> (0::'a::{field} star) ==> r pow n \<noteq> 0"
-by transfer (rule power_not_zero)
+lemma hyperpow: "star_n X pow star_n Y = star_n (\<lambda>n. X n ^ Y n)"
+  by (simp add: hyperpow_def starfun2_star_n)
+
+lemma hyperpow_zero [simp]: "\<And>n. (0::'a::{power,semiring_0} star) pow (n + (1::hypnat)) = 0"
+  by transfer simp
 
-lemma hyperpow_inverse:
-  "\<And>r n. r \<noteq> (0::'a::field star)
-   \<Longrightarrow> inverse (r pow n) = (inverse r) pow n"
-by transfer (rule power_inverse [symmetric])
+lemma hyperpow_not_zero: "\<And>r n. r \<noteq> (0::'a::{field} star) \<Longrightarrow> r pow n \<noteq> 0"
+  by transfer (rule power_not_zero)
 
-lemma hyperpow_hrabs:
-  "\<And>r n. \<bar>r::'a::{linordered_idom} star\<bar> pow n = \<bar>r pow n\<bar>"
-by transfer (rule power_abs [symmetric])
+lemma hyperpow_inverse: "\<And>r n. r \<noteq> (0::'a::field star) \<Longrightarrow> inverse (r pow n) = (inverse r) pow n"
+  by transfer (rule power_inverse [symmetric])
 
-lemma hyperpow_add:
-  "\<And>r n m. (r::'a::monoid_mult star) pow (n + m) = (r pow n) * (r pow m)"
-by transfer (rule power_add)
+lemma hyperpow_hrabs: "\<And>r n. \<bar>r::'a::{linordered_idom} star\<bar> pow n = \<bar>r pow n\<bar>"
+  by transfer (rule power_abs [symmetric])
 
-lemma hyperpow_one [simp]:
-  "\<And>r. (r::'a::monoid_mult star) pow (1::hypnat) = r"
-by transfer (rule power_one_right)
+lemma hyperpow_add: "\<And>r n m. (r::'a::monoid_mult star) pow (n + m) = (r pow n) * (r pow m)"
+  by transfer (rule power_add)
 
-lemma hyperpow_two:
-  "\<And>r. (r::'a::monoid_mult star) pow (2::hypnat) = r * r"
-by transfer (rule power2_eq_square)
+lemma hyperpow_one [simp]: "\<And>r. (r::'a::monoid_mult star) pow (1::hypnat) = r"
+  by transfer (rule power_one_right)
 
-lemma hyperpow_gt_zero:
-  "\<And>r n. (0::'a::{linordered_semidom} star) < r \<Longrightarrow> 0 < r pow n"
-by transfer (rule zero_less_power)
+lemma hyperpow_two: "\<And>r. (r::'a::monoid_mult star) pow (2::hypnat) = r * r"
+  by transfer (rule power2_eq_square)
 
-lemma hyperpow_ge_zero:
-  "\<And>r n. (0::'a::{linordered_semidom} star) \<le> r \<Longrightarrow> 0 \<le> r pow n"
-by transfer (rule zero_le_power)
+lemma hyperpow_gt_zero: "\<And>r n. (0::'a::{linordered_semidom} star) < r \<Longrightarrow> 0 < r pow n"
+  by transfer (rule zero_less_power)
+
+lemma hyperpow_ge_zero: "\<And>r n. (0::'a::{linordered_semidom} star) \<le> r \<Longrightarrow> 0 \<le> r pow n"
+  by transfer (rule zero_le_power)
 
-lemma hyperpow_le:
-  "\<And>x y n. \<lbrakk>(0::'a::{linordered_semidom} star) < x; x \<le> y\<rbrakk>
-   \<Longrightarrow> x pow n \<le> y pow n"
-by transfer (rule power_mono [OF _ order_less_imp_le])
+lemma hyperpow_le: "\<And>x y n. (0::'a::{linordered_semidom} star) < x \<Longrightarrow> x \<le> y \<Longrightarrow> x pow n \<le> y pow n"
+  by transfer (rule power_mono [OF _ order_less_imp_le])
 
-lemma hyperpow_eq_one [simp]:
-  "\<And>n. 1 pow n = (1::'a::monoid_mult star)"
-by transfer (rule power_one)
+lemma hyperpow_eq_one [simp]: "\<And>n. 1 pow n = (1::'a::monoid_mult star)"
+  by transfer (rule power_one)
 
-lemma hrabs_hyperpow_minus [simp]:
-  "\<And>(a::'a::{linordered_idom} star) n. \<bar>(-a) pow n\<bar> = \<bar>a pow n\<bar>"
-by transfer (rule abs_power_minus)
+lemma hrabs_hyperpow_minus [simp]: "\<And>(a::'a::linordered_idom star) n. \<bar>(-a) pow n\<bar> = \<bar>a pow n\<bar>"
+  by transfer (rule abs_power_minus)
 
-lemma hyperpow_mult:
-  "\<And>r s n. (r * s::'a::{comm_monoid_mult} star) pow n
-   = (r pow n) * (s pow n)"
-by transfer (rule power_mult_distrib)
+lemma hyperpow_mult: "\<And>r s n. (r * s::'a::comm_monoid_mult star) pow n = (r pow n) * (s pow n)"
+  by transfer (rule power_mult_distrib)
 
-lemma hyperpow_two_le [simp]:
-  "\<And>r. (0::'a::{monoid_mult,linordered_ring_strict} star) \<le> r pow 2"
-by (auto simp add: hyperpow_two zero_le_mult_iff)
+lemma hyperpow_two_le [simp]: "\<And>r. (0::'a::{monoid_mult,linordered_ring_strict} star) \<le> r pow 2"
+  by (auto simp add: hyperpow_two zero_le_mult_iff)
 
 lemma hrabs_hyperpow_two [simp]:
-  "\<bar>x pow 2\<bar> =
-   (x::'a::{monoid_mult,linordered_ring_strict} star) pow 2"
-by (simp only: abs_of_nonneg hyperpow_two_le)
+  "\<bar>x pow 2\<bar> = (x::'a::{monoid_mult,linordered_ring_strict} star) pow 2"
+  by (simp only: abs_of_nonneg hyperpow_two_le)
 
-lemma hyperpow_two_hrabs [simp]:
-  "\<bar>x::'a::{linordered_idom} star\<bar> pow 2 = x pow 2"
-by (simp add: hyperpow_hrabs)
+lemma hyperpow_two_hrabs [simp]: "\<bar>x::'a::linordered_idom star\<bar> pow 2 = x pow 2"
+  by (simp add: hyperpow_hrabs)
 
-text\<open>The precondition could be weakened to @{term "0\<le>x"}\<close>
-lemma hypreal_mult_less_mono:
-     "[| u<v;  x<y;  (0::hypreal) < v;  0 < x |] ==> u*x < v* y"
- by (simp add: mult_strict_mono order_less_imp_le)
+text \<open>The precondition could be weakened to @{term "0\<le>x"}.\<close>
+lemma hypreal_mult_less_mono: "u < v \<Longrightarrow> x < y \<Longrightarrow> 0 < v \<Longrightarrow> 0 < x \<Longrightarrow> u * x < v * y"
+  for u v x y :: hypreal
+  by (simp add: mult_strict_mono order_less_imp_le)
 
-lemma hyperpow_two_gt_one:
-  "\<And>r::'a::{linordered_semidom} star. 1 < r \<Longrightarrow> 1 < r pow 2"
-by transfer simp
+lemma hyperpow_two_gt_one: "\<And>r::'a::linordered_semidom star. 1 < r \<Longrightarrow> 1 < r pow 2"
+  by transfer simp
 
-lemma hyperpow_two_ge_one:
-  "\<And>r::'a::{linordered_semidom} star. 1 \<le> r \<Longrightarrow> 1 \<le> r pow 2"
-by transfer (rule one_le_power)
+lemma hyperpow_two_ge_one: "\<And>r::'a::linordered_semidom star. 1 \<le> r \<Longrightarrow> 1 \<le> r pow 2"
+  by transfer (rule one_le_power)
 
 lemma two_hyperpow_ge_one [simp]: "(1::hypreal) \<le> 2 pow n"
-apply (rule_tac y = "1 pow n" in order_trans)
-apply (rule_tac [2] hyperpow_le, auto)
-done
+  apply (rule_tac y = "1 pow n" in order_trans)
+   apply (rule_tac [2] hyperpow_le)
+    apply auto
+  done
 
-lemma hyperpow_minus_one2 [simp]:
-     "\<And>n. (- 1) pow (2*n) = (1::hypreal)"
-by transfer (rule power_minus1_even)
+lemma hyperpow_minus_one2 [simp]: "\<And>n. (- 1) pow (2 * n) = (1::hypreal)"
+  by transfer (rule power_minus1_even)
 
-lemma hyperpow_less_le:
-     "!!r n N. [|(0::hypreal) \<le> r; r \<le> 1; n < N|] ==> r pow N \<le> r pow n"
-by transfer (rule power_decreasing [OF order_less_imp_le])
+lemma hyperpow_less_le: "\<And>r n N. (0::hypreal) \<le> r \<Longrightarrow> r \<le> 1 \<Longrightarrow> n < N \<Longrightarrow> r pow N \<le> r pow n"
+  by transfer (rule power_decreasing [OF order_less_imp_le])
 
 lemma hyperpow_SHNat_le:
-     "[| 0 \<le> r;  r \<le> (1::hypreal);  N \<in> HNatInfinite |]
-      ==> ALL n: Nats. r pow N \<le> r pow n"
-by (auto intro!: hyperpow_less_le simp add: HNatInfinite_iff)
+  "0 \<le> r \<Longrightarrow> r \<le> (1::hypreal) \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> \<forall>n\<in>Nats. r pow N \<le> r pow n"
+  by (auto intro!: hyperpow_less_le simp: HNatInfinite_iff)
 
-lemma hyperpow_realpow:
-      "(hypreal_of_real r) pow (hypnat_of_nat n) = hypreal_of_real (r ^ n)"
-by transfer (rule refl)
+lemma hyperpow_realpow: "(hypreal_of_real r) pow (hypnat_of_nat n) = hypreal_of_real (r ^ n)"
+  by transfer (rule refl)
 
-lemma hyperpow_SReal [simp]:
-     "(hypreal_of_real r) pow (hypnat_of_nat n) \<in> \<real>"
-by (simp add: Reals_eq_Standard)
+lemma hyperpow_SReal [simp]: "(hypreal_of_real r) pow (hypnat_of_nat n) \<in> \<real>"
+  by (simp add: Reals_eq_Standard)
 
-lemma hyperpow_zero_HNatInfinite [simp]:
-     "N \<in> HNatInfinite ==> (0::hypreal) pow N = 0"
-by (drule HNatInfinite_is_Suc, auto)
+lemma hyperpow_zero_HNatInfinite [simp]: "N \<in> HNatInfinite \<Longrightarrow> (0::hypreal) pow N = 0"
+  by (drule HNatInfinite_is_Suc, auto)
 
-lemma hyperpow_le_le:
-     "[| (0::hypreal) \<le> r; r \<le> 1; n \<le> N |] ==> r pow N \<le> r pow n"
-apply (drule order_le_less [of n, THEN iffD1])
-apply (auto intro: hyperpow_less_le)
-done
+lemma hyperpow_le_le: "(0::hypreal) \<le> r \<Longrightarrow> r \<le> 1 \<Longrightarrow> n \<le> N \<Longrightarrow> r pow N \<le> r pow n"
+  apply (drule order_le_less [of n, THEN iffD1])
+  apply (auto intro: hyperpow_less_le)
+  done
 
-lemma hyperpow_Suc_le_self2:
-     "[| (0::hypreal) \<le> r; r < 1 |] ==> r pow (n + (1::hypnat)) \<le> r"
-apply (drule_tac n = " (1::hypnat) " in hyperpow_le_le)
-apply auto
-done
+lemma hyperpow_Suc_le_self2: "(0::hypreal) \<le> r \<Longrightarrow> r < 1 \<Longrightarrow> r pow (n + (1::hypnat)) \<le> r"
+  apply (drule_tac n = " (1::hypnat) " in hyperpow_le_le)
+    apply auto
+  done
 
 lemma hyperpow_hypnat_of_nat: "\<And>x. x pow hypnat_of_nat n = x ^ n"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma of_hypreal_hyperpow:
-  "\<And>x n. of_hypreal (x pow n) =
-   (of_hypreal x::'a::{real_algebra_1} star) pow n"
-by transfer (rule of_real_power)
+  "\<And>x n. of_hypreal (x pow n) = (of_hypreal x::'a::{real_algebra_1} star) pow n"
+  by transfer (rule of_real_power)
 
 end
--- a/src/HOL/Nonstandard_Analysis/HyperNat.thy	Mon Oct 31 16:26:36 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/HyperNat.thy	Tue Nov 01 00:44:24 2016 +0100
@@ -5,410 +5,395 @@
 Converted to Isar and polished by lcp
 *)
 
-section\<open>Hypernatural numbers\<close>
+section \<open>Hypernatural numbers\<close>
 
 theory HyperNat
-imports StarDef
+  imports StarDef
 begin
 
 type_synonym hypnat = "nat star"
 
-abbreviation
-  hypnat_of_nat :: "nat => nat star" where
-  "hypnat_of_nat == star_of"
+abbreviation hypnat_of_nat :: "nat \<Rightarrow> nat star"
+  where "hypnat_of_nat \<equiv> star_of"
 
-definition
-  hSuc :: "hypnat => hypnat" where
-  hSuc_def [transfer_unfold]: "hSuc = *f* Suc"
+definition hSuc :: "hypnat \<Rightarrow> hypnat"
+  where hSuc_def [transfer_unfold]: "hSuc = *f* Suc"
 
-subsection\<open>Properties Transferred from Naturals\<close>
+
+subsection \<open>Properties Transferred from Naturals\<close>
 
 lemma hSuc_not_zero [iff]: "\<And>m. hSuc m \<noteq> 0"
-by transfer (rule Suc_not_Zero)
+  by transfer (rule Suc_not_Zero)
 
 lemma zero_not_hSuc [iff]: "\<And>m. 0 \<noteq> hSuc m"
-by transfer (rule Zero_not_Suc)
+  by transfer (rule Zero_not_Suc)
 
-lemma hSuc_hSuc_eq [iff]: "\<And>m n. (hSuc m = hSuc n) = (m = n)"
-by transfer (rule nat.inject)
+lemma hSuc_hSuc_eq [iff]: "\<And>m n. hSuc m = hSuc n \<longleftrightarrow> m = n"
+  by transfer (rule nat.inject)
 
 lemma zero_less_hSuc [iff]: "\<And>n. 0 < hSuc n"
-by transfer (rule zero_less_Suc)
+  by transfer (rule zero_less_Suc)
 
-lemma hypnat_minus_zero [simp]: "!!z. z - z = (0::hypnat)"
-by transfer (rule diff_self_eq_0)
+lemma hypnat_minus_zero [simp]: "\<And>z::hypnat. z - z = 0"
+  by transfer (rule diff_self_eq_0)
 
-lemma hypnat_diff_0_eq_0 [simp]: "!!n. (0::hypnat) - n = 0"
-by transfer (rule diff_0_eq_0)
+lemma hypnat_diff_0_eq_0 [simp]: "\<And>n::hypnat. 0 - n = 0"
+  by transfer (rule diff_0_eq_0)
 
-lemma hypnat_add_is_0 [iff]: "!!m n. (m+n = (0::hypnat)) = (m=0 & n=0)"
-by transfer (rule add_is_0)
+lemma hypnat_add_is_0 [iff]: "\<And>m n::hypnat. m + n = 0 \<longleftrightarrow> m = 0 \<and> n = 0"
+  by transfer (rule add_is_0)
 
-lemma hypnat_diff_diff_left: "!!i j k. (i::hypnat) - j - k = i - (j+k)"
-by transfer (rule diff_diff_left)
+lemma hypnat_diff_diff_left: "\<And>i j k::hypnat. i - j - k = i - (j + k)"
+  by transfer (rule diff_diff_left)
 
-lemma hypnat_diff_commute: "!!i j k. (i::hypnat) - j - k = i-k-j"
-by transfer (rule diff_commute)
+lemma hypnat_diff_commute: "\<And>i j k::hypnat. i - j - k = i - k - j"
+  by transfer (rule diff_commute)
 
-lemma hypnat_diff_add_inverse [simp]: "!!m n. ((n::hypnat) + m) - n = m"
-by transfer (rule diff_add_inverse)
+lemma hypnat_diff_add_inverse [simp]: "\<And>m n::hypnat. n + m - n = m"
+  by transfer (rule diff_add_inverse)
 
-lemma hypnat_diff_add_inverse2 [simp]:  "!!m n. ((m::hypnat) + n) - n = m"
-by transfer (rule diff_add_inverse2)
+lemma hypnat_diff_add_inverse2 [simp]:  "\<And>m n::hypnat. m + n - n = m"
+  by transfer (rule diff_add_inverse2)
 
-lemma hypnat_diff_cancel [simp]: "!!k m n. ((k::hypnat) + m) - (k+n) = m - n"
-by transfer (rule diff_cancel)
+lemma hypnat_diff_cancel [simp]: "\<And>k m n::hypnat. (k + m) - (k + n) = m - n"
+  by transfer (rule diff_cancel)
 
-lemma hypnat_diff_cancel2 [simp]: "!!k m n. ((m::hypnat) + k) - (n+k) = m - n"
-by transfer (rule diff_cancel2)
+lemma hypnat_diff_cancel2 [simp]: "\<And>k m n::hypnat. (m + k) - (n + k) = m - n"
+  by transfer (rule diff_cancel2)
 
-lemma hypnat_diff_add_0 [simp]: "!!m n. (n::hypnat) - (n+m) = (0::hypnat)"
-by transfer (rule diff_add_0)
+lemma hypnat_diff_add_0 [simp]: "\<And>m n::hypnat. n - (n + m) = 0"
+  by transfer (rule diff_add_0)
 
-lemma hypnat_diff_mult_distrib: "!!k m n. ((m::hypnat) - n) * k = (m * k) - (n * k)"
-by transfer (rule diff_mult_distrib)
+lemma hypnat_diff_mult_distrib: "\<And>k m n::hypnat. (m - n) * k = (m * k) - (n * k)"
+  by transfer (rule diff_mult_distrib)
 
-lemma hypnat_diff_mult_distrib2: "!!k m n. (k::hypnat) * (m - n) = (k * m) - (k * n)"
-by transfer (rule diff_mult_distrib2)
+lemma hypnat_diff_mult_distrib2: "\<And>k m n::hypnat. k * (m - n) = (k * m) - (k * n)"
+  by transfer (rule diff_mult_distrib2)
 
-lemma hypnat_le_zero_cancel [iff]: "!!n. (n \<le> (0::hypnat)) = (n = 0)"
-by transfer (rule le_0_eq)
+lemma hypnat_le_zero_cancel [iff]: "\<And>n::hypnat. n \<le> 0 \<longleftrightarrow> n = 0"
+  by transfer (rule le_0_eq)
 
-lemma hypnat_mult_is_0 [simp]: "!!m n. (m*n = (0::hypnat)) = (m=0 | n=0)"
-by transfer (rule mult_is_0)
+lemma hypnat_mult_is_0 [simp]: "\<And>m n::hypnat. m * n = 0 \<longleftrightarrow> m = 0 \<or> n = 0"
+  by transfer (rule mult_is_0)
 
-lemma hypnat_diff_is_0_eq [simp]: "!!m n. ((m::hypnat) - n = 0) = (m \<le> n)"
-by transfer (rule diff_is_0_eq)
+lemma hypnat_diff_is_0_eq [simp]: "\<And>m n::hypnat. m - n = 0 \<longleftrightarrow> m \<le> n"
+  by transfer (rule diff_is_0_eq)
 
-lemma hypnat_not_less0 [iff]: "!!n. ~ n < (0::hypnat)"
-by transfer (rule not_less0)
+lemma hypnat_not_less0 [iff]: "\<And>n::hypnat. \<not> n < 0"
+  by transfer (rule not_less0)
+
+lemma hypnat_less_one [iff]: "\<And>n::hypnat. n < 1 \<longleftrightarrow> n = 0"
+  by transfer (rule less_one)
 
-lemma hypnat_less_one [iff]:
-      "!!n. (n < (1::hypnat)) = (n=0)"
-by transfer (rule less_one)
+lemma hypnat_add_diff_inverse: "\<And>m n::hypnat. \<not> m < n \<Longrightarrow> n + (m - n) = m"
+  by transfer (rule add_diff_inverse)
 
-lemma hypnat_add_diff_inverse: "!!m n. ~ m<n ==> n+(m-n) = (m::hypnat)"
-by transfer (rule add_diff_inverse)
+lemma hypnat_le_add_diff_inverse [simp]: "\<And>m n::hypnat. n \<le> m \<Longrightarrow> n + (m - n) = m"
+  by transfer (rule le_add_diff_inverse)
 
-lemma hypnat_le_add_diff_inverse [simp]: "!!m n. n \<le> m ==> n+(m-n) = (m::hypnat)"
-by transfer (rule le_add_diff_inverse)
-
-lemma hypnat_le_add_diff_inverse2 [simp]: "!!m n. n\<le>m ==> (m-n)+n = (m::hypnat)"
-by transfer (rule le_add_diff_inverse2)
+lemma hypnat_le_add_diff_inverse2 [simp]: "\<And>m n::hypnat. n \<le> m \<Longrightarrow> (m - n) + n = m"
+  by transfer (rule le_add_diff_inverse2)
 
 declare hypnat_le_add_diff_inverse2 [OF order_less_imp_le]
 
-lemma hypnat_le0 [iff]: "!!n. (0::hypnat) \<le> n"
-by transfer (rule le0)
+lemma hypnat_le0 [iff]: "\<And>n::hypnat. 0 \<le> n"
+  by transfer (rule le0)
 
-lemma hypnat_le_add1 [simp]: "!!x n. (x::hypnat) \<le> x + n"
-by transfer (rule le_add1)
+lemma hypnat_le_add1 [simp]: "\<And>x n::hypnat. x \<le> x + n"
+  by transfer (rule le_add1)
 
-lemma hypnat_add_self_le [simp]: "!!x n. (x::hypnat) \<le> n + x"
-by transfer (rule le_add2)
+lemma hypnat_add_self_le [simp]: "\<And>x n::hypnat. x \<le> n + x"
+  by transfer (rule le_add2)
 
-lemma hypnat_add_one_self_less [simp]: "(x::hypnat) < x + (1::hypnat)"
+lemma hypnat_add_one_self_less [simp]: "x < x + 1" for x :: hypnat
   by (fact less_add_one)
 
-lemma hypnat_neq0_conv [iff]: "!!n. (n \<noteq> 0) = (0 < (n::hypnat))"
-by transfer (rule neq0_conv)
+lemma hypnat_neq0_conv [iff]: "\<And>n::hypnat. n \<noteq> 0 \<longleftrightarrow> 0 < n"
+  by transfer (rule neq0_conv)
 
-lemma hypnat_gt_zero_iff: "((0::hypnat) < n) = ((1::hypnat) \<le> n)"
-by (auto simp add: linorder_not_less [symmetric])
+lemma hypnat_gt_zero_iff: "0 < n \<longleftrightarrow> 1 \<le> n" for n :: hypnat
+  by (auto simp add: linorder_not_less [symmetric])
 
-lemma hypnat_gt_zero_iff2: "(0 < n) = (\<exists>m. n = m + (1::hypnat))"
+lemma hypnat_gt_zero_iff2: "0 < n \<longleftrightarrow> (\<exists>m. n = m + 1)" for n :: hypnat
   by (auto intro!: add_nonneg_pos exI[of _ "n - 1"] simp: hypnat_gt_zero_iff)
 
-lemma hypnat_add_self_not_less: "~ (x + y < (x::hypnat))"
-by (simp add: linorder_not_le [symmetric] add.commute [of x])
+lemma hypnat_add_self_not_less: "\<not> x + y < x" for x y :: hypnat
+  by (simp add: linorder_not_le [symmetric] add.commute [of x])
 
-lemma hypnat_diff_split:
-    "P(a - b::hypnat) = ((a<b --> P 0) & (ALL d. a = b + d --> P d))"
-    \<comment> \<open>elimination of \<open>-\<close> on \<open>hypnat\<close>\<close>
-proof (cases "a<b" rule: case_split)
+lemma hypnat_diff_split: "P (a - b) \<longleftrightarrow> (a < b \<longrightarrow> P 0) \<and> (\<forall>d. a = b + d \<longrightarrow> P d)"
+  for a b :: hypnat
+  \<comment> \<open>elimination of \<open>-\<close> on \<open>hypnat\<close>\<close>
+proof (cases "a < b" rule: case_split)
   case True
-    thus ?thesis
-      by (auto simp add: hypnat_add_self_not_less order_less_imp_le
-                         hypnat_diff_is_0_eq [THEN iffD2])
+  then show ?thesis
+    by (auto simp add: hypnat_add_self_not_less order_less_imp_le hypnat_diff_is_0_eq [THEN iffD2])
 next
   case False
-    thus ?thesis
-      by (auto simp add: linorder_not_less dest: order_le_less_trans)
+  then show ?thesis
+    by (auto simp add: linorder_not_less dest: order_le_less_trans)
 qed
 
-subsection\<open>Properties of the set of embedded natural numbers\<close>
+
+subsection \<open>Properties of the set of embedded natural numbers\<close>
 
 lemma of_nat_eq_star_of [simp]: "of_nat = star_of"
 proof
-  fix n :: nat
-  show "of_nat n = star_of n" by transfer simp
+  show "of_nat n = star_of n" for n
+    by transfer simp
 qed
 
 lemma Nats_eq_Standard: "(Nats :: nat star set) = Standard"
-by (auto simp add: Nats_def Standard_def)
+  by (auto simp: Nats_def Standard_def)
 
 lemma hypnat_of_nat_mem_Nats [simp]: "hypnat_of_nat n \<in> Nats"
-by (simp add: Nats_eq_Standard)
+  by (simp add: Nats_eq_Standard)
 
-lemma hypnat_of_nat_one [simp]: "hypnat_of_nat (Suc 0) = (1::hypnat)"
-by transfer simp
+lemma hypnat_of_nat_one [simp]: "hypnat_of_nat (Suc 0) = 1"
+  by transfer simp
 
-lemma hypnat_of_nat_Suc [simp]:
-     "hypnat_of_nat (Suc n) = hypnat_of_nat n + (1::hypnat)"
-by transfer simp
+lemma hypnat_of_nat_Suc [simp]: "hypnat_of_nat (Suc n) = hypnat_of_nat n + 1"
+  by transfer simp
 
-lemma of_nat_eq_add [rule_format]:
-     "\<forall>d::hypnat. of_nat m = of_nat n + d --> d \<in> range of_nat"
-apply (induct n)
-apply (auto simp add: add.assoc)
-apply (case_tac x)
-apply (auto simp add: add.commute [of 1])
-done
+lemma of_nat_eq_add [rule_format]: "\<forall>d::hypnat. of_nat m = of_nat n + d --> d \<in> range of_nat"
+  apply (induct n)
+   apply (auto simp add: add.assoc)
+  apply (case_tac x)
+   apply (auto simp add: add.commute [of 1])
+  done
 
-lemma Nats_diff [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> (a-b :: hypnat) \<in> Nats"
-by (simp add: Nats_eq_Standard)
+lemma Nats_diff [simp]: "a \<in> Nats \<Longrightarrow> b \<in> Nats \<Longrightarrow> a - b \<in> Nats" for a b :: hypnat
+  by (simp add: Nats_eq_Standard)
 
 
-subsection\<open>Infinite Hypernatural Numbers -- @{term HNatInfinite}\<close>
+subsection \<open>Infinite Hypernatural Numbers -- @{term HNatInfinite}\<close>
+
+text \<open>The set of infinite hypernatural numbers.\<close>
+definition HNatInfinite :: "hypnat set"
+  where "HNatInfinite = {n. n \<notin> Nats}"
 
-definition
-  (* the set of infinite hypernatural numbers *)
-  HNatInfinite :: "hypnat set" where
-  "HNatInfinite = {n. n \<notin> Nats}"
+lemma Nats_not_HNatInfinite_iff: "x \<in> Nats \<longleftrightarrow> x \<notin> HNatInfinite"
+  by (simp add: HNatInfinite_def)
 
-lemma Nats_not_HNatInfinite_iff: "(x \<in> Nats) = (x \<notin> HNatInfinite)"
-by (simp add: HNatInfinite_def)
-
-lemma HNatInfinite_not_Nats_iff: "(x \<in> HNatInfinite) = (x \<notin> Nats)"
-by (simp add: HNatInfinite_def)
+lemma HNatInfinite_not_Nats_iff: "x \<in> HNatInfinite \<longleftrightarrow> x \<notin> Nats"
+  by (simp add: HNatInfinite_def)
 
 lemma star_of_neq_HNatInfinite: "N \<in> HNatInfinite \<Longrightarrow> star_of n \<noteq> N"
-by (auto simp add: HNatInfinite_def Nats_eq_Standard)
+  by (auto simp add: HNatInfinite_def Nats_eq_Standard)
 
-lemma star_of_Suc_lessI:
-  "\<And>N. \<lbrakk>star_of n < N; star_of (Suc n) \<noteq> N\<rbrakk> \<Longrightarrow> star_of (Suc n) < N"
-by transfer (rule Suc_lessI)
+lemma star_of_Suc_lessI: "\<And>N. star_of n < N \<Longrightarrow> star_of (Suc n) \<noteq> N \<Longrightarrow> star_of (Suc n) < N"
+  by transfer (rule Suc_lessI)
 
 lemma star_of_less_HNatInfinite:
   assumes N: "N \<in> HNatInfinite"
   shows "star_of n < N"
 proof (induct n)
   case 0
-  from N have "star_of 0 \<noteq> N" by (rule star_of_neq_HNatInfinite)
-  thus "star_of 0 < N" by simp
+  from N have "star_of 0 \<noteq> N"
+    by (rule star_of_neq_HNatInfinite)
+  then show ?case by simp
 next
   case (Suc n)
-  from N have "star_of (Suc n) \<noteq> N" by (rule star_of_neq_HNatInfinite)
-  with Suc show "star_of (Suc n) < N" by (rule star_of_Suc_lessI)
+  from N have "star_of (Suc n) \<noteq> N"
+    by (rule star_of_neq_HNatInfinite)
+  with Suc show ?case
+    by (rule star_of_Suc_lessI)
 qed
 
 lemma star_of_le_HNatInfinite: "N \<in> HNatInfinite \<Longrightarrow> star_of n \<le> N"
-by (rule star_of_less_HNatInfinite [THEN order_less_imp_le])
+  by (rule star_of_less_HNatInfinite [THEN order_less_imp_le])
+
 
 subsubsection \<open>Closure Rules\<close>
 
-lemma Nats_less_HNatInfinite: "\<lbrakk>x \<in> Nats; y \<in> HNatInfinite\<rbrakk> \<Longrightarrow> x < y"
-by (auto simp add: Nats_def star_of_less_HNatInfinite)
+lemma Nats_less_HNatInfinite: "x \<in> Nats \<Longrightarrow> y \<in> HNatInfinite \<Longrightarrow> x < y"
+  by (auto simp add: Nats_def star_of_less_HNatInfinite)
 
-lemma Nats_le_HNatInfinite: "\<lbrakk>x \<in> Nats; y \<in> HNatInfinite\<rbrakk> \<Longrightarrow> x \<le> y"
-by (rule Nats_less_HNatInfinite [THEN order_less_imp_le])
+lemma Nats_le_HNatInfinite: "x \<in> Nats \<Longrightarrow> y \<in> HNatInfinite \<Longrightarrow> x \<le> y"
+  by (rule Nats_less_HNatInfinite [THEN order_less_imp_le])
 
 lemma zero_less_HNatInfinite: "x \<in> HNatInfinite \<Longrightarrow> 0 < x"
-by (simp add: Nats_less_HNatInfinite)
+  by (simp add: Nats_less_HNatInfinite)
 
 lemma one_less_HNatInfinite: "x \<in> HNatInfinite \<Longrightarrow> 1 < x"
-by (simp add: Nats_less_HNatInfinite)
+  by (simp add: Nats_less_HNatInfinite)
 
 lemma one_le_HNatInfinite: "x \<in> HNatInfinite \<Longrightarrow> 1 \<le> x"
-by (simp add: Nats_le_HNatInfinite)
+  by (simp add: Nats_le_HNatInfinite)
 
 lemma zero_not_mem_HNatInfinite [simp]: "0 \<notin> HNatInfinite"
-by (simp add: HNatInfinite_def)
+  by (simp add: HNatInfinite_def)
 
-lemma Nats_downward_closed:
-  "\<lbrakk>x \<in> Nats; (y::hypnat) \<le> x\<rbrakk> \<Longrightarrow> y \<in> Nats"
-apply (simp only: linorder_not_less [symmetric])
-apply (erule contrapos_np)
-apply (drule HNatInfinite_not_Nats_iff [THEN iffD2])
-apply (erule (1) Nats_less_HNatInfinite)
-done
+lemma Nats_downward_closed: "x \<in> Nats \<Longrightarrow> y \<le> x \<Longrightarrow> y \<in> Nats" for x y :: hypnat
+  apply (simp only: linorder_not_less [symmetric])
+  apply (erule contrapos_np)
+  apply (drule HNatInfinite_not_Nats_iff [THEN iffD2])
+  apply (erule (1) Nats_less_HNatInfinite)
+  done
 
-lemma HNatInfinite_upward_closed:
-  "\<lbrakk>x \<in> HNatInfinite; x \<le> y\<rbrakk> \<Longrightarrow> y \<in> HNatInfinite"
-apply (simp only: HNatInfinite_not_Nats_iff)
-apply (erule contrapos_nn)
-apply (erule (1) Nats_downward_closed)
-done
+lemma HNatInfinite_upward_closed: "x \<in> HNatInfinite \<Longrightarrow> x \<le> y \<Longrightarrow> y \<in> HNatInfinite"
+  apply (simp only: HNatInfinite_not_Nats_iff)
+  apply (erule contrapos_nn)
+  apply (erule (1) Nats_downward_closed)
+  done
 
 lemma HNatInfinite_add: "x \<in> HNatInfinite \<Longrightarrow> x + y \<in> HNatInfinite"
-apply (erule HNatInfinite_upward_closed)
-apply (rule hypnat_le_add1)
-done
+  apply (erule HNatInfinite_upward_closed)
+  apply (rule hypnat_le_add1)
+  done
 
 lemma HNatInfinite_add_one: "x \<in> HNatInfinite \<Longrightarrow> x + 1 \<in> HNatInfinite"
-by (rule HNatInfinite_add)
+  by (rule HNatInfinite_add)
 
-lemma HNatInfinite_diff:
-  "\<lbrakk>x \<in> HNatInfinite; y \<in> Nats\<rbrakk> \<Longrightarrow> x - y \<in> HNatInfinite"
-apply (frule (1) Nats_le_HNatInfinite)
-apply (simp only: HNatInfinite_not_Nats_iff)
-apply (erule contrapos_nn)
-apply (drule (1) Nats_add, simp)
-done
+lemma HNatInfinite_diff: "x \<in> HNatInfinite \<Longrightarrow> y \<in> Nats \<Longrightarrow> x - y \<in> HNatInfinite"
+  apply (frule (1) Nats_le_HNatInfinite)
+  apply (simp only: HNatInfinite_not_Nats_iff)
+  apply (erule contrapos_nn)
+  apply (drule (1) Nats_add, simp)
+  done
 
-lemma HNatInfinite_is_Suc: "x \<in> HNatInfinite ==> \<exists>y. x = y + (1::hypnat)"
-apply (rule_tac x = "x - (1::hypnat) " in exI)
-apply (simp add: Nats_le_HNatInfinite)
-done
+lemma HNatInfinite_is_Suc: "x \<in> HNatInfinite \<Longrightarrow> \<exists>y. x = y + 1" for x :: hypnat
+  apply (rule_tac x = "x - (1::hypnat) " in exI)
+  apply (simp add: Nats_le_HNatInfinite)
+  done
 
 
-subsection\<open>Existence of an infinite hypernatural number\<close>
+subsection \<open>Existence of an infinite hypernatural number\<close>
 
-definition
-  (* \<omega> is in fact an infinite hypernatural number = [<1,2,3,...>] *)
-  whn :: hypnat where
-  hypnat_omega_def: "whn = star_n (%n::nat. n)"
+text \<open>\<open>\<omega>\<close> is in fact an infinite hypernatural number = \<open>[<1, 2, 3, \<dots>>]\<close>\<close>
+definition whn :: hypnat
+  where hypnat_omega_def: "whn = star_n (\<lambda>n::nat. n)"
 
 lemma hypnat_of_nat_neq_whn: "hypnat_of_nat n \<noteq> whn"
-by (simp add: FreeUltrafilterNat.singleton' hypnat_omega_def star_of_def star_n_eq_iff)
+  by (simp add: FreeUltrafilterNat.singleton' hypnat_omega_def star_of_def star_n_eq_iff)
 
 lemma whn_neq_hypnat_of_nat: "whn \<noteq> hypnat_of_nat n"
-by (simp add: FreeUltrafilterNat.singleton hypnat_omega_def star_of_def star_n_eq_iff)
+  by (simp add: FreeUltrafilterNat.singleton hypnat_omega_def star_of_def star_n_eq_iff)
 
 lemma whn_not_Nats [simp]: "whn \<notin> Nats"
-by (simp add: Nats_def image_def whn_neq_hypnat_of_nat)
+  by (simp add: Nats_def image_def whn_neq_hypnat_of_nat)
 
 lemma HNatInfinite_whn [simp]: "whn \<in> HNatInfinite"
-by (simp add: HNatInfinite_def)
+  by (simp add: HNatInfinite_def)
 
 lemma lemma_unbounded_set [simp]: "eventually (\<lambda>n::nat. m < n) \<U>"
   by (rule filter_leD[OF FreeUltrafilterNat.le_cofinite])
      (auto simp add: cofinite_eq_sequentially eventually_at_top_dense)
 
-lemma hypnat_of_nat_eq:
-     "hypnat_of_nat m  = star_n (%n::nat. m)"
-by (simp add: star_of_def)
+lemma hypnat_of_nat_eq: "hypnat_of_nat m  = star_n (\<lambda>n::nat. m)"
+  by (simp add: star_of_def)
 
 lemma SHNat_eq: "Nats = {n. \<exists>N. n = hypnat_of_nat N}"
-by (simp add: Nats_def image_def)
+  by (simp add: Nats_def image_def)
 
 lemma Nats_less_whn: "n \<in> Nats \<Longrightarrow> n < whn"
-by (simp add: Nats_less_HNatInfinite)
+  by (simp add: Nats_less_HNatInfinite)
 
 lemma Nats_le_whn: "n \<in> Nats \<Longrightarrow> n \<le> whn"
-by (simp add: Nats_le_HNatInfinite)
+  by (simp add: Nats_le_HNatInfinite)
 
 lemma hypnat_of_nat_less_whn [simp]: "hypnat_of_nat n < whn"
-by (simp add: Nats_less_whn)
+  by (simp add: Nats_less_whn)
 
 lemma hypnat_of_nat_le_whn [simp]: "hypnat_of_nat n \<le> whn"
-by (simp add: Nats_le_whn)
+  by (simp add: Nats_le_whn)
 
 lemma hypnat_zero_less_hypnat_omega [simp]: "0 < whn"
-by (simp add: Nats_less_whn)
+  by (simp add: Nats_less_whn)
 
 lemma hypnat_one_less_hypnat_omega [simp]: "1 < whn"
-by (simp add: Nats_less_whn)
+  by (simp add: Nats_less_whn)
 
 
-subsubsection\<open>Alternative characterization of the set of infinite hypernaturals\<close>
+subsubsection \<open>Alternative characterization of the set of infinite hypernaturals\<close>
 
-text\<open>@{term "HNatInfinite = {N. \<forall>n \<in> Nats. n < N}"}\<close>
+text \<open>@{term "HNatInfinite = {N. \<forall>n \<in> Nats. n < N}"}\<close>
 
 (*??delete? similar reasoning in hypnat_omega_gt_SHNat above*)
 lemma HNatInfinite_FreeUltrafilterNat_lemma:
   assumes "\<forall>N::nat. eventually (\<lambda>n. f n \<noteq> N) \<U>"
   shows "eventually (\<lambda>n. N < f n) \<U>"
-apply (induct N)
-using assms
-apply (drule_tac x = 0 in spec, simp)
-using assms
-apply (drule_tac x = "Suc N" in spec)
-apply (auto elim: eventually_elim2)
-done
+  apply (induct N)
+  using assms
+   apply (drule_tac x = 0 in spec, simp)
+  using assms
+  apply (drule_tac x = "Suc N" in spec)
+  apply (auto elim: eventually_elim2)
+  done
 
 lemma HNatInfinite_iff: "HNatInfinite = {N. \<forall>n \<in> Nats. n < N}"
-apply (safe intro!: Nats_less_HNatInfinite)
-apply (auto simp add: HNatInfinite_def)
-done
+  apply (safe intro!: Nats_less_HNatInfinite)
+  apply (auto simp add: HNatInfinite_def)
+  done
 
 
-subsubsection\<open>Alternative Characterization of @{term HNatInfinite} using
-Free Ultrafilter\<close>
+subsubsection \<open>Alternative Characterization of @{term HNatInfinite} using Free Ultrafilter\<close>
 
 lemma HNatInfinite_FreeUltrafilterNat:
-     "star_n X \<in> HNatInfinite ==> \<forall>u. eventually (\<lambda>n. u < X n) FreeUltrafilterNat"
-apply (auto simp add: HNatInfinite_iff SHNat_eq)
-apply (drule_tac x="star_of u" in spec, simp)
-apply (simp add: star_of_def star_less_def starP2_star_n)
-done
+  "star_n X \<in> HNatInfinite \<Longrightarrow> \<forall>u. eventually (\<lambda>n. u < X n) FreeUltrafilterNat"
+  apply (auto simp add: HNatInfinite_iff SHNat_eq)
+  apply (drule_tac x="star_of u" in spec, simp)
+  apply (simp add: star_of_def star_less_def starP2_star_n)
+  done
 
 lemma FreeUltrafilterNat_HNatInfinite:
-     "\<forall>u. eventually (\<lambda>n. u < X n) FreeUltrafilterNat ==> star_n X \<in> HNatInfinite"
-by (auto simp add: star_less_def starP2_star_n HNatInfinite_iff SHNat_eq hypnat_of_nat_eq)
+  "\<forall>u. eventually (\<lambda>n. u < X n) FreeUltrafilterNat \<Longrightarrow> star_n X \<in> HNatInfinite"
+  by (auto simp add: star_less_def starP2_star_n HNatInfinite_iff SHNat_eq hypnat_of_nat_eq)
 
 lemma HNatInfinite_FreeUltrafilterNat_iff:
-     "(star_n X \<in> HNatInfinite) = (\<forall>u. eventually (\<lambda>n. u < X n) FreeUltrafilterNat)"
-by (rule iffI [OF HNatInfinite_FreeUltrafilterNat
-                 FreeUltrafilterNat_HNatInfinite])
+  "(star_n X \<in> HNatInfinite) = (\<forall>u. eventually (\<lambda>n. u < X n) FreeUltrafilterNat)"
+  by (rule iffI [OF HNatInfinite_FreeUltrafilterNat FreeUltrafilterNat_HNatInfinite])
+
 
 subsection \<open>Embedding of the Hypernaturals into other types\<close>
 
-definition
-  of_hypnat :: "hypnat \<Rightarrow> 'a::semiring_1_cancel star" where
-  of_hypnat_def [transfer_unfold]: "of_hypnat = *f* of_nat"
+definition of_hypnat :: "hypnat \<Rightarrow> 'a::semiring_1_cancel star"
+  where of_hypnat_def [transfer_unfold]: "of_hypnat = *f* of_nat"
 
 lemma of_hypnat_0 [simp]: "of_hypnat 0 = 0"
-by transfer (rule of_nat_0)
+  by transfer (rule of_nat_0)
 
 lemma of_hypnat_1 [simp]: "of_hypnat 1 = 1"
-by transfer (rule of_nat_1)
+  by transfer (rule of_nat_1)
 
 lemma of_hypnat_hSuc: "\<And>m. of_hypnat (hSuc m) = 1 + of_hypnat m"
-by transfer (rule of_nat_Suc)
+  by transfer (rule of_nat_Suc)
 
-lemma of_hypnat_add [simp]:
-  "\<And>m n. of_hypnat (m + n) = of_hypnat m + of_hypnat n"
-by transfer (rule of_nat_add)
+lemma of_hypnat_add [simp]: "\<And>m n. of_hypnat (m + n) = of_hypnat m + of_hypnat n"
+  by transfer (rule of_nat_add)
 
-lemma of_hypnat_mult [simp]:
-  "\<And>m n. of_hypnat (m * n) = of_hypnat m * of_hypnat n"
-by transfer (rule of_nat_mult)
+lemma of_hypnat_mult [simp]: "\<And>m n. of_hypnat (m * n) = of_hypnat m * of_hypnat n"
+  by transfer (rule of_nat_mult)
 
 lemma of_hypnat_less_iff [simp]:
-  "\<And>m n. (of_hypnat m < (of_hypnat n::'a::linordered_semidom star)) = (m < n)"
-by transfer (rule of_nat_less_iff)
+  "\<And>m n. of_hypnat m < (of_hypnat n::'a::linordered_semidom star) \<longleftrightarrow> m < n"
+  by transfer (rule of_nat_less_iff)
 
 lemma of_hypnat_0_less_iff [simp]:
-  "\<And>n. (0 < (of_hypnat n::'a::linordered_semidom star)) = (0 < n)"
-by transfer (rule of_nat_0_less_iff)
+  "\<And>n. 0 < (of_hypnat n::'a::linordered_semidom star) \<longleftrightarrow> 0 < n"
+  by transfer (rule of_nat_0_less_iff)
 
-lemma of_hypnat_less_0_iff [simp]:
-  "\<And>m. \<not> (of_hypnat m::'a::linordered_semidom star) < 0"
-by transfer (rule of_nat_less_0_iff)
+lemma of_hypnat_less_0_iff [simp]: "\<And>m. \<not> (of_hypnat m::'a::linordered_semidom star) < 0"
+  by transfer (rule of_nat_less_0_iff)
 
 lemma of_hypnat_le_iff [simp]:
-  "\<And>m n. (of_hypnat m \<le> (of_hypnat n::'a::linordered_semidom star)) = (m \<le> n)"
-by transfer (rule of_nat_le_iff)
+  "\<And>m n. of_hypnat m \<le> (of_hypnat n::'a::linordered_semidom star) \<longleftrightarrow> m \<le> n"
+  by transfer (rule of_nat_le_iff)
 
-lemma of_hypnat_0_le_iff [simp]:
-  "\<And>n. 0 \<le> (of_hypnat n::'a::linordered_semidom star)"
-by transfer (rule of_nat_0_le_iff)
+lemma of_hypnat_0_le_iff [simp]: "\<And>n. 0 \<le> (of_hypnat n::'a::linordered_semidom star)"
+  by transfer (rule of_nat_0_le_iff)
 
-lemma of_hypnat_le_0_iff [simp]:
-  "\<And>m. ((of_hypnat m::'a::linordered_semidom star) \<le> 0) = (m = 0)"
-by transfer (rule of_nat_le_0_iff)
+lemma of_hypnat_le_0_iff [simp]: "\<And>m. (of_hypnat m::'a::linordered_semidom star) \<le> 0 \<longleftrightarrow> m = 0"
+  by transfer (rule of_nat_le_0_iff)
 
 lemma of_hypnat_eq_iff [simp]:
-  "\<And>m n. (of_hypnat m = (of_hypnat n::'a::linordered_semidom star)) = (m = n)"
-by transfer (rule of_nat_eq_iff)
+  "\<And>m n. of_hypnat m = (of_hypnat n::'a::linordered_semidom star) \<longleftrightarrow> m = n"
+  by transfer (rule of_nat_eq_iff)
 
-lemma of_hypnat_eq_0_iff [simp]:
-  "\<And>m. ((of_hypnat m::'a::linordered_semidom star) = 0) = (m = 0)"
-by transfer (rule of_nat_eq_0_iff)
+lemma of_hypnat_eq_0_iff [simp]: "\<And>m. (of_hypnat m::'a::linordered_semidom star) = 0 \<longleftrightarrow> m = 0"
+  by transfer (rule of_nat_eq_0_iff)
 
 lemma HNatInfinite_of_hypnat_gt_zero:
   "N \<in> HNatInfinite \<Longrightarrow> (0::'a::linordered_semidom star) < of_hypnat N"
-by (rule ccontr, simp add: linorder_not_less)
+  by (rule ccontr) (simp add: linorder_not_less)
 
 end
--- a/src/HOL/Nonstandard_Analysis/NSA.thy	Mon Oct 31 16:26:36 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/NSA.thy	Tue Nov 01 00:44:24 2016 +0100
@@ -3,649 +3,584 @@
     Author:     Lawrence C Paulson, University of Cambridge
 *)
 
-section\<open>Infinite Numbers, Infinitesimals, Infinitely Close Relation\<close>
+section \<open>Infinite Numbers, Infinitesimals, Infinitely Close Relation\<close>
 
 theory NSA
-imports HyperDef "~~/src/HOL/Library/Lub_Glb"
+  imports HyperDef "~~/src/HOL/Library/Lub_Glb"
 begin
 
-definition
-  hnorm :: "'a::real_normed_vector star \<Rightarrow> real star" where
-  [transfer_unfold]: "hnorm = *f* norm"
+definition hnorm :: "'a::real_normed_vector star \<Rightarrow> real star"
+  where [transfer_unfold]: "hnorm = *f* norm"
 
-definition
-  Infinitesimal  :: "('a::real_normed_vector) star set" where
-  "Infinitesimal = {x. \<forall>r \<in> Reals. 0 < r --> hnorm x < r}"
+definition Infinitesimal  :: "('a::real_normed_vector) star set"
+  where "Infinitesimal = {x. \<forall>r \<in> Reals. 0 < r \<longrightarrow> hnorm x < r}"
 
-definition
-  HFinite :: "('a::real_normed_vector) star set" where
-  "HFinite = {x. \<exists>r \<in> Reals. hnorm x < r}"
+definition HFinite :: "('a::real_normed_vector) star set"
+  where "HFinite = {x. \<exists>r \<in> Reals. hnorm x < r}"
 
-definition
-  HInfinite :: "('a::real_normed_vector) star set" where
-  "HInfinite = {x. \<forall>r \<in> Reals. r < hnorm x}"
+definition HInfinite :: "('a::real_normed_vector) star set"
+  where "HInfinite = {x. \<forall>r \<in> Reals. r < hnorm x}"
 
-definition
-  approx :: "['a::real_normed_vector star, 'a star] => bool"  (infixl "\<approx>" 50) where
-    \<comment>\<open>the `infinitely close' relation\<close>
-  "(x \<approx> y) = ((x - y) \<in> Infinitesimal)"
+definition approx :: "'a::real_normed_vector star \<Rightarrow> 'a star \<Rightarrow> bool"  (infixl "\<approx>" 50)
+  where "x \<approx> y \<longleftrightarrow> x - y \<in> Infinitesimal"
+    \<comment> \<open>the ``infinitely close'' relation\<close>
 
-definition
-  st        :: "hypreal => hypreal" where
-    \<comment>\<open>the standard part of a hyperreal\<close>
-  "st = (%x. @r. x \<in> HFinite & r \<in> \<real> & r \<approx> x)"
+definition st :: "hypreal \<Rightarrow> hypreal"
+  where "st = (\<lambda>x. SOME r. x \<in> HFinite \<and> r \<in> \<real> \<and> r \<approx> x)"
+    \<comment> \<open>the standard part of a hyperreal\<close>
 
-definition
-  monad     :: "'a::real_normed_vector star => 'a star set" where
-  "monad x = {y. x \<approx> y}"
+definition monad :: "'a::real_normed_vector star \<Rightarrow> 'a star set"
+  where "monad x = {y. x \<approx> y}"
 
-definition
-  galaxy    :: "'a::real_normed_vector star => 'a star set" where
-  "galaxy x = {y. (x + -y) \<in> HFinite}"
+definition galaxy :: "'a::real_normed_vector star \<Rightarrow> 'a star set"
+  where "galaxy x = {y. (x + -y) \<in> HFinite}"
 
-lemma SReal_def: "\<real> == {x. \<exists>r. x = hypreal_of_real r}"
-by (simp add: Reals_def image_def)
+lemma SReal_def: "\<real> \<equiv> {x. \<exists>r. x = hypreal_of_real r}"
+  by (simp add: Reals_def image_def)
+
 
 subsection \<open>Nonstandard Extension of the Norm Function\<close>
 
-definition
-  scaleHR :: "real star \<Rightarrow> 'a star \<Rightarrow> 'a::real_normed_vector star" where
-  [transfer_unfold]: "scaleHR = starfun2 scaleR"
+definition scaleHR :: "real star \<Rightarrow> 'a star \<Rightarrow> 'a::real_normed_vector star"
+  where [transfer_unfold]: "scaleHR = starfun2 scaleR"
 
 lemma Standard_hnorm [simp]: "x \<in> Standard \<Longrightarrow> hnorm x \<in> Standard"
-by (simp add: hnorm_def)
+  by (simp add: hnorm_def)
 
 lemma star_of_norm [simp]: "star_of (norm x) = hnorm (star_of x)"
-by transfer (rule refl)
+  by transfer (rule refl)
 
-lemma hnorm_ge_zero [simp]:
-  "\<And>x::'a::real_normed_vector star. 0 \<le> hnorm x"
-by transfer (rule norm_ge_zero)
+lemma hnorm_ge_zero [simp]: "\<And>x::'a::real_normed_vector star. 0 \<le> hnorm x"
+  by transfer (rule norm_ge_zero)
 
-lemma hnorm_eq_zero [simp]:
-  "\<And>x::'a::real_normed_vector star. (hnorm x = 0) = (x = 0)"
-by transfer (rule norm_eq_zero)
+lemma hnorm_eq_zero [simp]: "\<And>x::'a::real_normed_vector star. hnorm x = 0 \<longleftrightarrow> x = 0"
+  by transfer (rule norm_eq_zero)
 
-lemma hnorm_triangle_ineq:
-  "\<And>x y::'a::real_normed_vector star. hnorm (x + y) \<le> hnorm x + hnorm y"
-by transfer (rule norm_triangle_ineq)
-
-lemma hnorm_triangle_ineq3:
-  "\<And>x y::'a::real_normed_vector star. \<bar>hnorm x - hnorm y\<bar> \<le> hnorm (x - y)"
-by transfer (rule norm_triangle_ineq3)
+lemma hnorm_triangle_ineq: "\<And>x y::'a::real_normed_vector star. hnorm (x + y) \<le> hnorm x + hnorm y"
+  by transfer (rule norm_triangle_ineq)
 
-lemma hnorm_scaleR:
-  "\<And>x::'a::real_normed_vector star.
-   hnorm (a *\<^sub>R x) = \<bar>star_of a\<bar> * hnorm x"
-by transfer (rule norm_scaleR)
+lemma hnorm_triangle_ineq3: "\<And>x y::'a::real_normed_vector star. \<bar>hnorm x - hnorm y\<bar> \<le> hnorm (x - y)"
+  by transfer (rule norm_triangle_ineq3)
+
+lemma hnorm_scaleR: "\<And>x::'a::real_normed_vector star. hnorm (a *\<^sub>R x) = \<bar>star_of a\<bar> * hnorm x"
+  by transfer (rule norm_scaleR)
 
-lemma hnorm_scaleHR:
-  "\<And>a (x::'a::real_normed_vector star).
-   hnorm (scaleHR a x) = \<bar>a\<bar> * hnorm x"
-by transfer (rule norm_scaleR)
+lemma hnorm_scaleHR: "\<And>a (x::'a::real_normed_vector star). hnorm (scaleHR a x) = \<bar>a\<bar> * hnorm x"
+  by transfer (rule norm_scaleR)
 
-lemma hnorm_mult_ineq:
-  "\<And>x y::'a::real_normed_algebra star. hnorm (x * y) \<le> hnorm x * hnorm y"
-by transfer (rule norm_mult_ineq)
+lemma hnorm_mult_ineq: "\<And>x y::'a::real_normed_algebra star. hnorm (x * y) \<le> hnorm x * hnorm y"
+  by transfer (rule norm_mult_ineq)
 
-lemma hnorm_mult:
-  "\<And>x y::'a::real_normed_div_algebra star. hnorm (x * y) = hnorm x * hnorm y"
-by transfer (rule norm_mult)
+lemma hnorm_mult: "\<And>x y::'a::real_normed_div_algebra star. hnorm (x * y) = hnorm x * hnorm y"
+  by transfer (rule norm_mult)
 
-lemma hnorm_hyperpow:
-  "\<And>(x::'a::{real_normed_div_algebra} star) n.
-   hnorm (x pow n) = hnorm x pow n"
-by transfer (rule norm_power)
+lemma hnorm_hyperpow: "\<And>(x::'a::{real_normed_div_algebra} star) n. hnorm (x pow n) = hnorm x pow n"
+  by transfer (rule norm_power)
+
+lemma hnorm_one [simp]: "hnorm (1::'a::real_normed_div_algebra star) = 1"
+  by transfer (rule norm_one)
 
-lemma hnorm_one [simp]:
-  "hnorm (1::'a::real_normed_div_algebra star) = 1"
-by transfer (rule norm_one)
+lemma hnorm_zero [simp]: "hnorm (0::'a::real_normed_vector star) = 0"
+  by transfer (rule norm_zero)
 
-lemma hnorm_zero [simp]:
-  "hnorm (0::'a::real_normed_vector star) = 0"
-by transfer (rule norm_zero)
+lemma zero_less_hnorm_iff [simp]: "\<And>x::'a::real_normed_vector star. 0 < hnorm x \<longleftrightarrow> x \<noteq> 0"
+  by transfer (rule zero_less_norm_iff)
 
-lemma zero_less_hnorm_iff [simp]:
-  "\<And>x::'a::real_normed_vector star. (0 < hnorm x) = (x \<noteq> 0)"
-by transfer (rule zero_less_norm_iff)
+lemma hnorm_minus_cancel [simp]: "\<And>x::'a::real_normed_vector star. hnorm (- x) = hnorm x"
+  by transfer (rule norm_minus_cancel)
 
-lemma hnorm_minus_cancel [simp]:
-  "\<And>x::'a::real_normed_vector star. hnorm (- x) = hnorm x"
-by transfer (rule norm_minus_cancel)
+lemma hnorm_minus_commute: "\<And>a b::'a::real_normed_vector star. hnorm (a - b) = hnorm (b - a)"
+  by transfer (rule norm_minus_commute)
 
-lemma hnorm_minus_commute:
-  "\<And>a b::'a::real_normed_vector star. hnorm (a - b) = hnorm (b - a)"
-by transfer (rule norm_minus_commute)
-
-lemma hnorm_triangle_ineq2:
-  "\<And>a b::'a::real_normed_vector star. hnorm a - hnorm b \<le> hnorm (a - b)"
-by transfer (rule norm_triangle_ineq2)
+lemma hnorm_triangle_ineq2: "\<And>a b::'a::real_normed_vector star. hnorm a - hnorm b \<le> hnorm (a - b)"
+  by transfer (rule norm_triangle_ineq2)
 
-lemma hnorm_triangle_ineq4:
-  "\<And>a b::'a::real_normed_vector star. hnorm (a - b) \<le> hnorm a + hnorm b"
-by transfer (rule norm_triangle_ineq4)
+lemma hnorm_triangle_ineq4: "\<And>a b::'a::real_normed_vector star. hnorm (a - b) \<le> hnorm a + hnorm b"
+  by transfer (rule norm_triangle_ineq4)
 
-lemma abs_hnorm_cancel [simp]:
-  "\<And>a::'a::real_normed_vector star. \<bar>hnorm a\<bar> = hnorm a"
-by transfer (rule abs_norm_cancel)
+lemma abs_hnorm_cancel [simp]: "\<And>a::'a::real_normed_vector star. \<bar>hnorm a\<bar> = hnorm a"
+  by transfer (rule abs_norm_cancel)
 
-lemma hnorm_of_hypreal [simp]:
-  "\<And>r. hnorm (of_hypreal r::'a::real_normed_algebra_1 star) = \<bar>r\<bar>"
-by transfer (rule norm_of_real)
+lemma hnorm_of_hypreal [simp]: "\<And>r. hnorm (of_hypreal r::'a::real_normed_algebra_1 star) = \<bar>r\<bar>"
+  by transfer (rule norm_of_real)
 
 lemma nonzero_hnorm_inverse:
-  "\<And>a::'a::real_normed_div_algebra star.
-   a \<noteq> 0 \<Longrightarrow> hnorm (inverse a) = inverse (hnorm a)"
-by transfer (rule nonzero_norm_inverse)
+  "\<And>a::'a::real_normed_div_algebra star. a \<noteq> 0 \<Longrightarrow> hnorm (inverse a) = inverse (hnorm a)"
+  by transfer (rule nonzero_norm_inverse)
 
 lemma hnorm_inverse:
-  "\<And>a::'a::{real_normed_div_algebra, division_ring} star.
-   hnorm (inverse a) = inverse (hnorm a)"
-by transfer (rule norm_inverse)
+  "\<And>a::'a::{real_normed_div_algebra, division_ring} star. hnorm (inverse a) = inverse (hnorm a)"
+  by transfer (rule norm_inverse)
 
-lemma hnorm_divide:
-  "\<And>a b::'a::{real_normed_field, field} star.
-   hnorm (a / b) = hnorm a / hnorm b"
-by transfer (rule norm_divide)
+lemma hnorm_divide: "\<And>a b::'a::{real_normed_field, field} star. hnorm (a / b) = hnorm a / hnorm b"
+  by transfer (rule norm_divide)
 
-lemma hypreal_hnorm_def [simp]:
-  "\<And>r::hypreal. hnorm r = \<bar>r\<bar>"
-by transfer (rule real_norm_def)
+lemma hypreal_hnorm_def [simp]: "\<And>r::hypreal. hnorm r = \<bar>r\<bar>"
+  by transfer (rule real_norm_def)
 
 lemma hnorm_add_less:
-  "\<And>(x::'a::real_normed_vector star) y r s.
-   \<lbrakk>hnorm x < r; hnorm y < s\<rbrakk> \<Longrightarrow> hnorm (x + y) < r + s"
-by transfer (rule norm_add_less)
+  "\<And>(x::'a::real_normed_vector star) y r s. hnorm x < r \<Longrightarrow> hnorm y < s \<Longrightarrow> hnorm (x + y) < r + s"
+  by transfer (rule norm_add_less)
 
 lemma hnorm_mult_less:
-  "\<And>(x::'a::real_normed_algebra star) y r s.
-   \<lbrakk>hnorm x < r; hnorm y < s\<rbrakk> \<Longrightarrow> hnorm (x * y) < r * s"
-by transfer (rule norm_mult_less)
+  "\<And>(x::'a::real_normed_algebra star) y r s. hnorm x < r \<Longrightarrow> hnorm y < s \<Longrightarrow> hnorm (x * y) < r * s"
+  by transfer (rule norm_mult_less)
 
-lemma hnorm_scaleHR_less:
-  "\<lbrakk>\<bar>x\<bar> < r; hnorm y < s\<rbrakk> \<Longrightarrow> hnorm (scaleHR x y) < r * s"
-apply (simp only: hnorm_scaleHR)
-apply (simp add: mult_strict_mono')
-done
+lemma hnorm_scaleHR_less: "\<bar>x\<bar> < r \<Longrightarrow> hnorm y < s \<Longrightarrow> hnorm (scaleHR x y) < r * s"
+ by (simp only: hnorm_scaleHR) (simp add: mult_strict_mono')
+
+
+subsection \<open>Closure Laws for the Standard Reals\<close>
 
-subsection\<open>Closure Laws for the Standard Reals\<close>
+lemma Reals_minus_iff [simp]: "- x \<in> \<real> \<longleftrightarrow> x \<in> \<real>"
+  apply auto
+  apply (drule Reals_minus)
+  apply auto
+  done
 
-lemma Reals_minus_iff [simp]: "(-x \<in> \<real>) = (x \<in> \<real>)"
-apply auto
-apply (drule Reals_minus, auto)
-done
+lemma Reals_add_cancel: "x + y \<in> \<real> \<Longrightarrow> y \<in> \<real> \<Longrightarrow> x \<in> \<real>"
+  by (drule (1) Reals_diff) simp
 
-lemma Reals_add_cancel: "\<lbrakk>x + y \<in> \<real>; y \<in> \<real>\<rbrakk> \<Longrightarrow> x \<in> \<real>"
-by (drule (1) Reals_diff, simp)
-
-lemma SReal_hrabs: "(x::hypreal) \<in> \<real> ==> \<bar>x\<bar> \<in> \<real>"
-by (simp add: Reals_eq_Standard)
+lemma SReal_hrabs: "x \<in> \<real> \<Longrightarrow> \<bar>x\<bar> \<in> \<real>"
+  for x :: hypreal
+  by (simp add: Reals_eq_Standard)
 
 lemma SReal_hypreal_of_real [simp]: "hypreal_of_real x \<in> \<real>"
-by (simp add: Reals_eq_Standard)
+  by (simp add: Reals_eq_Standard)
 
-lemma SReal_divide_numeral: "r \<in> \<real> ==> r/(numeral w::hypreal) \<in> \<real>"
-by simp
+lemma SReal_divide_numeral: "r \<in> \<real> \<Longrightarrow> r / (numeral w::hypreal) \<in> \<real>"
+  by simp
 
 text \<open>\<open>\<epsilon>\<close> is not in Reals because it is an infinitesimal\<close>
 lemma SReal_epsilon_not_mem: "\<epsilon> \<notin> \<real>"
-apply (simp add: SReal_def)
-apply (auto simp add: hypreal_of_real_not_eq_epsilon [THEN not_sym])
-done
+  by (auto simp: SReal_def hypreal_of_real_not_eq_epsilon [symmetric])
 
 lemma SReal_omega_not_mem: "\<omega> \<notin> \<real>"
-apply (simp add: SReal_def)
-apply (auto simp add: hypreal_of_real_not_eq_omega [THEN not_sym])
-done
+  by (auto simp: SReal_def hypreal_of_real_not_eq_omega [symmetric])
 
 lemma SReal_UNIV_real: "{x. hypreal_of_real x \<in> \<real>} = (UNIV::real set)"
-by simp
+  by simp
 
-lemma SReal_iff: "(x \<in> \<real>) = (\<exists>y. x = hypreal_of_real y)"
-by (simp add: SReal_def)
+lemma SReal_iff: "x \<in> \<real> \<longleftrightarrow> (\<exists>y. x = hypreal_of_real y)"
+  by (simp add: SReal_def)
 
 lemma hypreal_of_real_image: "hypreal_of_real `(UNIV::real set) = \<real>"
-by (simp add: Reals_eq_Standard Standard_def)
+  by (simp add: Reals_eq_Standard Standard_def)
 
 lemma inv_hypreal_of_real_image: "inv hypreal_of_real ` \<real> = UNIV"
-apply (auto simp add: SReal_def)
-apply (rule inj_star_of [THEN inv_f_f, THEN subst], blast)
-done
+  apply (auto simp add: SReal_def)
+  apply (rule inj_star_of [THEN inv_f_f, THEN subst], blast)
+  done
 
-lemma SReal_hypreal_of_real_image:
-      "[| \<exists>x. x: P; P \<subseteq> \<real> |] ==> \<exists>Q. P = hypreal_of_real ` Q"
-by (simp add: SReal_def image_def, blast)
+lemma SReal_hypreal_of_real_image: "\<exists>x. x \<in> P \<Longrightarrow> P \<subseteq> \<real> \<Longrightarrow> \<exists>Q. P = hypreal_of_real ` Q"
+  unfolding SReal_def image_def by blast
 
-lemma SReal_dense:
-     "[| (x::hypreal) \<in> \<real>; y \<in> \<real>;  x<y |] ==> \<exists>r \<in> Reals. x<r & r<y"
-apply (auto simp add: SReal_def)
-apply (drule dense, auto)
-done
+lemma SReal_dense: "x \<in> \<real> \<Longrightarrow> y \<in> \<real> \<Longrightarrow> x < y \<Longrightarrow> \<exists>r \<in> Reals. x < r \<and> r < y"
+  for x y :: hypreal
+  apply (auto simp: SReal_def)
+  apply (drule dense)
+  apply auto
+  done
 
-text\<open>Completeness of Reals, but both lemmas are unused.\<close>
+
+text \<open>Completeness of Reals, but both lemmas are unused.\<close>
 
 lemma SReal_sup_lemma:
-     "P \<subseteq> \<real> ==> ((\<exists>x \<in> P. y < x) =
-      (\<exists>X. hypreal_of_real X \<in> P & y < hypreal_of_real X))"
-by (blast dest!: SReal_iff [THEN iffD1])
+  "P \<subseteq> \<real> \<Longrightarrow> (\<exists>x \<in> P. y < x) = (\<exists>X. hypreal_of_real X \<in> P \<and> y < hypreal_of_real X)"
+  by (blast dest!: SReal_iff [THEN iffD1])
 
 lemma SReal_sup_lemma2:
-     "[| P \<subseteq> \<real>; \<exists>x. x \<in> P; \<exists>y \<in> Reals. \<forall>x \<in> P. x < y |]
-      ==> (\<exists>X. X \<in> {w. hypreal_of_real w \<in> P}) &
-          (\<exists>Y. \<forall>X \<in> {w. hypreal_of_real w \<in> P}. X < Y)"
-apply (rule conjI)
-apply (fast dest!: SReal_iff [THEN iffD1])
-apply (auto, frule subsetD, assumption)
-apply (drule SReal_iff [THEN iffD1])
-apply (auto, rule_tac x = ya in exI, auto)
-done
+  "P \<subseteq> \<real> \<Longrightarrow> \<exists>x. x \<in> P \<Longrightarrow> \<exists>y \<in> Reals. \<forall>x \<in> P. x < y \<Longrightarrow>
+    (\<exists>X. X \<in> {w. hypreal_of_real w \<in> P}) \<and>
+    (\<exists>Y. \<forall>X \<in> {w. hypreal_of_real w \<in> P}. X < Y)"
+  apply (rule conjI)
+   apply (fast dest!: SReal_iff [THEN iffD1])
+  apply (auto, frule subsetD, assumption)
+  apply (drule SReal_iff [THEN iffD1])
+  apply (auto, rule_tac x = ya in exI, auto)
+  done
 
 
-subsection\<open>Set of Finite Elements is a Subring of the Extended Reals\<close>
+subsection \<open>Set of Finite Elements is a Subring of the Extended Reals\<close>
 
-lemma HFinite_add: "[|x \<in> HFinite; y \<in> HFinite|] ==> (x+y) \<in> HFinite"
-apply (simp add: HFinite_def)
-apply (blast intro!: Reals_add hnorm_add_less)
-done
+lemma HFinite_add: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x + y \<in> HFinite"
+  unfolding HFinite_def by (blast intro!: Reals_add hnorm_add_less)
 
-lemma HFinite_mult:
-  fixes x y :: "'a::real_normed_algebra star"
-  shows "[|x \<in> HFinite; y \<in> HFinite|] ==> x*y \<in> HFinite"
-apply (simp add: HFinite_def)
-apply (blast intro!: Reals_mult hnorm_mult_less)
-done
+lemma HFinite_mult: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x * y \<in> HFinite"
+  for x y :: "'a::real_normed_algebra star"
+  unfolding HFinite_def by (blast intro!: Reals_mult hnorm_mult_less)
 
-lemma HFinite_scaleHR:
-  "[|x \<in> HFinite; y \<in> HFinite|] ==> scaleHR x y \<in> HFinite"
-apply (simp add: HFinite_def)
-apply (blast intro!: Reals_mult hnorm_scaleHR_less)
-done
+lemma HFinite_scaleHR: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> scaleHR x y \<in> HFinite"
+  by (auto simp: HFinite_def intro!: Reals_mult hnorm_scaleHR_less)
 
-lemma HFinite_minus_iff: "(-x \<in> HFinite) = (x \<in> HFinite)"
-by (simp add: HFinite_def)
+lemma HFinite_minus_iff: "- x \<in> HFinite \<longleftrightarrow> x \<in> HFinite"
+  by (simp add: HFinite_def)
 
 lemma HFinite_star_of [simp]: "star_of x \<in> HFinite"
-apply (simp add: HFinite_def)
-apply (rule_tac x="star_of (norm x) + 1" in bexI)
-apply (transfer, simp)
-apply (blast intro: Reals_add SReal_hypreal_of_real Reals_1)
-done
+  apply (simp add: HFinite_def)
+  apply (rule_tac x="star_of (norm x) + 1" in bexI)
+   apply (transfer, simp)
+  apply (blast intro: Reals_add SReal_hypreal_of_real Reals_1)
+  done
 
 lemma SReal_subset_HFinite: "(\<real>::hypreal set) \<subseteq> HFinite"
-by (auto simp add: SReal_def)
+  by (auto simp add: SReal_def)
 
-lemma HFiniteD: "x \<in> HFinite ==> \<exists>t \<in> Reals. hnorm x < t"
-by (simp add: HFinite_def)
+lemma HFiniteD: "x \<in> HFinite \<Longrightarrow> \<exists>t \<in> Reals. hnorm x < t"
+  by (simp add: HFinite_def)
 
-lemma HFinite_hrabs_iff [iff]: "(\<bar>x::hypreal\<bar> \<in> HFinite) = (x \<in> HFinite)"
-by (simp add: HFinite_def)
+lemma HFinite_hrabs_iff [iff]: "\<bar>x\<bar> \<in> HFinite \<longleftrightarrow> x \<in> HFinite"
+  for x :: hypreal
+  by (simp add: HFinite_def)
 
-lemma HFinite_hnorm_iff [iff]:
-  "(hnorm (x::hypreal) \<in> HFinite) = (x \<in> HFinite)"
-by (simp add: HFinite_def)
+lemma HFinite_hnorm_iff [iff]: "hnorm x \<in> HFinite \<longleftrightarrow> x \<in> HFinite"
+  for x :: hypreal
+  by (simp add: HFinite_def)
 
 lemma HFinite_numeral [simp]: "numeral w \<in> HFinite"
-unfolding star_numeral_def by (rule HFinite_star_of)
+  unfolding star_numeral_def by (rule HFinite_star_of)
 
-(** As always with numerals, 0 and 1 are special cases **)
+text \<open>As always with numerals, \<open>0\<close> and \<open>1\<close> are special cases.\<close>
 
 lemma HFinite_0 [simp]: "0 \<in> HFinite"
-unfolding star_zero_def by (rule HFinite_star_of)
+  unfolding star_zero_def by (rule HFinite_star_of)
 
 lemma HFinite_1 [simp]: "1 \<in> HFinite"
-unfolding star_one_def by (rule HFinite_star_of)
+  unfolding star_one_def by (rule HFinite_star_of)
 
-lemma hrealpow_HFinite:
-  fixes x :: "'a::{real_normed_algebra,monoid_mult} star"
-  shows "x \<in> HFinite ==> x ^ n \<in> HFinite"
-apply (induct n)
-apply (auto simp add: power_Suc intro: HFinite_mult)
-done
+lemma hrealpow_HFinite: "x \<in> HFinite \<Longrightarrow> x ^ n \<in> HFinite"
+  for x :: "'a::{real_normed_algebra,monoid_mult} star"
+  by (induct n) (auto simp add: power_Suc intro: HFinite_mult)
 
-lemma HFinite_bounded:
-  "[|(x::hypreal) \<in> HFinite; y \<le> x; 0 \<le> y |] ==> y \<in> HFinite"
-apply (cases "x \<le> 0")
-apply (drule_tac y = x in order_trans)
-apply (drule_tac [2] order_antisym)
-apply (auto simp add: linorder_not_le)
-apply (auto intro: order_le_less_trans simp add: abs_if HFinite_def)
-done
+lemma HFinite_bounded: "x \<in> HFinite \<Longrightarrow> y \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> y \<in> HFinite"
+  for x y :: hypreal
+  apply (cases "x \<le> 0")
+   apply (drule_tac y = x in order_trans)
+    apply (drule_tac [2] order_antisym)
+     apply (auto simp add: linorder_not_le)
+  apply (auto intro: order_le_less_trans simp add: abs_if HFinite_def)
+  done
 
 
-subsection\<open>Set of Infinitesimals is a Subring of the Hyperreals\<close>
+subsection \<open>Set of Infinitesimals is a Subring of the Hyperreals\<close>
 
-lemma InfinitesimalI:
-  "(\<And>r. \<lbrakk>r \<in> \<real>; 0 < r\<rbrakk> \<Longrightarrow> hnorm x < r) \<Longrightarrow> x \<in> Infinitesimal"
-by (simp add: Infinitesimal_def)
+lemma InfinitesimalI: "(\<And>r. r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> hnorm x < r) \<Longrightarrow> x \<in> Infinitesimal"
+  by (simp add: Infinitesimal_def)
 
-lemma InfinitesimalD:
-      "x \<in> Infinitesimal ==> \<forall>r \<in> Reals. 0 < r --> hnorm x < r"
-by (simp add: Infinitesimal_def)
+lemma InfinitesimalD: "x \<in> Infinitesimal \<Longrightarrow> \<forall>r \<in> Reals. 0 < r \<longrightarrow> hnorm x < r"
+  by (simp add: Infinitesimal_def)
 
-lemma InfinitesimalI2:
-  "(\<And>r. 0 < r \<Longrightarrow> hnorm x < star_of r) \<Longrightarrow> x \<in> Infinitesimal"
-by (auto simp add: Infinitesimal_def SReal_def)
+lemma InfinitesimalI2: "(\<And>r. 0 < r \<Longrightarrow> hnorm x < star_of r) \<Longrightarrow> x \<in> Infinitesimal"
+  by (auto simp add: Infinitesimal_def SReal_def)
 
-lemma InfinitesimalD2:
-  "\<lbrakk>x \<in> Infinitesimal; 0 < r\<rbrakk> \<Longrightarrow> hnorm x < star_of r"
-by (auto simp add: Infinitesimal_def SReal_def)
+lemma InfinitesimalD2: "x \<in> Infinitesimal \<Longrightarrow> 0 < r \<Longrightarrow> hnorm x < star_of r"
+  by (auto simp add: Infinitesimal_def SReal_def)
 
 lemma Infinitesimal_zero [iff]: "0 \<in> Infinitesimal"
-by (simp add: Infinitesimal_def)
+  by (simp add: Infinitesimal_def)
 
-lemma hypreal_sum_of_halves: "x/(2::hypreal) + x/(2::hypreal) = x"
-by auto
+lemma hypreal_sum_of_halves: "x / 2 + x / 2 = x"
+  for x :: hypreal
+  by auto
 
-lemma Infinitesimal_add:
-     "[| x \<in> Infinitesimal; y \<in> Infinitesimal |] ==> (x+y) \<in> Infinitesimal"
-apply (rule InfinitesimalI)
-apply (rule hypreal_sum_of_halves [THEN subst])
-apply (drule half_gt_zero)
-apply (blast intro: hnorm_add_less SReal_divide_numeral dest: InfinitesimalD)
-done
+lemma Infinitesimal_add: "x \<in> Infinitesimal \<Longrightarrow> y \<in> Infinitesimal \<Longrightarrow> x + y \<in> Infinitesimal"
+  apply (rule InfinitesimalI)
+  apply (rule hypreal_sum_of_halves [THEN subst])
+  apply (drule half_gt_zero)
+  apply (blast intro: hnorm_add_less SReal_divide_numeral dest: InfinitesimalD)
+  done
 
-lemma Infinitesimal_minus_iff [simp]: "(-x:Infinitesimal) = (x:Infinitesimal)"
-by (simp add: Infinitesimal_def)
+lemma Infinitesimal_minus_iff [simp]: "- x \<in> Infinitesimal \<longleftrightarrow> x \<in> Infinitesimal"
+  by (simp add: Infinitesimal_def)
 
-lemma Infinitesimal_hnorm_iff:
-  "(hnorm x \<in> Infinitesimal) = (x \<in> Infinitesimal)"
-by (simp add: Infinitesimal_def)
+lemma Infinitesimal_hnorm_iff: "hnorm x \<in> Infinitesimal \<longleftrightarrow> x \<in> Infinitesimal"
+  by (simp add: Infinitesimal_def)
 
-lemma Infinitesimal_hrabs_iff [iff]:
-  "(\<bar>x::hypreal\<bar> \<in> Infinitesimal) = (x \<in> Infinitesimal)"
-by (simp add: abs_if)
+lemma Infinitesimal_hrabs_iff [iff]: "\<bar>x\<bar> \<in> Infinitesimal \<longleftrightarrow> x \<in> Infinitesimal"
+  for x :: hypreal
+  by (simp add: abs_if)
 
 lemma Infinitesimal_of_hypreal_iff [simp]:
-  "((of_hypreal x::'a::real_normed_algebra_1 star) \<in> Infinitesimal) =
-   (x \<in> Infinitesimal)"
-by (subst Infinitesimal_hnorm_iff [symmetric], simp)
+  "(of_hypreal x::'a::real_normed_algebra_1 star) \<in> Infinitesimal \<longleftrightarrow> x \<in> Infinitesimal"
+  by (subst Infinitesimal_hnorm_iff [symmetric]) simp
 
-lemma Infinitesimal_diff:
-     "[| x \<in> Infinitesimal;  y \<in> Infinitesimal |] ==> x-y \<in> Infinitesimal"
+lemma Infinitesimal_diff: "x \<in> Infinitesimal \<Longrightarrow> y \<in> Infinitesimal \<Longrightarrow> x - y \<in> Infinitesimal"
   using Infinitesimal_add [of x "- y"] by simp
 
-lemma Infinitesimal_mult:
-  fixes x y :: "'a::real_normed_algebra star"
-  shows "[|x \<in> Infinitesimal; y \<in> Infinitesimal|] ==> (x * y) \<in> Infinitesimal"
-apply (rule InfinitesimalI)
-apply (subgoal_tac "hnorm (x * y) < 1 * r", simp only: mult_1)
-apply (rule hnorm_mult_less)
-apply (simp_all add: InfinitesimalD)
-done
+lemma Infinitesimal_mult: "x \<in> Infinitesimal \<Longrightarrow> y \<in> Infinitesimal \<Longrightarrow> x * y \<in> Infinitesimal"
+  for x y :: "'a::real_normed_algebra star"
+  apply (rule InfinitesimalI)
+  apply (subgoal_tac "hnorm (x * y) < 1 * r")
+   apply (simp only: mult_1)
+  apply (rule hnorm_mult_less)
+   apply (simp_all add: InfinitesimalD)
+  done
 
-lemma Infinitesimal_HFinite_mult:
-  fixes x y :: "'a::real_normed_algebra star"
-  shows "[| x \<in> Infinitesimal; y \<in> HFinite |] ==> (x * y) \<in> Infinitesimal"
-apply (rule InfinitesimalI)
-apply (drule HFiniteD, clarify)
-apply (subgoal_tac "0 < t")
-apply (subgoal_tac "hnorm (x * y) < (r / t) * t", simp)
-apply (subgoal_tac "0 < r / t")
-apply (rule hnorm_mult_less)
-apply (simp add: InfinitesimalD)
-apply assumption
-apply simp
-apply (erule order_le_less_trans [OF hnorm_ge_zero])
-done
+lemma Infinitesimal_HFinite_mult: "x \<in> Infinitesimal \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x * y \<in> Infinitesimal"
+  for x y :: "'a::real_normed_algebra star"
+  apply (rule InfinitesimalI)
+  apply (drule HFiniteD, clarify)
+  apply (subgoal_tac "0 < t")
+   apply (subgoal_tac "hnorm (x * y) < (r / t) * t", simp)
+   apply (subgoal_tac "0 < r / t")
+    apply (rule hnorm_mult_less)
+     apply (simp add: InfinitesimalD)
+    apply assumption
+   apply simp
+  apply (erule order_le_less_trans [OF hnorm_ge_zero])
+  done
 
 lemma Infinitesimal_HFinite_scaleHR:
-  "[| x \<in> Infinitesimal; y \<in> HFinite |] ==> scaleHR x y \<in> Infinitesimal"
-apply (rule InfinitesimalI)
-apply (drule HFiniteD, clarify)
-apply (drule InfinitesimalD)
-apply (simp add: hnorm_scaleHR)
-apply (subgoal_tac "0 < t")
-apply (subgoal_tac "\<bar>x\<bar> * hnorm y < (r / t) * t", simp)
-apply (subgoal_tac "0 < r / t")
-apply (rule mult_strict_mono', simp_all)
-apply (erule order_le_less_trans [OF hnorm_ge_zero])
-done
+  "x \<in> Infinitesimal \<Longrightarrow> y \<in> HFinite \<Longrightarrow> scaleHR x y \<in> Infinitesimal"
+  apply (rule InfinitesimalI)
+  apply (drule HFiniteD, clarify)
+  apply (drule InfinitesimalD)
+  apply (simp add: hnorm_scaleHR)
+  apply (subgoal_tac "0 < t")
+   apply (subgoal_tac "\<bar>x\<bar> * hnorm y < (r / t) * t", simp)
+   apply (subgoal_tac "0 < r / t")
+    apply (rule mult_strict_mono', simp_all)
+  apply (erule order_le_less_trans [OF hnorm_ge_zero])
+  done
 
 lemma Infinitesimal_HFinite_mult2:
-  fixes x y :: "'a::real_normed_algebra star"
-  shows "[| x \<in> Infinitesimal; y \<in> HFinite |] ==> (y * x) \<in> Infinitesimal"
-apply (rule InfinitesimalI)
-apply (drule HFiniteD, clarify)
-apply (subgoal_tac "0 < t")
-apply (subgoal_tac "hnorm (y * x) < t * (r / t)", simp)
-apply (subgoal_tac "0 < r / t")
-apply (rule hnorm_mult_less)
-apply assumption
-apply (simp add: InfinitesimalD)
-apply simp
-apply (erule order_le_less_trans [OF hnorm_ge_zero])
-done
+  "x \<in> Infinitesimal \<Longrightarrow> y \<in> HFinite \<Longrightarrow> y * x \<in> Infinitesimal"
+  for x y :: "'a::real_normed_algebra star"
+  apply (rule InfinitesimalI)
+  apply (drule HFiniteD, clarify)
+  apply (subgoal_tac "0 < t")
+   apply (subgoal_tac "hnorm (y * x) < t * (r / t)", simp)
+   apply (subgoal_tac "0 < r / t")
+    apply (rule hnorm_mult_less)
+     apply assumption
+    apply (simp add: InfinitesimalD)
+   apply simp
+  apply (erule order_le_less_trans [OF hnorm_ge_zero])
+  done
 
-lemma Infinitesimal_scaleR2:
-  "x \<in> Infinitesimal ==> a *\<^sub>R x \<in> Infinitesimal"
-apply (case_tac "a = 0", simp)
-apply (rule InfinitesimalI)
-apply (drule InfinitesimalD)
-apply (drule_tac x="r / \<bar>star_of a\<bar>" in bspec)
-apply (simp add: Reals_eq_Standard)
-apply simp
-apply (simp add: hnorm_scaleR pos_less_divide_eq mult.commute)
-done
+lemma Infinitesimal_scaleR2: "x \<in> Infinitesimal \<Longrightarrow> a *\<^sub>R x \<in> Infinitesimal"
+  apply (case_tac "a = 0", simp)
+  apply (rule InfinitesimalI)
+  apply (drule InfinitesimalD)
+  apply (drule_tac x="r / \<bar>star_of a\<bar>" in bspec)
+   apply (simp add: Reals_eq_Standard)
+  apply simp
+  apply (simp add: hnorm_scaleR pos_less_divide_eq mult.commute)
+  done
 
 lemma Compl_HFinite: "- HFinite = HInfinite"
-apply (auto simp add: HInfinite_def HFinite_def linorder_not_less)
-apply (rule_tac y="r + 1" in order_less_le_trans, simp)
-apply simp
-done
+  apply (auto simp add: HInfinite_def HFinite_def linorder_not_less)
+  apply (rule_tac y="r + 1" in order_less_le_trans, simp)
+  apply simp
+  done
 
-lemma HInfinite_inverse_Infinitesimal:
-  fixes x :: "'a::real_normed_div_algebra star"
-  shows "x \<in> HInfinite ==> inverse x \<in> Infinitesimal"
-apply (rule InfinitesimalI)
-apply (subgoal_tac "x \<noteq> 0")
-apply (rule inverse_less_imp_less)
-apply (simp add: nonzero_hnorm_inverse)
-apply (simp add: HInfinite_def Reals_inverse)
-apply assumption
-apply (clarify, simp add: Compl_HFinite [symmetric])
-done
+lemma HInfinite_inverse_Infinitesimal: "x \<in> HInfinite \<Longrightarrow> inverse x \<in> Infinitesimal"
+  for x :: "'a::real_normed_div_algebra star"
+  apply (rule InfinitesimalI)
+  apply (subgoal_tac "x \<noteq> 0")
+   apply (rule inverse_less_imp_less)
+    apply (simp add: nonzero_hnorm_inverse)
+    apply (simp add: HInfinite_def Reals_inverse)
+   apply assumption
+  apply (clarify, simp add: Compl_HFinite [symmetric])
+  done
 
 lemma HInfiniteI: "(\<And>r. r \<in> \<real> \<Longrightarrow> r < hnorm x) \<Longrightarrow> x \<in> HInfinite"
-by (simp add: HInfinite_def)
+  by (simp add: HInfinite_def)
 
-lemma HInfiniteD: "\<lbrakk>x \<in> HInfinite; r \<in> \<real>\<rbrakk> \<Longrightarrow> r < hnorm x"
-by (simp add: HInfinite_def)
+lemma HInfiniteD: "x \<in> HInfinite \<Longrightarrow> r \<in> \<real> \<Longrightarrow> r < hnorm x"
+  by (simp add: HInfinite_def)
 
-lemma HInfinite_mult:
-  fixes x y :: "'a::real_normed_div_algebra star"
-  shows "[|x \<in> HInfinite; y \<in> HInfinite|] ==> (x*y) \<in> HInfinite"
-apply (rule HInfiniteI, simp only: hnorm_mult)
-apply (subgoal_tac "r * 1 < hnorm x * hnorm y", simp only: mult_1)
-apply (case_tac "x = 0", simp add: HInfinite_def)
-apply (rule mult_strict_mono)
-apply (simp_all add: HInfiniteD)
-done
+lemma HInfinite_mult: "x \<in> HInfinite \<Longrightarrow> y \<in> HInfinite \<Longrightarrow> x * y \<in> HInfinite"
+  for x y :: "'a::real_normed_div_algebra star"
+  apply (rule HInfiniteI, simp only: hnorm_mult)
+  apply (subgoal_tac "r * 1 < hnorm x * hnorm y", simp only: mult_1)
+  apply (case_tac "x = 0", simp add: HInfinite_def)
+  apply (rule mult_strict_mono)
+     apply (simp_all add: HInfiniteD)
+  done
 
-lemma hypreal_add_zero_less_le_mono: "[|r < x; (0::hypreal) \<le> y|] ==> r < x+y"
-by (auto dest: add_less_le_mono)
+lemma hypreal_add_zero_less_le_mono: "r < x \<Longrightarrow> 0 \<le> y \<Longrightarrow> r < x + y"
+  for r x y :: hypreal
+  by (auto dest: add_less_le_mono)
 
-lemma HInfinite_add_ge_zero:
-     "[|(x::hypreal) \<in> HInfinite; 0 \<le> y; 0 \<le> x|] ==> (x + y): HInfinite"
-by (auto intro!: hypreal_add_zero_less_le_mono
-       simp add: abs_if add.commute add_nonneg_nonneg HInfinite_def)
+lemma HInfinite_add_ge_zero: "x \<in> HInfinite \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> x \<Longrightarrow> x + y \<in> HInfinite"
+  for x y :: hypreal
+  by (auto simp: abs_if add.commute HInfinite_def)
 
-lemma HInfinite_add_ge_zero2:
-     "[|(x::hypreal) \<in> HInfinite; 0 \<le> y; 0 \<le> x|] ==> (y + x): HInfinite"
-by (auto intro!: HInfinite_add_ge_zero simp add: add.commute)
+lemma HInfinite_add_ge_zero2: "x \<in> HInfinite \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> x \<Longrightarrow> y + x \<in> HInfinite"
+  for x y :: hypreal
+  by (auto intro!: HInfinite_add_ge_zero simp add: add.commute)
 
-lemma HInfinite_add_gt_zero:
-     "[|(x::hypreal) \<in> HInfinite; 0 < y; 0 < x|] ==> (x + y): HInfinite"
-by (blast intro: HInfinite_add_ge_zero order_less_imp_le)
+lemma HInfinite_add_gt_zero: "x \<in> HInfinite \<Longrightarrow> 0 < y \<Longrightarrow> 0 < x \<Longrightarrow> x + y \<in> HInfinite"
+  for x y :: hypreal
+  by (blast intro: HInfinite_add_ge_zero order_less_imp_le)
 
-lemma HInfinite_minus_iff: "(-x \<in> HInfinite) = (x \<in> HInfinite)"
-by (simp add: HInfinite_def)
+lemma HInfinite_minus_iff: "- x \<in> HInfinite \<longleftrightarrow> x \<in> HInfinite"
+  by (simp add: HInfinite_def)
 
-lemma HInfinite_add_le_zero:
-     "[|(x::hypreal) \<in> HInfinite; y \<le> 0; x \<le> 0|] ==> (x + y): HInfinite"
-apply (drule HInfinite_minus_iff [THEN iffD2])
-apply (rule HInfinite_minus_iff [THEN iffD1])
-apply (simp only: minus_add add.commute)
-apply (rule HInfinite_add_ge_zero)
-apply simp_all
-done
+lemma HInfinite_add_le_zero: "x \<in> HInfinite \<Longrightarrow> y \<le> 0 \<Longrightarrow> x \<le> 0 \<Longrightarrow> x + y \<in> HInfinite"
+  for x y :: hypreal
+  apply (drule HInfinite_minus_iff [THEN iffD2])
+  apply (rule HInfinite_minus_iff [THEN iffD1])
+  apply (simp only: minus_add add.commute)
+  apply (rule HInfinite_add_ge_zero)
+    apply simp_all
+  done
 
-lemma HInfinite_add_lt_zero:
-     "[|(x::hypreal) \<in> HInfinite; y < 0; x < 0|] ==> (x + y): HInfinite"
-by (blast intro: HInfinite_add_le_zero order_less_imp_le)
+lemma HInfinite_add_lt_zero: "x \<in> HInfinite \<Longrightarrow> y < 0 \<Longrightarrow> x < 0 \<Longrightarrow> x + y \<in> HInfinite"
+  for x y :: hypreal
+  by (blast intro: HInfinite_add_le_zero order_less_imp_le)
 
 lemma HFinite_sum_squares:
-  fixes a b c :: "'a::real_normed_algebra star"
-  shows "[|a: HFinite; b: HFinite; c: HFinite|]
-      ==> a*a + b*b + c*c \<in> HFinite"
-by (auto intro: HFinite_mult HFinite_add)
+  "a \<in> HFinite \<Longrightarrow> b \<in> HFinite \<Longrightarrow> c \<in> HFinite \<Longrightarrow> a * a + b * b + c * c \<in> HFinite"
+  for a b c :: "'a::real_normed_algebra star"
+  by (auto intro: HFinite_mult HFinite_add)
 
-lemma not_Infinitesimal_not_zero: "x \<notin> Infinitesimal ==> x \<noteq> 0"
-by auto
+lemma not_Infinitesimal_not_zero: "x \<notin> Infinitesimal \<Longrightarrow> x \<noteq> 0"
+  by auto
 
-lemma not_Infinitesimal_not_zero2: "x \<in> HFinite - Infinitesimal ==> x \<noteq> 0"
-by auto
+lemma not_Infinitesimal_not_zero2: "x \<in> HFinite - Infinitesimal \<Longrightarrow> x \<noteq> 0"
+  by auto
 
 lemma HFinite_diff_Infinitesimal_hrabs:
-  "(x::hypreal) \<in> HFinite - Infinitesimal ==> \<bar>x\<bar> \<in> HFinite - Infinitesimal"
-by blast
+  "x \<in> HFinite - Infinitesimal \<Longrightarrow> \<bar>x\<bar> \<in> HFinite - Infinitesimal"
+  for x :: hypreal
+  by blast
 
-lemma hnorm_le_Infinitesimal:
-  "\<lbrakk>e \<in> Infinitesimal; hnorm x \<le> e\<rbrakk> \<Longrightarrow> x \<in> Infinitesimal"
-by (auto simp add: Infinitesimal_def abs_less_iff)
+lemma hnorm_le_Infinitesimal: "e \<in> Infinitesimal \<Longrightarrow> hnorm x \<le> e \<Longrightarrow> x \<in> Infinitesimal"
+  by (auto simp: Infinitesimal_def abs_less_iff)
 
-lemma hnorm_less_Infinitesimal:
-  "\<lbrakk>e \<in> Infinitesimal; hnorm x < e\<rbrakk> \<Longrightarrow> x \<in> Infinitesimal"
-by (erule hnorm_le_Infinitesimal, erule order_less_imp_le)
+lemma hnorm_less_Infinitesimal: "e \<in> Infinitesimal \<Longrightarrow> hnorm x < e \<Longrightarrow> x \<in> Infinitesimal"
+  by (erule hnorm_le_Infinitesimal, erule order_less_imp_le)
 
-lemma hrabs_le_Infinitesimal:
-     "[| e \<in> Infinitesimal; \<bar>x::hypreal\<bar> \<le> e |] ==> x \<in> Infinitesimal"
-by (erule hnorm_le_Infinitesimal, simp)
+lemma hrabs_le_Infinitesimal: "e \<in> Infinitesimal \<Longrightarrow> \<bar>x\<bar> \<le> e \<Longrightarrow> x \<in> Infinitesimal"
+  for x :: hypreal
+  by (erule hnorm_le_Infinitesimal) simp
 
-lemma hrabs_less_Infinitesimal:
-      "[| e \<in> Infinitesimal; \<bar>x::hypreal\<bar> < e |] ==> x \<in> Infinitesimal"
-by (erule hnorm_less_Infinitesimal, simp)
+lemma hrabs_less_Infinitesimal: "e \<in> Infinitesimal \<Longrightarrow> \<bar>x\<bar> < e \<Longrightarrow> x \<in> Infinitesimal"
+  for x :: hypreal
+  by (erule hnorm_less_Infinitesimal) simp
 
 lemma Infinitesimal_interval:
-      "[| e \<in> Infinitesimal; e' \<in> Infinitesimal; e' < x ; x < e |]
-       ==> (x::hypreal) \<in> Infinitesimal"
-by (auto simp add: Infinitesimal_def abs_less_iff)
+  "e \<in> Infinitesimal \<Longrightarrow> e' \<in> Infinitesimal \<Longrightarrow> e' < x \<Longrightarrow> x < e \<Longrightarrow> x \<in> Infinitesimal"
+  for x :: hypreal
+  by (auto simp add: Infinitesimal_def abs_less_iff)
 
 lemma Infinitesimal_interval2:
-     "[| e \<in> Infinitesimal; e' \<in> Infinitesimal;
-         e' \<le> x ; x \<le> e |] ==> (x::hypreal) \<in> Infinitesimal"
-by (auto intro: Infinitesimal_interval simp add: order_le_less)
+  "e \<in> Infinitesimal \<Longrightarrow> e' \<in> Infinitesimal \<Longrightarrow> e' \<le> x \<Longrightarrow> x \<le> e \<Longrightarrow> x \<in> Infinitesimal"
+  for x :: hypreal
+  by (auto intro: Infinitesimal_interval simp add: order_le_less)
 
 
-lemma lemma_Infinitesimal_hyperpow:
-     "[| (x::hypreal) \<in> Infinitesimal; 0 < N |] ==> \<bar>x pow N\<bar> \<le> \<bar>x\<bar>"
-apply (unfold Infinitesimal_def)
-apply (auto intro!: hyperpow_Suc_le_self2
-          simp add: hyperpow_hrabs [symmetric] hypnat_gt_zero_iff2 abs_ge_zero)
-done
+lemma lemma_Infinitesimal_hyperpow: "x \<in> Infinitesimal \<Longrightarrow> 0 < N \<Longrightarrow> \<bar>x pow N\<bar> \<le> \<bar>x\<bar>"
+  for x :: hypreal
+  apply (unfold Infinitesimal_def)
+  apply (auto intro!: hyperpow_Suc_le_self2
+      simp: hyperpow_hrabs [symmetric] hypnat_gt_zero_iff2 abs_ge_zero)
+  done
 
-lemma Infinitesimal_hyperpow:
-     "[| (x::hypreal) \<in> Infinitesimal; 0 < N |] ==> x pow N \<in> Infinitesimal"
-apply (rule hrabs_le_Infinitesimal)
-apply (rule_tac [2] lemma_Infinitesimal_hyperpow, auto)
-done
+lemma Infinitesimal_hyperpow: "x \<in> Infinitesimal \<Longrightarrow> 0 < N \<Longrightarrow> x pow N \<in> Infinitesimal"
+  for x :: hypreal
+  apply (rule hrabs_le_Infinitesimal)
+   apply (rule_tac [2] lemma_Infinitesimal_hyperpow)
+  apply auto
+  done
 
 lemma hrealpow_hyperpow_Infinitesimal_iff:
-     "(x ^ n \<in> Infinitesimal) = (x pow (hypnat_of_nat n) \<in> Infinitesimal)"
-by (simp only: hyperpow_hypnat_of_nat)
+  "(x ^ n \<in> Infinitesimal) \<longleftrightarrow> x pow (hypnat_of_nat n) \<in> Infinitesimal"
+  by (simp only: hyperpow_hypnat_of_nat)
 
-lemma Infinitesimal_hrealpow:
-     "[| (x::hypreal) \<in> Infinitesimal; 0 < n |] ==> x ^ n \<in> Infinitesimal"
-by (simp add: hrealpow_hyperpow_Infinitesimal_iff Infinitesimal_hyperpow)
+lemma Infinitesimal_hrealpow: "x \<in> Infinitesimal \<Longrightarrow> 0 < n \<Longrightarrow> x ^ n \<in> Infinitesimal"
+  for x :: hypreal
+  by (simp add: hrealpow_hyperpow_Infinitesimal_iff Infinitesimal_hyperpow)
 
 lemma not_Infinitesimal_mult:
-  fixes x y :: "'a::real_normed_div_algebra star"
-  shows "[| x \<notin> Infinitesimal;  y \<notin> Infinitesimal|] ==> (x*y) \<notin>Infinitesimal"
-apply (unfold Infinitesimal_def, clarify, rename_tac r s)
-apply (simp only: linorder_not_less hnorm_mult)
-apply (drule_tac x = "r * s" in bspec)
-apply (fast intro: Reals_mult)
-apply (simp)
-apply (drule_tac c = s and d = "hnorm y" and a = r and b = "hnorm x" in mult_mono)
-apply (simp_all (no_asm_simp))
-done
+  "x \<notin> Infinitesimal \<Longrightarrow> y \<notin> Infinitesimal \<Longrightarrow> x * y \<notin> Infinitesimal"
+  for x y :: "'a::real_normed_div_algebra star"
+  apply (unfold Infinitesimal_def, clarify, rename_tac r s)
+  apply (simp only: linorder_not_less hnorm_mult)
+  apply (drule_tac x = "r * s" in bspec)
+   apply (fast intro: Reals_mult)
+  apply simp
+  apply (drule_tac c = s and d = "hnorm y" and a = r and b = "hnorm x" in mult_mono)
+     apply simp_all
+  done
 
-lemma Infinitesimal_mult_disj:
-  fixes x y :: "'a::real_normed_div_algebra star"
-  shows "x*y \<in> Infinitesimal ==> x \<in> Infinitesimal | y \<in> Infinitesimal"
-apply (rule ccontr)
-apply (drule de_Morgan_disj [THEN iffD1])
-apply (fast dest: not_Infinitesimal_mult)
-done
+lemma Infinitesimal_mult_disj: "x * y \<in> Infinitesimal \<Longrightarrow> x \<in> Infinitesimal \<or> y \<in> Infinitesimal"
+  for x y :: "'a::real_normed_div_algebra star"
+  apply (rule ccontr)
+  apply (drule de_Morgan_disj [THEN iffD1])
+  apply (fast dest: not_Infinitesimal_mult)
+  done
 
-lemma HFinite_Infinitesimal_not_zero: "x \<in> HFinite-Infinitesimal ==> x \<noteq> 0"
-by blast
+lemma HFinite_Infinitesimal_not_zero: "x \<in> HFinite-Infinitesimal \<Longrightarrow> x \<noteq> 0"
+  by blast
 
 lemma HFinite_Infinitesimal_diff_mult:
-  fixes x y :: "'a::real_normed_div_algebra star"
-  shows "[| x \<in> HFinite - Infinitesimal;
-                   y \<in> HFinite - Infinitesimal
-                |] ==> (x*y) \<in> HFinite - Infinitesimal"
-apply clarify
-apply (blast dest: HFinite_mult not_Infinitesimal_mult)
-done
+  "x \<in> HFinite - Infinitesimal \<Longrightarrow> y \<in> HFinite - Infinitesimal \<Longrightarrow> x * y \<in> HFinite - Infinitesimal"
+  for x y :: "'a::real_normed_div_algebra star"
+  apply clarify
+  apply (blast dest: HFinite_mult not_Infinitesimal_mult)
+  done
 
-lemma Infinitesimal_subset_HFinite:
-      "Infinitesimal \<subseteq> HFinite"
-apply (simp add: Infinitesimal_def HFinite_def, auto)
-apply (rule_tac x = 1 in bexI, auto)
-done
+lemma Infinitesimal_subset_HFinite: "Infinitesimal \<subseteq> HFinite"
+  apply (simp add: Infinitesimal_def HFinite_def)
+  apply auto
+  apply (rule_tac x = 1 in bexI)
+  apply auto
+  done
 
-lemma Infinitesimal_star_of_mult:
-  fixes x :: "'a::real_normed_algebra star"
-  shows "x \<in> Infinitesimal ==> x * star_of r \<in> Infinitesimal"
-by (erule HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult])
+lemma Infinitesimal_star_of_mult: "x \<in> Infinitesimal \<Longrightarrow> x * star_of r \<in> Infinitesimal"
+  for x :: "'a::real_normed_algebra star"
+  by (erule HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult])
 
-lemma Infinitesimal_star_of_mult2:
-  fixes x :: "'a::real_normed_algebra star"
-  shows "x \<in> Infinitesimal ==> star_of r * x \<in> Infinitesimal"
-by (erule HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult2])
+lemma Infinitesimal_star_of_mult2: "x \<in> Infinitesimal \<Longrightarrow> star_of r * x \<in> Infinitesimal"
+  for x :: "'a::real_normed_algebra star"
+  by (erule HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult2])
 
 
-subsection\<open>The Infinitely Close Relation\<close>
+subsection \<open>The Infinitely Close Relation\<close>
 
-lemma mem_infmal_iff: "(x \<in> Infinitesimal) = (x \<approx> 0)"
-by (simp add: Infinitesimal_def approx_def)
+lemma mem_infmal_iff: "x \<in> Infinitesimal \<longleftrightarrow> x \<approx> 0"
+  by (simp add: Infinitesimal_def approx_def)
 
-lemma approx_minus_iff: " (x \<approx> y) = (x - y \<approx> 0)"
-by (simp add: approx_def)
+lemma approx_minus_iff: "x \<approx> y \<longleftrightarrow> x - y \<approx> 0"
+  by (simp add: approx_def)
 
-lemma approx_minus_iff2: " (x \<approx> y) = (-y + x \<approx> 0)"
-by (simp add: approx_def add.commute)
+lemma approx_minus_iff2: "x \<approx> y \<longleftrightarrow> - y + x \<approx> 0"
+  by (simp add: approx_def add.commute)
 
 lemma approx_refl [iff]: "x \<approx> x"
-by (simp add: approx_def Infinitesimal_def)
+  by (simp add: approx_def Infinitesimal_def)
 
-lemma hypreal_minus_distrib1: "-(y + -(x::'a::ab_group_add)) = x + -y"
-by (simp add: add.commute)
+lemma hypreal_minus_distrib1: "- (y + - x) = x + -y"
+  for x y :: "'a::ab_group_add"
+  by (simp add: add.commute)
 
-lemma approx_sym: "x \<approx> y ==> y \<approx> x"
-apply (simp add: approx_def)
-apply (drule Infinitesimal_minus_iff [THEN iffD2])
-apply simp
-done
+lemma approx_sym: "x \<approx> y \<Longrightarrow> y \<approx> x"
+  apply (simp add: approx_def)
+  apply (drule Infinitesimal_minus_iff [THEN iffD2])
+  apply simp
+  done
 
-lemma approx_trans: "[| x \<approx> y; y \<approx> z |] ==> x \<approx> z"
-apply (simp add: approx_def)
-apply (drule (1) Infinitesimal_add)
-apply simp
-done
+lemma approx_trans: "x \<approx> y \<Longrightarrow> y \<approx> z \<Longrightarrow> x \<approx> z"
+  apply (simp add: approx_def)
+  apply (drule (1) Infinitesimal_add)
+  apply simp
+  done
 
-lemma approx_trans2: "[| r \<approx> x; s \<approx> x |] ==> r \<approx> s"
-by (blast intro: approx_sym approx_trans)
+lemma approx_trans2: "r \<approx> x \<Longrightarrow> s \<approx> x \<Longrightarrow> r \<approx> s"
+  by (blast intro: approx_sym approx_trans)
 
-lemma approx_trans3: "[| x \<approx> r; x \<approx> s|] ==> r \<approx> s"
-by (blast intro: approx_sym approx_trans)
+lemma approx_trans3: "x \<approx> r \<Longrightarrow> x \<approx> s \<Longrightarrow> r \<approx> s"
+  by (blast intro: approx_sym approx_trans)
 
-lemma approx_reorient: "(x \<approx> y) = (y \<approx> x)"
-by (blast intro: approx_sym)
+lemma approx_reorient: "x \<approx> y \<longleftrightarrow> y \<approx> x"
+  by (blast intro: approx_sym)
 
-(*reorientation simplification procedure: reorients (polymorphic)
-  0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*)
+text \<open>Reorientation simplification procedure: reorients (polymorphic)
+  \<open>0 = x\<close>, \<open>1 = x\<close>, \<open>nnn = x\<close> provided \<open>x\<close> isn't \<open>0\<close>, \<open>1\<close> or a numeral.\<close>
 simproc_setup approx_reorient_simproc
   ("0 \<approx> x" | "1 \<approx> y" | "numeral w \<approx> z" | "- 1 \<approx> y" | "- numeral w \<approx> r") =
 \<open>
@@ -658,1133 +593,1067 @@
   in proc end
 \<close>
 
-lemma Infinitesimal_approx_minus: "(x-y \<in> Infinitesimal) = (x \<approx> y)"
-by (simp add: approx_minus_iff [symmetric] mem_infmal_iff)
+lemma Infinitesimal_approx_minus: "x - y \<in> Infinitesimal \<longleftrightarrow> x \<approx> y"
+  by (simp add: approx_minus_iff [symmetric] mem_infmal_iff)
 
-lemma approx_monad_iff: "(x \<approx> y) = (monad(x)=monad(y))"
-apply (simp add: monad_def)
-apply (auto dest: approx_sym elim!: approx_trans equalityCE)
-done
+lemma approx_monad_iff: "x \<approx> y \<longleftrightarrow> monad x = monad y"
+  by (auto simp add: monad_def dest: approx_sym elim!: approx_trans equalityCE)
 
-lemma Infinitesimal_approx:
-     "[| x \<in> Infinitesimal; y \<in> Infinitesimal |] ==> x \<approx> y"
-apply (simp add: mem_infmal_iff)
-apply (blast intro: approx_trans approx_sym)
-done
+lemma Infinitesimal_approx: "x \<in> Infinitesimal \<Longrightarrow> y \<in> Infinitesimal \<Longrightarrow> x \<approx> y"
+  apply (simp add: mem_infmal_iff)
+  apply (blast intro: approx_trans approx_sym)
+  done
 
-lemma approx_add: "[| a \<approx> b; c \<approx> d |] ==> a+c \<approx> b+d"
+lemma approx_add: "a \<approx> b \<Longrightarrow> c \<approx> d \<Longrightarrow> a + c \<approx> b + d"
 proof (unfold approx_def)
   assume inf: "a - b \<in> Infinitesimal" "c - d \<in> Infinitesimal"
   have "a + c - (b + d) = (a - b) + (c - d)" by simp
-  also have "... \<in> Infinitesimal" using inf by (rule Infinitesimal_add)
+  also have "... \<in> Infinitesimal"
+    using inf by (rule Infinitesimal_add)
   finally show "a + c - (b + d) \<in> Infinitesimal" .
 qed
 
-lemma approx_minus: "a \<approx> b ==> -a \<approx> -b"
-apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])
-apply (drule approx_minus_iff [THEN iffD1])
-apply (simp add: add.commute)
-done
+lemma approx_minus: "a \<approx> b \<Longrightarrow> - a \<approx> - b"
+  apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])
+  apply (drule approx_minus_iff [THEN iffD1])
+  apply (simp add: add.commute)
+  done
 
-lemma approx_minus2: "-a \<approx> -b ==> a \<approx> b"
-by (auto dest: approx_minus)
+lemma approx_minus2: "- a \<approx> - b \<Longrightarrow> a \<approx> b"
+  by (auto dest: approx_minus)
 
-lemma approx_minus_cancel [simp]: "(-a \<approx> -b) = (a \<approx> b)"
-by (blast intro: approx_minus approx_minus2)
+lemma approx_minus_cancel [simp]: "- a \<approx> - b \<longleftrightarrow> a \<approx> b"
+  by (blast intro: approx_minus approx_minus2)
 
-lemma approx_add_minus: "[| a \<approx> b; c \<approx> d |] ==> a + -c \<approx> b + -d"
-by (blast intro!: approx_add approx_minus)
+lemma approx_add_minus: "a \<approx> b \<Longrightarrow> c \<approx> d \<Longrightarrow> a + - c \<approx> b + - d"
+  by (blast intro!: approx_add approx_minus)
 
-lemma approx_diff: "[| a \<approx> b; c \<approx> d |] ==> a - c \<approx> b - d"
+lemma approx_diff: "a \<approx> b \<Longrightarrow> c \<approx> d \<Longrightarrow> a - c \<approx> b - d"
   using approx_add [of a b "- c" "- d"] by simp
 
-lemma approx_mult1:
-  fixes a b c :: "'a::real_normed_algebra star"
-  shows "[| a \<approx> b; c: HFinite|] ==> a*c \<approx> b*c"
-by (simp add: approx_def Infinitesimal_HFinite_mult
-              left_diff_distrib [symmetric])
+lemma approx_mult1: "a \<approx> b \<Longrightarrow> c \<in> HFinite \<Longrightarrow> a * c \<approx> b * c"
+  for a b c :: "'a::real_normed_algebra star"
+  by (simp add: approx_def Infinitesimal_HFinite_mult left_diff_distrib [symmetric])
+
+lemma approx_mult2: "a \<approx> b \<Longrightarrow> c \<in> HFinite \<Longrightarrow> c * a \<approx> c * b"
+  for a b c :: "'a::real_normed_algebra star"
+  by (simp add: approx_def Infinitesimal_HFinite_mult2 right_diff_distrib [symmetric])
 
-lemma approx_mult2:
-  fixes a b c :: "'a::real_normed_algebra star"
-  shows "[|a \<approx> b; c: HFinite|] ==> c*a \<approx> c*b"
-by (simp add: approx_def Infinitesimal_HFinite_mult2
-              right_diff_distrib [symmetric])
+lemma approx_mult_subst: "u \<approx> v * x \<Longrightarrow> x \<approx> y \<Longrightarrow> v \<in> HFinite \<Longrightarrow> u \<approx> v * y"
+  for u v x y :: "'a::real_normed_algebra star"
+  by (blast intro: approx_mult2 approx_trans)
 
-lemma approx_mult_subst:
-  fixes u v x y :: "'a::real_normed_algebra star"
-  shows "[|u \<approx> v*x; x \<approx> y; v \<in> HFinite|] ==> u \<approx> v*y"
-by (blast intro: approx_mult2 approx_trans)
-
-lemma approx_mult_subst2:
-  fixes u v x y :: "'a::real_normed_algebra star"
-  shows "[| u \<approx> x*v; x \<approx> y; v \<in> HFinite |] ==> u \<approx> y*v"
-by (blast intro: approx_mult1 approx_trans)
+lemma approx_mult_subst2: "u \<approx> x * v \<Longrightarrow> x \<approx> y \<Longrightarrow> v \<in> HFinite \<Longrightarrow> u \<approx> y * v"
+  for u v x y :: "'a::real_normed_algebra star"
+  by (blast intro: approx_mult1 approx_trans)
 
-lemma approx_mult_subst_star_of:
-  fixes u x y :: "'a::real_normed_algebra star"
-  shows "[| u \<approx> x*star_of v; x \<approx> y |] ==> u \<approx> y*star_of v"
-by (auto intro: approx_mult_subst2)
+lemma approx_mult_subst_star_of: "u \<approx> x * star_of v \<Longrightarrow> x \<approx> y \<Longrightarrow> u \<approx> y * star_of v"
+  for u x y :: "'a::real_normed_algebra star"
+  by (auto intro: approx_mult_subst2)
 
-lemma approx_eq_imp: "a = b ==> a \<approx> b"
-by (simp add: approx_def)
+lemma approx_eq_imp: "a = b \<Longrightarrow> a \<approx> b"
+  by (simp add: approx_def)
 
-lemma Infinitesimal_minus_approx: "x \<in> Infinitesimal ==> -x \<approx> x"
-by (blast intro: Infinitesimal_minus_iff [THEN iffD2]
-                    mem_infmal_iff [THEN iffD1] approx_trans2)
+lemma Infinitesimal_minus_approx: "x \<in> Infinitesimal \<Longrightarrow> - x \<approx> x"
+  by (blast intro: Infinitesimal_minus_iff [THEN iffD2] mem_infmal_iff [THEN iffD1] approx_trans2)
 
-lemma bex_Infinitesimal_iff: "(\<exists>y \<in> Infinitesimal. x - z = y) = (x \<approx> z)"
-by (simp add: approx_def)
+lemma bex_Infinitesimal_iff: "(\<exists>y \<in> Infinitesimal. x - z = y) \<longleftrightarrow> x \<approx> z"
+  by (simp add: approx_def)
 
-lemma bex_Infinitesimal_iff2: "(\<exists>y \<in> Infinitesimal. x = z + y) = (x \<approx> z)"
-by (force simp add: bex_Infinitesimal_iff [symmetric])
+lemma bex_Infinitesimal_iff2: "(\<exists>y \<in> Infinitesimal. x = z + y) \<longleftrightarrow> x \<approx> z"
+  by (force simp add: bex_Infinitesimal_iff [symmetric])
 
-lemma Infinitesimal_add_approx: "[| y \<in> Infinitesimal; x + y = z |] ==> x \<approx> z"
-apply (rule bex_Infinitesimal_iff [THEN iffD1])
-apply (drule Infinitesimal_minus_iff [THEN iffD2])
-apply (auto simp add: add.assoc [symmetric])
-done
+lemma Infinitesimal_add_approx: "y \<in> Infinitesimal \<Longrightarrow> x + y = z \<Longrightarrow> x \<approx> z"
+  apply (rule bex_Infinitesimal_iff [THEN iffD1])
+  apply (drule Infinitesimal_minus_iff [THEN iffD2])
+  apply (auto simp add: add.assoc [symmetric])
+  done
 
-lemma Infinitesimal_add_approx_self: "y \<in> Infinitesimal ==> x \<approx> x + y"
-apply (rule bex_Infinitesimal_iff [THEN iffD1])
-apply (drule Infinitesimal_minus_iff [THEN iffD2])
-apply (auto simp add: add.assoc [symmetric])
-done
+lemma Infinitesimal_add_approx_self: "y \<in> Infinitesimal \<Longrightarrow> x \<approx> x + y"
+  apply (rule bex_Infinitesimal_iff [THEN iffD1])
+  apply (drule Infinitesimal_minus_iff [THEN iffD2])
+  apply (auto simp add: add.assoc [symmetric])
+  done
 
-lemma Infinitesimal_add_approx_self2: "y \<in> Infinitesimal ==> x \<approx> y + x"
-by (auto dest: Infinitesimal_add_approx_self simp add: add.commute)
+lemma Infinitesimal_add_approx_self2: "y \<in> Infinitesimal \<Longrightarrow> x \<approx> y + x"
+  by (auto dest: Infinitesimal_add_approx_self simp add: add.commute)
 
-lemma Infinitesimal_add_minus_approx_self: "y \<in> Infinitesimal ==> x \<approx> x + -y"
-by (blast intro!: Infinitesimal_add_approx_self Infinitesimal_minus_iff [THEN iffD2])
+lemma Infinitesimal_add_minus_approx_self: "y \<in> Infinitesimal \<Longrightarrow> x \<approx> x + - y"
+  by (blast intro!: Infinitesimal_add_approx_self Infinitesimal_minus_iff [THEN iffD2])
 
-lemma Infinitesimal_add_cancel: "[| y \<in> Infinitesimal; x+y \<approx> z|] ==> x \<approx> z"
-apply (drule_tac x = x in Infinitesimal_add_approx_self [THEN approx_sym])
-apply (erule approx_trans3 [THEN approx_sym], assumption)
-done
+lemma Infinitesimal_add_cancel: "y \<in> Infinitesimal \<Longrightarrow> x + y \<approx> z \<Longrightarrow> x \<approx> z"
+  apply (drule_tac x = x in Infinitesimal_add_approx_self [THEN approx_sym])
+  apply (erule approx_trans3 [THEN approx_sym], assumption)
+  done
 
-lemma Infinitesimal_add_right_cancel:
-     "[| y \<in> Infinitesimal; x \<approx> z + y|] ==> x \<approx> z"
-apply (drule_tac x = z in Infinitesimal_add_approx_self2 [THEN approx_sym])
-apply (erule approx_trans3 [THEN approx_sym])
-apply (simp add: add.commute)
-apply (erule approx_sym)
-done
+lemma Infinitesimal_add_right_cancel: "y \<in> Infinitesimal \<Longrightarrow> x \<approx> z + y \<Longrightarrow> x \<approx> z"
+  apply (drule_tac x = z in Infinitesimal_add_approx_self2 [THEN approx_sym])
+  apply (erule approx_trans3 [THEN approx_sym])
+  apply (simp add: add.commute)
+  apply (erule approx_sym)
+  done
 
-lemma approx_add_left_cancel: "d + b  \<approx> d + c ==> b \<approx> c"
-apply (drule approx_minus_iff [THEN iffD1])
-apply (simp add: approx_minus_iff [symmetric] ac_simps)
-done
+lemma approx_add_left_cancel: "d + b  \<approx> d + c \<Longrightarrow> b \<approx> c"
+  apply (drule approx_minus_iff [THEN iffD1])
+  apply (simp add: approx_minus_iff [symmetric] ac_simps)
+  done
 
-lemma approx_add_right_cancel: "b + d \<approx> c + d ==> b \<approx> c"
-apply (rule approx_add_left_cancel)
-apply (simp add: add.commute)
-done
+lemma approx_add_right_cancel: "b + d \<approx> c + d \<Longrightarrow> b \<approx> c"
+  apply (rule approx_add_left_cancel)
+  apply (simp add: add.commute)
+  done
 
-lemma approx_add_mono1: "b \<approx> c ==> d + b \<approx> d + c"
-apply (rule approx_minus_iff [THEN iffD2])
-apply (simp add: approx_minus_iff [symmetric] ac_simps)
-done
+lemma approx_add_mono1: "b \<approx> c \<Longrightarrow> d + b \<approx> d + c"
+  apply (rule approx_minus_iff [THEN iffD2])
+  apply (simp add: approx_minus_iff [symmetric] ac_simps)
+  done
 
-lemma approx_add_mono2: "b \<approx> c ==> b + a \<approx> c + a"
-by (simp add: add.commute approx_add_mono1)
+lemma approx_add_mono2: "b \<approx> c \<Longrightarrow> b + a \<approx> c + a"
+  by (simp add: add.commute approx_add_mono1)
 
-lemma approx_add_left_iff [simp]: "(a + b \<approx> a + c) = (b \<approx> c)"
-by (fast elim: approx_add_left_cancel approx_add_mono1)
+lemma approx_add_left_iff [simp]: "a + b \<approx> a + c \<longleftrightarrow> b \<approx> c"
+  by (fast elim: approx_add_left_cancel approx_add_mono1)
 
-lemma approx_add_right_iff [simp]: "(b + a \<approx> c + a) = (b \<approx> c)"
-by (simp add: add.commute)
+lemma approx_add_right_iff [simp]: "b + a \<approx> c + a \<longleftrightarrow> b \<approx> c"
+  by (simp add: add.commute)
 
-lemma approx_HFinite: "[| x \<in> HFinite; x \<approx> y |] ==> y \<in> HFinite"
-apply (drule bex_Infinitesimal_iff2 [THEN iffD2], safe)
-apply (drule Infinitesimal_subset_HFinite [THEN subsetD, THEN HFinite_minus_iff [THEN iffD2]])
-apply (drule HFinite_add)
-apply (auto simp add: add.assoc)
-done
+lemma approx_HFinite: "x \<in> HFinite \<Longrightarrow> x \<approx> y \<Longrightarrow> y \<in> HFinite"
+  apply (drule bex_Infinitesimal_iff2 [THEN iffD2], safe)
+  apply (drule Infinitesimal_subset_HFinite [THEN subsetD, THEN HFinite_minus_iff [THEN iffD2]])
+  apply (drule HFinite_add)
+   apply (auto simp add: add.assoc)
+  done
 
-lemma approx_star_of_HFinite: "x \<approx> star_of D ==> x \<in> HFinite"
-by (rule approx_sym [THEN [2] approx_HFinite], auto)
+lemma approx_star_of_HFinite: "x \<approx> star_of D \<Longrightarrow> x \<in> HFinite"
+  by (rule approx_sym [THEN [2] approx_HFinite], auto)
 
-lemma approx_mult_HFinite:
-  fixes a b c d :: "'a::real_normed_algebra star"
-  shows "[|a \<approx> b; c \<approx> d; b: HFinite; d: HFinite|] ==> a*c \<approx> b*d"
-apply (rule approx_trans)
-apply (rule_tac [2] approx_mult2)
-apply (rule approx_mult1)
-prefer 2 apply (blast intro: approx_HFinite approx_sym, auto)
-done
+lemma approx_mult_HFinite: "a \<approx> b \<Longrightarrow> c \<approx> d \<Longrightarrow> b \<in> HFinite \<Longrightarrow> d \<in> HFinite \<Longrightarrow> a * c \<approx> b * d"
+  for a b c d :: "'a::real_normed_algebra star"
+  apply (rule approx_trans)
+   apply (rule_tac [2] approx_mult2)
+    apply (rule approx_mult1)
+     prefer 2 apply (blast intro: approx_HFinite approx_sym, auto)
+  done
 
-lemma scaleHR_left_diff_distrib:
-  "\<And>a b x. scaleHR (a - b) x = scaleHR a x - scaleHR b x"
-by transfer (rule scaleR_left_diff_distrib)
+lemma scaleHR_left_diff_distrib: "\<And>a b x. scaleHR (a - b) x = scaleHR a x - scaleHR b x"
+  by transfer (rule scaleR_left_diff_distrib)
 
-lemma approx_scaleR1:
-  "[| a \<approx> star_of b; c: HFinite|] ==> scaleHR a c \<approx> b *\<^sub>R c"
-apply (unfold approx_def)
-apply (drule (1) Infinitesimal_HFinite_scaleHR)
-apply (simp only: scaleHR_left_diff_distrib)
-apply (simp add: scaleHR_def star_scaleR_def [symmetric])
-done
+lemma approx_scaleR1: "a \<approx> star_of b \<Longrightarrow> c \<in> HFinite \<Longrightarrow> scaleHR a c \<approx> b *\<^sub>R c"
+  apply (unfold approx_def)
+  apply (drule (1) Infinitesimal_HFinite_scaleHR)
+  apply (simp only: scaleHR_left_diff_distrib)
+  apply (simp add: scaleHR_def star_scaleR_def [symmetric])
+  done
 
-lemma approx_scaleR2:
-  "a \<approx> b ==> c *\<^sub>R a \<approx> c *\<^sub>R b"
-by (simp add: approx_def Infinitesimal_scaleR2
-              scaleR_right_diff_distrib [symmetric])
+lemma approx_scaleR2: "a \<approx> b \<Longrightarrow> c *\<^sub>R a \<approx> c *\<^sub>R b"
+  by (simp add: approx_def Infinitesimal_scaleR2 scaleR_right_diff_distrib [symmetric])
+
+lemma approx_scaleR_HFinite: "a \<approx> star_of b \<Longrightarrow> c \<approx> d \<Longrightarrow> d \<in> HFinite \<Longrightarrow> scaleHR a c \<approx> b *\<^sub>R d"
+  apply (rule approx_trans)
+   apply (rule_tac [2] approx_scaleR2)
+   apply (rule approx_scaleR1)
+    prefer 2 apply (blast intro: approx_HFinite approx_sym, auto)
+  done
 
-lemma approx_scaleR_HFinite:
-  "[|a \<approx> star_of b; c \<approx> d; d: HFinite|] ==> scaleHR a c \<approx> b *\<^sub>R d"
-apply (rule approx_trans)
-apply (rule_tac [2] approx_scaleR2)
-apply (rule approx_scaleR1)
-prefer 2 apply (blast intro: approx_HFinite approx_sym, auto)
-done
+lemma approx_mult_star_of: "a \<approx> star_of b \<Longrightarrow> c \<approx> star_of d \<Longrightarrow> a * c \<approx> star_of b * star_of d"
+  for a c :: "'a::real_normed_algebra star"
+  by (blast intro!: approx_mult_HFinite approx_star_of_HFinite HFinite_star_of)
+
+lemma approx_SReal_mult_cancel_zero: "a \<in> \<real> \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> a * x \<approx> 0 \<Longrightarrow> x \<approx> 0"
+  for a x :: hypreal
+  apply (drule Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]])
+  apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric])
+  done
 
-lemma approx_mult_star_of:
-  fixes a c :: "'a::real_normed_algebra star"
-  shows "[|a \<approx> star_of b; c \<approx> star_of d |]
-      ==> a*c \<approx> star_of b*star_of d"
-by (blast intro!: approx_mult_HFinite approx_star_of_HFinite HFinite_star_of)
+lemma approx_mult_SReal1: "a \<in> \<real> \<Longrightarrow> x \<approx> 0 \<Longrightarrow> x * a \<approx> 0"
+  for a x :: hypreal
+  by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult1)
 
-lemma approx_SReal_mult_cancel_zero:
-     "[| (a::hypreal) \<in> \<real>; a \<noteq> 0; a*x \<approx> 0 |] ==> x \<approx> 0"
-apply (drule Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]])
-apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric])
-done
+lemma approx_mult_SReal2: "a \<in> \<real> \<Longrightarrow> x \<approx> 0 \<Longrightarrow> a * x \<approx> 0"
+  for a x :: hypreal
+  by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult2)
 
-lemma approx_mult_SReal1: "[| (a::hypreal) \<in> \<real>; x \<approx> 0 |] ==> x*a \<approx> 0"
-by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult1)
-
-lemma approx_mult_SReal2: "[| (a::hypreal) \<in> \<real>; x \<approx> 0 |] ==> a*x \<approx> 0"
-by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult2)
+lemma approx_mult_SReal_zero_cancel_iff [simp]: "a \<in> \<real> \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> a * x \<approx> 0 \<longleftrightarrow> x \<approx> 0"
+  for a x :: hypreal
+  by (blast intro: approx_SReal_mult_cancel_zero approx_mult_SReal2)
 
-lemma approx_mult_SReal_zero_cancel_iff [simp]:
-     "[|(a::hypreal) \<in> \<real>; a \<noteq> 0 |] ==> (a*x \<approx> 0) = (x \<approx> 0)"
-by (blast intro: approx_SReal_mult_cancel_zero approx_mult_SReal2)
+lemma approx_SReal_mult_cancel: "a \<in> \<real> \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> a * w \<approx> a * z \<Longrightarrow> w \<approx> z"
+  for a w z :: hypreal
+  apply (drule Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]])
+  apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric])
+  done
 
-lemma approx_SReal_mult_cancel:
-     "[| (a::hypreal) \<in> \<real>; a \<noteq> 0; a* w \<approx> a*z |] ==> w \<approx> z"
-apply (drule Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]])
-apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric])
-done
+lemma approx_SReal_mult_cancel_iff1 [simp]: "a \<in> \<real> \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> a * w \<approx> a * z \<longleftrightarrow> w \<approx> z"
+  for a w z :: hypreal
+  by (auto intro!: approx_mult2 SReal_subset_HFinite [THEN subsetD]
+      intro: approx_SReal_mult_cancel)
 
-lemma approx_SReal_mult_cancel_iff1 [simp]:
-     "[| (a::hypreal) \<in> \<real>; a \<noteq> 0|] ==> (a* w \<approx> a*z) = (w \<approx> z)"
-by (auto intro!: approx_mult2 SReal_subset_HFinite [THEN subsetD]
-         intro: approx_SReal_mult_cancel)
+lemma approx_le_bound: "z \<le> f \<Longrightarrow> f \<approx> g \<Longrightarrow> g \<le> z ==> f \<approx> z"
+  for z :: hypreal
+  apply (simp add: bex_Infinitesimal_iff2 [symmetric], auto)
+  apply (rule_tac x = "g + y - z" in bexI)
+   apply simp
+  apply (rule Infinitesimal_interval2)
+     apply (rule_tac [2] Infinitesimal_zero, auto)
+  done
 
-lemma approx_le_bound: "[| (z::hypreal) \<le> f; f \<approx> g; g \<le> z |] ==> f \<approx> z"
-apply (simp add: bex_Infinitesimal_iff2 [symmetric], auto)
-apply (rule_tac x = "g+y-z" in bexI)
-apply (simp (no_asm))
-apply (rule Infinitesimal_interval2)
-apply (rule_tac [2] Infinitesimal_zero, auto)
-done
-
-lemma approx_hnorm:
-  fixes x y :: "'a::real_normed_vector star"
-  shows "x \<approx> y \<Longrightarrow> hnorm x \<approx> hnorm y"
+lemma approx_hnorm: "x \<approx> y \<Longrightarrow> hnorm x \<approx> hnorm y"
+  for x y :: "'a::real_normed_vector star"
 proof (unfold approx_def)
   assume "x - y \<in> Infinitesimal"
-  hence 1: "hnorm (x - y) \<in> Infinitesimal"
+  then have "hnorm (x - y) \<in> Infinitesimal"
     by (simp only: Infinitesimal_hnorm_iff)
-  moreover have 2: "(0::real star) \<in> Infinitesimal"
+  moreover have "(0::real star) \<in> Infinitesimal"
     by (rule Infinitesimal_zero)
-  moreover have 3: "0 \<le> \<bar>hnorm x - hnorm y\<bar>"
+  moreover have "0 \<le> \<bar>hnorm x - hnorm y\<bar>"
     by (rule abs_ge_zero)
-  moreover have 4: "\<bar>hnorm x - hnorm y\<bar> \<le> hnorm (x - y)"
+  moreover have "\<bar>hnorm x - hnorm y\<bar> \<le> hnorm (x - y)"
     by (rule hnorm_triangle_ineq3)
   ultimately have "\<bar>hnorm x - hnorm y\<bar> \<in> Infinitesimal"
     by (rule Infinitesimal_interval2)
-  thus "hnorm x - hnorm y \<in> Infinitesimal"
+  then show "hnorm x - hnorm y \<in> Infinitesimal"
     by (simp only: Infinitesimal_hrabs_iff)
 qed
 
 
-subsection\<open>Zero is the Only Infinitesimal that is also a Real\<close>
+subsection \<open>Zero is the Only Infinitesimal that is also a Real\<close>
 
-lemma Infinitesimal_less_SReal:
-     "[| (x::hypreal) \<in> \<real>; y \<in> Infinitesimal; 0 < x |] ==> y < x"
-apply (simp add: Infinitesimal_def)
-apply (rule abs_ge_self [THEN order_le_less_trans], auto)
-done
+lemma Infinitesimal_less_SReal: "x \<in> \<real> \<Longrightarrow> y \<in> Infinitesimal \<Longrightarrow> 0 < x \<Longrightarrow> y < x"
+  for x y :: hypreal
+  apply (simp add: Infinitesimal_def)
+  apply (rule abs_ge_self [THEN order_le_less_trans], auto)
+  done
 
-lemma Infinitesimal_less_SReal2:
-     "(y::hypreal) \<in> Infinitesimal ==> \<forall>r \<in> Reals. 0 < r --> y < r"
-by (blast intro: Infinitesimal_less_SReal)
+lemma Infinitesimal_less_SReal2: "y \<in> Infinitesimal \<Longrightarrow> \<forall>r \<in> Reals. 0 < r \<longrightarrow> y < r"
+  for y :: hypreal
+  by (blast intro: Infinitesimal_less_SReal)
 
-lemma SReal_not_Infinitesimal:
-     "[| 0 < y;  (y::hypreal) \<in> \<real>|] ==> y \<notin> Infinitesimal"
-apply (simp add: Infinitesimal_def)
-apply (auto simp add: abs_if)
-done
+lemma SReal_not_Infinitesimal: "0 < y \<Longrightarrow> y \<in> \<real> ==> y \<notin> Infinitesimal"
+  for y :: hypreal
+  apply (simp add: Infinitesimal_def)
+  apply (auto simp add: abs_if)
+  done
 
-lemma SReal_minus_not_Infinitesimal:
-     "[| y < 0;  (y::hypreal) \<in> \<real> |] ==> y \<notin> Infinitesimal"
-apply (subst Infinitesimal_minus_iff [symmetric])
-apply (rule SReal_not_Infinitesimal, auto)
-done
+lemma SReal_minus_not_Infinitesimal: "y < 0 \<Longrightarrow> y \<in> \<real> \<Longrightarrow> y \<notin> Infinitesimal"
+  for y :: hypreal
+  apply (subst Infinitesimal_minus_iff [symmetric])
+  apply (rule SReal_not_Infinitesimal, auto)
+  done
 
 lemma SReal_Int_Infinitesimal_zero: "\<real> Int Infinitesimal = {0::hypreal}"
-apply auto
-apply (cut_tac x = x and y = 0 in linorder_less_linear)
-apply (blast dest: SReal_not_Infinitesimal SReal_minus_not_Infinitesimal)
-done
+  apply auto
+  apply (cut_tac x = x and y = 0 in linorder_less_linear)
+  apply (blast dest: SReal_not_Infinitesimal SReal_minus_not_Infinitesimal)
+  done
 
-lemma SReal_Infinitesimal_zero:
-  "[| (x::hypreal) \<in> \<real>; x \<in> Infinitesimal|] ==> x = 0"
-by (cut_tac SReal_Int_Infinitesimal_zero, blast)
+lemma SReal_Infinitesimal_zero: "x \<in> \<real> \<Longrightarrow> x \<in> Infinitesimal \<Longrightarrow> x = 0"
+  for x :: hypreal
+  using SReal_Int_Infinitesimal_zero by blast
 
-lemma SReal_HFinite_diff_Infinitesimal:
-     "[| (x::hypreal) \<in> \<real>; x \<noteq> 0 |] ==> x \<in> HFinite - Infinitesimal"
-by (auto dest: SReal_Infinitesimal_zero SReal_subset_HFinite [THEN subsetD])
+lemma SReal_HFinite_diff_Infinitesimal: "x \<in> \<real> \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> x \<in> HFinite - Infinitesimal"
+  for x :: hypreal
+  by (auto dest: SReal_Infinitesimal_zero SReal_subset_HFinite [THEN subsetD])
 
 lemma hypreal_of_real_HFinite_diff_Infinitesimal:
-     "hypreal_of_real x \<noteq> 0 ==> hypreal_of_real x \<in> HFinite - Infinitesimal"
-by (rule SReal_HFinite_diff_Infinitesimal, auto)
+  "hypreal_of_real x \<noteq> 0 \<Longrightarrow> hypreal_of_real x \<in> HFinite - Infinitesimal"
+  by (rule SReal_HFinite_diff_Infinitesimal) auto
 
-lemma star_of_Infinitesimal_iff_0 [iff]:
-  "(star_of x \<in> Infinitesimal) = (x = 0)"
-apply (auto simp add: Infinitesimal_def)
-apply (drule_tac x="hnorm (star_of x)" in bspec)
-apply (simp add: SReal_def)
-apply (rule_tac x="norm x" in exI, simp)
-apply simp
-done
+lemma star_of_Infinitesimal_iff_0 [iff]: "star_of x \<in> Infinitesimal \<longleftrightarrow> x = 0"
+  apply (auto simp add: Infinitesimal_def)
+  apply (drule_tac x="hnorm (star_of x)" in bspec)
+   apply (simp add: SReal_def)
+   apply (rule_tac x="norm x" in exI, simp)
+  apply simp
+  done
 
-lemma star_of_HFinite_diff_Infinitesimal:
-     "x \<noteq> 0 ==> star_of x \<in> HFinite - Infinitesimal"
-by simp
+lemma star_of_HFinite_diff_Infinitesimal: "x \<noteq> 0 \<Longrightarrow> star_of x \<in> HFinite - Infinitesimal"
+  by simp
 
 lemma numeral_not_Infinitesimal [simp]:
-     "numeral w \<noteq> (0::hypreal) ==> (numeral w :: hypreal) \<notin> Infinitesimal"
-by (fast dest: Reals_numeral [THEN SReal_Infinitesimal_zero])
+  "numeral w \<noteq> (0::hypreal) \<Longrightarrow> (numeral w :: hypreal) \<notin> Infinitesimal"
+  by (fast dest: Reals_numeral [THEN SReal_Infinitesimal_zero])
 
-(*again: 1 is a special case, but not 0 this time*)
+text \<open>Again: \<open>1\<close> is a special case, but not \<open>0\<close> this time.\<close>
 lemma one_not_Infinitesimal [simp]:
   "(1::'a::{real_normed_vector,zero_neq_one} star) \<notin> Infinitesimal"
-apply (simp only: star_one_def star_of_Infinitesimal_iff_0)
-apply simp
-done
+  apply (simp only: star_one_def star_of_Infinitesimal_iff_0)
+  apply simp
+  done
 
-lemma approx_SReal_not_zero:
-  "[| (y::hypreal) \<in> \<real>; x \<approx> y; y\<noteq> 0 |] ==> x \<noteq> 0"
-apply (cut_tac x = 0 and y = y in linorder_less_linear, simp)
-apply (blast dest: approx_sym [THEN mem_infmal_iff [THEN iffD2]] SReal_not_Infinitesimal SReal_minus_not_Infinitesimal)
-done
+lemma approx_SReal_not_zero: "y \<in> \<real> \<Longrightarrow> x \<approx> y \<Longrightarrow> y \<noteq> 0 \<Longrightarrow> x \<noteq> 0"
+  for x y :: hypreal
+  apply (cut_tac x = 0 and y = y in linorder_less_linear, simp)
+  apply (blast dest: approx_sym [THEN mem_infmal_iff [THEN iffD2]]
+      SReal_not_Infinitesimal SReal_minus_not_Infinitesimal)
+  done
 
 lemma HFinite_diff_Infinitesimal_approx:
-     "[| x \<approx> y; y \<in> HFinite - Infinitesimal |]
-      ==> x \<in> HFinite - Infinitesimal"
-apply (auto intro: approx_sym [THEN [2] approx_HFinite]
-            simp add: mem_infmal_iff)
-apply (drule approx_trans3, assumption)
-apply (blast dest: approx_sym)
-done
+  "x \<approx> y \<Longrightarrow> y \<in> HFinite - Infinitesimal \<Longrightarrow> x \<in> HFinite - Infinitesimal"
+  apply (auto intro: approx_sym [THEN [2] approx_HFinite] simp: mem_infmal_iff)
+  apply (drule approx_trans3, assumption)
+  apply (blast dest: approx_sym)
+  done
 
-(*The premise y\<noteq>0 is essential; otherwise x/y =0 and we lose the
-  HFinite premise.*)
+text \<open>The premise \<open>y \<noteq> 0\<close> is essential; otherwise \<open>x / y = 0\<close> and we lose the
+  \<open>HFinite\<close> premise.\<close>
 lemma Infinitesimal_ratio:
-  fixes x y :: "'a::{real_normed_div_algebra,field} star"
-  shows "[| y \<noteq> 0;  y \<in> Infinitesimal;  x/y \<in> HFinite |]
-         ==> x \<in> Infinitesimal"
-apply (drule Infinitesimal_HFinite_mult2, assumption)
-apply (simp add: divide_inverse mult.assoc)
-done
+  "y \<noteq> 0 \<Longrightarrow> y \<in> Infinitesimal \<Longrightarrow> x / y \<in> HFinite \<Longrightarrow> x \<in> Infinitesimal"
+  for x y :: "'a::{real_normed_div_algebra,field} star"
+  apply (drule Infinitesimal_HFinite_mult2, assumption)
+  apply (simp add: divide_inverse mult.assoc)
+  done
+
+lemma Infinitesimal_SReal_divide: "x \<in> Infinitesimal \<Longrightarrow> y \<in> \<real> \<Longrightarrow> x / y \<in> Infinitesimal"
+  for x y :: hypreal
+  apply (simp add: divide_inverse)
+  apply (auto intro!: Infinitesimal_HFinite_mult
+      dest!: Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]])
+  done
+
+
+section \<open>Standard Part Theorem\<close>
 
-lemma Infinitesimal_SReal_divide:
-  "[| (x::hypreal) \<in> Infinitesimal; y \<in> \<real> |] ==> x/y \<in> Infinitesimal"
-apply (simp add: divide_inverse)
-apply (auto intro!: Infinitesimal_HFinite_mult
-            dest!: Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]])
-done
+text \<open>
+  Every finite \<open>x \<in> R*\<close> is infinitely close to a unique real number
+  (i.e. a member of \<open>Reals\<close>).
+\<close>
 
-(*------------------------------------------------------------------
-       Standard Part Theorem: Every finite x: R* is infinitely
-       close to a unique real number (i.e a member of Reals)
- ------------------------------------------------------------------*)
 
-subsection\<open>Uniqueness: Two Infinitely Close Reals are Equal\<close>
+subsection \<open>Uniqueness: Two Infinitely Close Reals are Equal\<close>
 
-lemma star_of_approx_iff [simp]: "(star_of x \<approx> star_of y) = (x = y)"
-apply safe
-apply (simp add: approx_def)
-apply (simp only: star_of_diff [symmetric])
-apply (simp only: star_of_Infinitesimal_iff_0)
-apply simp
-done
+lemma star_of_approx_iff [simp]: "star_of x \<approx> star_of y \<longleftrightarrow> x = y"
+  apply safe
+  apply (simp add: approx_def)
+  apply (simp only: star_of_diff [symmetric])
+  apply (simp only: star_of_Infinitesimal_iff_0)
+  apply simp
+  done
 
-lemma SReal_approx_iff:
-  "[|(x::hypreal) \<in> \<real>; y \<in> \<real>|] ==> (x \<approx> y) = (x = y)"
-apply auto
-apply (simp add: approx_def)
-apply (drule (1) Reals_diff)
-apply (drule (1) SReal_Infinitesimal_zero)
-apply simp
-done
+lemma SReal_approx_iff: "x \<in> \<real> \<Longrightarrow> y \<in> \<real> \<Longrightarrow> x \<approx> y \<longleftrightarrow> x = y"
+  for x y :: hypreal
+  apply auto
+  apply (simp add: approx_def)
+  apply (drule (1) Reals_diff)
+  apply (drule (1) SReal_Infinitesimal_zero)
+  apply simp
+  done
 
 lemma numeral_approx_iff [simp]:
-     "(numeral v \<approx> (numeral w :: 'a::{numeral,real_normed_vector} star)) =
-      (numeral v = (numeral w :: 'a))"
-apply (unfold star_numeral_def)
-apply (rule star_of_approx_iff)
-done
+  "(numeral v \<approx> (numeral w :: 'a::{numeral,real_normed_vector} star)) =
+    (numeral v = (numeral w :: 'a))"
+  apply (unfold star_numeral_def)
+  apply (rule star_of_approx_iff)
+  done
 
-(*And also for 0 \<approx> #nn and 1 \<approx> #nn, #nn \<approx> 0 and #nn \<approx> 1.*)
+text \<open>And also for \<open>0 \<approx> #nn\<close> and \<open>1 \<approx> #nn\<close>, \<open>#nn \<approx> 0\<close> and \<open>#nn \<approx> 1\<close>.\<close>
 lemma [simp]:
-  "(numeral w \<approx> (0::'a::{numeral,real_normed_vector} star)) =
-   (numeral w = (0::'a))"
-  "((0::'a::{numeral,real_normed_vector} star) \<approx> numeral w) =
-   (numeral w = (0::'a))"
-  "(numeral w \<approx> (1::'b::{numeral,one,real_normed_vector} star)) =
-   (numeral w = (1::'b))"
-  "((1::'b::{numeral,one,real_normed_vector} star) \<approx> numeral w) =
-   (numeral w = (1::'b))"
-  "~ (0 \<approx> (1::'c::{zero_neq_one,real_normed_vector} star))"
-  "~ (1 \<approx> (0::'c::{zero_neq_one,real_normed_vector} star))"
-apply (unfold star_numeral_def star_zero_def star_one_def)
-apply (unfold star_of_approx_iff)
-by (auto intro: sym)
+  "(numeral w \<approx> (0::'a::{numeral,real_normed_vector} star)) = (numeral w = (0::'a))"
+  "((0::'a::{numeral,real_normed_vector} star) \<approx> numeral w) = (numeral w = (0::'a))"
+  "(numeral w \<approx> (1::'b::{numeral,one,real_normed_vector} star)) = (numeral w = (1::'b))"
+  "((1::'b::{numeral,one,real_normed_vector} star) \<approx> numeral w) = (numeral w = (1::'b))"
+  "\<not> (0 \<approx> (1::'c::{zero_neq_one,real_normed_vector} star))"
+  "\<not> (1 \<approx> (0::'c::{zero_neq_one,real_normed_vector} star))"
+       apply (unfold star_numeral_def star_zero_def star_one_def)
+       apply (unfold star_of_approx_iff)
+       apply (auto intro: sym)
+  done
 
-lemma star_of_approx_numeral_iff [simp]:
-     "(star_of k \<approx> numeral w) = (k = numeral w)"
-by (subst star_of_approx_iff [symmetric], auto)
+lemma star_of_approx_numeral_iff [simp]: "star_of k \<approx> numeral w \<longleftrightarrow> k = numeral w"
+  by (subst star_of_approx_iff [symmetric]) auto
 
-lemma star_of_approx_zero_iff [simp]: "(star_of k \<approx> 0) = (k = 0)"
-by (simp_all add: star_of_approx_iff [symmetric])
+lemma star_of_approx_zero_iff [simp]: "star_of k \<approx> 0 \<longleftrightarrow> k = 0"
+  by (simp_all add: star_of_approx_iff [symmetric])
 
-lemma star_of_approx_one_iff [simp]: "(star_of k \<approx> 1) = (k = 1)"
-by (simp_all add: star_of_approx_iff [symmetric])
+lemma star_of_approx_one_iff [simp]: "star_of k \<approx> 1 \<longleftrightarrow> k = 1"
+  by (simp_all add: star_of_approx_iff [symmetric])
 
-lemma approx_unique_real:
-     "[| (r::hypreal) \<in> \<real>; s \<in> \<real>; r \<approx> x; s \<approx> x|] ==> r = s"
-by (blast intro: SReal_approx_iff [THEN iffD1] approx_trans2)
+lemma approx_unique_real: "r \<in> \<real> \<Longrightarrow> s \<in> \<real> \<Longrightarrow> r \<approx> x \<Longrightarrow> s \<approx> x \<Longrightarrow> r = s"
+  for r s :: hypreal
+  by (blast intro: SReal_approx_iff [THEN iffD1] approx_trans2)
 
 
-subsection\<open>Existence of Unique Real Infinitely Close\<close>
+subsection \<open>Existence of Unique Real Infinitely Close\<close>
 
-subsubsection\<open>Lifting of the Ub and Lub Properties\<close>
+subsubsection \<open>Lifting of the Ub and Lub Properties\<close>
 
-lemma hypreal_of_real_isUb_iff:
-      "(isUb \<real> (hypreal_of_real ` Q) (hypreal_of_real Y)) =
-       (isUb (UNIV :: real set) Q Y)"
-by (simp add: isUb_def setle_def)
+lemma hypreal_of_real_isUb_iff: "isUb \<real> (hypreal_of_real ` Q) (hypreal_of_real Y) = isUb UNIV Q Y"
+  for Q :: "real set" and Y :: real
+  by (simp add: isUb_def setle_def)
 
-lemma hypreal_of_real_isLub1:
-     "isLub \<real> (hypreal_of_real ` Q) (hypreal_of_real Y)
-      ==> isLub (UNIV :: real set) Q Y"
-apply (simp add: isLub_def leastP_def)
-apply (auto intro: hypreal_of_real_isUb_iff [THEN iffD2]
-            simp add: hypreal_of_real_isUb_iff setge_def)
-done
+lemma hypreal_of_real_isLub1: "isLub \<real> (hypreal_of_real ` Q) (hypreal_of_real Y) \<Longrightarrow> isLub UNIV Q Y"
+  for Q :: "real set" and Y :: real
+  apply (simp add: isLub_def leastP_def)
+  apply (auto intro: hypreal_of_real_isUb_iff [THEN iffD2]
+      simp add: hypreal_of_real_isUb_iff setge_def)
+  done
 
-lemma hypreal_of_real_isLub2:
-      "isLub (UNIV :: real set) Q Y
-       ==> isLub \<real> (hypreal_of_real ` Q) (hypreal_of_real Y)"
-apply (auto simp add: isLub_def leastP_def hypreal_of_real_isUb_iff setge_def)
-by (metis SReal_iff hypreal_of_real_isUb_iff isUbD2a star_of_le)
+lemma hypreal_of_real_isLub2: "isLub UNIV Q Y \<Longrightarrow> isLub \<real> (hypreal_of_real ` Q) (hypreal_of_real Y)"
+  for Q :: "real set" and Y :: real
+  apply (auto simp add: isLub_def leastP_def hypreal_of_real_isUb_iff setge_def)
+  apply (metis SReal_iff hypreal_of_real_isUb_iff isUbD2a star_of_le)
+  done
 
 lemma hypreal_of_real_isLub_iff:
-     "(isLub \<real> (hypreal_of_real ` Q) (hypreal_of_real Y)) =
-      (isLub (UNIV :: real set) Q Y)"
-by (blast intro: hypreal_of_real_isLub1 hypreal_of_real_isLub2)
+  "isLub \<real> (hypreal_of_real ` Q) (hypreal_of_real Y) = isLub (UNIV :: real set) Q Y"
+  for Q :: "real set" and Y :: real
+  by (blast intro: hypreal_of_real_isLub1 hypreal_of_real_isLub2)
 
-lemma lemma_isUb_hypreal_of_real:
-     "isUb \<real> P Y ==> \<exists>Yo. isUb \<real> P (hypreal_of_real Yo)"
-by (auto simp add: SReal_iff isUb_def)
+lemma lemma_isUb_hypreal_of_real: "isUb \<real> P Y \<Longrightarrow> \<exists>Yo. isUb \<real> P (hypreal_of_real Yo)"
+  by (auto simp add: SReal_iff isUb_def)
+
+lemma lemma_isLub_hypreal_of_real: "isLub \<real> P Y \<Longrightarrow> \<exists>Yo. isLub \<real> P (hypreal_of_real Yo)"
+  by (auto simp add: isLub_def leastP_def isUb_def SReal_iff)
 
-lemma lemma_isLub_hypreal_of_real:
-     "isLub \<real> P Y ==> \<exists>Yo. isLub \<real> P (hypreal_of_real Yo)"
-by (auto simp add: isLub_def leastP_def isUb_def SReal_iff)
+lemma lemma_isLub_hypreal_of_real2: "\<exists>Yo. isLub \<real> P (hypreal_of_real Yo) \<Longrightarrow> \<exists>Y. isLub \<real> P Y"
+  by (auto simp add: isLub_def leastP_def isUb_def)
 
-lemma lemma_isLub_hypreal_of_real2:
-     "\<exists>Yo. isLub \<real> P (hypreal_of_real Yo) ==> \<exists>Y. isLub \<real> P Y"
-by (auto simp add: isLub_def leastP_def isUb_def)
+lemma SReal_complete: "P \<subseteq> \<real> \<Longrightarrow> \<exists>x. x \<in> P \<Longrightarrow> \<exists>Y. isUb \<real> P Y \<Longrightarrow> \<exists>t::hypreal. isLub \<real> P t"
+  apply (frule SReal_hypreal_of_real_image)
+   apply (auto, drule lemma_isUb_hypreal_of_real)
+  apply (auto intro!: reals_complete lemma_isLub_hypreal_of_real2
+      simp add: hypreal_of_real_isLub_iff hypreal_of_real_isUb_iff)
+  done
+
 
-lemma SReal_complete:
-     "[| P \<subseteq> \<real>;  \<exists>x. x \<in> P;  \<exists>Y. isUb \<real> P Y |]
-      ==> \<exists>t::hypreal. isLub \<real> P t"
-apply (frule SReal_hypreal_of_real_image)
-apply (auto, drule lemma_isUb_hypreal_of_real)
-apply (auto intro!: reals_complete lemma_isLub_hypreal_of_real2
-            simp add: hypreal_of_real_isLub_iff hypreal_of_real_isUb_iff)
-done
+text \<open>Lemmas about lubs.\<close>
 
-(* lemma about lubs *)
+lemma lemma_st_part_ub: "x \<in> HFinite \<Longrightarrow> \<exists>u. isUb \<real> {s. s \<in> \<real> \<and> s < x} u"
+  for x :: hypreal
+  apply (drule HFiniteD, safe)
+  apply (rule exI, rule isUbI)
+   apply (auto intro: setleI isUbI simp add: abs_less_iff)
+  done
 
-lemma lemma_st_part_ub:
-     "(x::hypreal) \<in> HFinite ==> \<exists>u. isUb \<real> {s. s \<in> \<real> & s < x} u"
-apply (drule HFiniteD, safe)
-apply (rule exI, rule isUbI)
-apply (auto intro: setleI isUbI simp add: abs_less_iff)
-done
+lemma lemma_st_part_nonempty: "x \<in> HFinite \<Longrightarrow> \<exists>y. y \<in> {s. s \<in> \<real> \<and> s < x}"
+  for x :: hypreal
+  apply (drule HFiniteD, safe)
+  apply (drule Reals_minus)
+  apply (rule_tac x = "-t" in exI)
+  apply (auto simp add: abs_less_iff)
+  done
 
-lemma lemma_st_part_nonempty:
-  "(x::hypreal) \<in> HFinite ==> \<exists>y. y \<in> {s. s \<in> \<real> & s < x}"
-apply (drule HFiniteD, safe)
-apply (drule Reals_minus)
-apply (rule_tac x = "-t" in exI)
-apply (auto simp add: abs_less_iff)
-done
-
-lemma lemma_st_part_lub:
-     "(x::hypreal) \<in> HFinite ==> \<exists>t. isLub \<real> {s. s \<in> \<real> & s < x} t"
-by (blast intro!: SReal_complete lemma_st_part_ub lemma_st_part_nonempty Collect_restrict)
+lemma lemma_st_part_lub: "x \<in> HFinite \<Longrightarrow> \<exists>t. isLub \<real> {s. s \<in> \<real> \<and> s < x} t"
+  for x :: hypreal
+  by (blast intro!: SReal_complete lemma_st_part_ub lemma_st_part_nonempty Collect_restrict)
 
 lemma lemma_st_part_le1:
-     "[| (x::hypreal) \<in> HFinite;  isLub \<real> {s. s \<in> \<real> & s < x} t;
-         r \<in> \<real>;  0 < r |] ==> x \<le> t + r"
-apply (frule isLubD1a)
-apply (rule ccontr, drule linorder_not_le [THEN iffD2])
-apply (drule (1) Reals_add)
-apply (drule_tac y = "r + t" in isLubD1 [THEN setleD], auto)
-done
+  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> x \<le> t + r"
+  for x r t :: hypreal
+  apply (frule isLubD1a)
+  apply (rule ccontr, drule linorder_not_le [THEN iffD2])
+  apply (drule (1) Reals_add)
+  apply (drule_tac y = "r + t" in isLubD1 [THEN setleD], auto)
+  done
 
-lemma hypreal_setle_less_trans:
-     "[| S *<= (x::hypreal); x < y |] ==> S *<= y"
-apply (simp add: setle_def)
-apply (auto dest!: bspec order_le_less_trans intro: order_less_imp_le)
-done
+lemma hypreal_setle_less_trans: "S *<= x \<Longrightarrow> x < y \<Longrightarrow> S *<= y"
+  for x y :: hypreal
+  apply (simp add: setle_def)
+  apply (auto dest!: bspec order_le_less_trans intro: order_less_imp_le)
+  done
 
-lemma hypreal_gt_isUb:
-     "[| isUb R S (x::hypreal); x < y; y \<in> R |] ==> isUb R S y"
-apply (simp add: isUb_def)
-apply (blast intro: hypreal_setle_less_trans)
-done
+lemma hypreal_gt_isUb: "isUb R S x \<Longrightarrow> x < y \<Longrightarrow> y \<in> R \<Longrightarrow> isUb R S y"
+  for x y :: hypreal
+  apply (simp add: isUb_def)
+  apply (blast intro: hypreal_setle_less_trans)
+  done
 
-lemma lemma_st_part_gt_ub:
-     "[| (x::hypreal) \<in> HFinite; x < y; y \<in> \<real> |]
-      ==> isUb \<real> {s. s \<in> \<real> & s < x} y"
-by (auto dest: order_less_trans intro: order_less_imp_le intro!: isUbI setleI)
+lemma lemma_st_part_gt_ub: "x \<in> HFinite \<Longrightarrow> x < y \<Longrightarrow> y \<in> \<real> \<Longrightarrow> isUb \<real> {s. s \<in> \<real> \<and> s < x} y"
+  for x y :: hypreal
+  by (auto dest: order_less_trans intro: order_less_imp_le intro!: isUbI setleI)
 
-lemma lemma_minus_le_zero: "t \<le> t + -r ==> r \<le> (0::hypreal)"
-apply (drule_tac c = "-t" in add_left_mono)
-apply (auto simp add: add.assoc [symmetric])
-done
+lemma lemma_minus_le_zero: "t \<le> t + -r \<Longrightarrow> r \<le> 0"
+  for r t :: hypreal
+  apply (drule_tac c = "-t" in add_left_mono)
+  apply (auto simp add: add.assoc [symmetric])
+  done
 
 lemma lemma_st_part_le2:
-     "[| (x::hypreal) \<in> HFinite;
-         isLub \<real> {s. s \<in> \<real> & s < x} t;
-         r \<in> \<real>; 0 < r |]
-      ==> t + -r \<le> x"
-apply (frule isLubD1a)
-apply (rule ccontr, drule linorder_not_le [THEN iffD1])
-apply (drule Reals_minus, drule_tac a = t in Reals_add, assumption)
-apply (drule lemma_st_part_gt_ub, assumption+)
-apply (drule isLub_le_isUb, assumption)
-apply (drule lemma_minus_le_zero)
-apply (auto dest: order_less_le_trans)
-done
+  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> t + -r \<le> x"
+  for x r t :: hypreal
+  apply (frule isLubD1a)
+  apply (rule ccontr, drule linorder_not_le [THEN iffD1])
+  apply (drule Reals_minus, drule_tac a = t in Reals_add, assumption)
+  apply (drule lemma_st_part_gt_ub, assumption+)
+  apply (drule isLub_le_isUb, assumption)
+  apply (drule lemma_minus_le_zero)
+  apply (auto dest: order_less_le_trans)
+  done
 
 lemma lemma_st_part1a:
-     "[| (x::hypreal) \<in> HFinite;
-         isLub \<real> {s. s \<in> \<real> & s < x} t;
-         r \<in> \<real>; 0 < r |]
-      ==> x + -t \<le> r"
-apply (subgoal_tac "x \<le> t+r")
-apply (auto intro: lemma_st_part_le1)
-done
+  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> x + -t \<le> r"
+  for x r t :: hypreal
+  apply (subgoal_tac "x \<le> t + r")
+   apply (auto intro: lemma_st_part_le1)
+  done
 
 lemma lemma_st_part2a:
-     "[| (x::hypreal) \<in> HFinite;
-         isLub \<real> {s. s \<in> \<real> & s < x} t;
-         r \<in> \<real>;  0 < r |]
-      ==> -(x + -t) \<le> r"
-apply (subgoal_tac "(t + -r \<le> x)")
-apply simp
-apply (rule lemma_st_part_le2)
-apply auto
-done
+  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> - (x + -t) \<le> r"
+  for x r t :: hypreal
+  apply (subgoal_tac "(t + -r \<le> x)")
+   apply simp
+  apply (rule lemma_st_part_le2)
+     apply auto
+  done
 
-lemma lemma_SReal_ub:
-     "(x::hypreal) \<in> \<real> ==> isUb \<real> {s. s \<in> \<real> & s < x} x"
-by (auto intro: isUbI setleI order_less_imp_le)
+lemma lemma_SReal_ub: "x \<in> \<real> \<Longrightarrow> isUb \<real> {s. s \<in> \<real> \<and> s < x} x"
+  for x :: hypreal
+  by (auto intro: isUbI setleI order_less_imp_le)
 
-lemma lemma_SReal_lub:
-     "(x::hypreal) \<in> \<real> ==> isLub \<real> {s. s \<in> \<real> & s < x} x"
-apply (auto intro!: isLubI2 lemma_SReal_ub setgeI)
-apply (frule isUbD2a)
-apply (rule_tac x = x and y = y in linorder_cases)
-apply (auto intro!: order_less_imp_le)
-apply (drule SReal_dense, assumption, assumption, safe)
-apply (drule_tac y = r in isUbD)
-apply (auto dest: order_less_le_trans)
-done
+lemma lemma_SReal_lub: "x \<in> \<real> \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} x"
+  for x :: hypreal
+  apply (auto intro!: isLubI2 lemma_SReal_ub setgeI)
+  apply (frule isUbD2a)
+  apply (rule_tac x = x and y = y in linorder_cases)
+    apply (auto intro!: order_less_imp_le)
+  apply (drule SReal_dense, assumption, assumption, safe)
+  apply (drule_tac y = r in isUbD)
+   apply (auto dest: order_less_le_trans)
+  done
 
 lemma lemma_st_part_not_eq1:
-     "[| (x::hypreal) \<in> HFinite;
-         isLub \<real> {s. s \<in> \<real> & s < x} t;
-         r \<in> \<real>; 0 < r |]
-      ==> x + -t \<noteq> r"
-apply auto
-apply (frule isLubD1a [THEN Reals_minus])
-using Reals_add_cancel [of x "- t"] apply simp
-apply (drule_tac x = x in lemma_SReal_lub)
-apply (drule isLub_unique, assumption, auto)
-done
+  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> x + - t \<noteq> r"
+  for x r t :: hypreal
+  apply auto
+  apply (frule isLubD1a [THEN Reals_minus])
+  using Reals_add_cancel [of x "- t"] apply simp
+  apply (drule_tac x = x in lemma_SReal_lub)
+  apply (drule isLub_unique, assumption, auto)
+  done
 
 lemma lemma_st_part_not_eq2:
-     "[| (x::hypreal) \<in> HFinite;
-         isLub \<real> {s. s \<in> \<real> & s < x} t;
-         r \<in> \<real>; 0 < r |]
-      ==> -(x + -t) \<noteq> r"
-apply (auto)
-apply (frule isLubD1a)
-using Reals_add_cancel [of "- x" t] apply simp
-apply (drule_tac x = x in lemma_SReal_lub)
-apply (drule isLub_unique, assumption, auto)
-done
+  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> - (x + -t) \<noteq> r"
+  for x r t :: hypreal
+  apply (auto)
+  apply (frule isLubD1a)
+  using Reals_add_cancel [of "- x" t] apply simp
+  apply (drule_tac x = x in lemma_SReal_lub)
+  apply (drule isLub_unique, assumption, auto)
+  done
 
 lemma lemma_st_part_major:
-     "[| (x::hypreal) \<in> HFinite;
-         isLub \<real> {s. s \<in> \<real> & s < x} t;
-         r \<in> \<real>; 0 < r |]
-      ==> \<bar>x - t\<bar> < r"
-apply (frule lemma_st_part1a)
-apply (frule_tac [4] lemma_st_part2a, auto)
-apply (drule order_le_imp_less_or_eq)+
-apply auto
-using lemma_st_part_not_eq2 apply fastforce
-using lemma_st_part_not_eq1 apply fastforce
-done
+  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> \<bar>x - t\<bar> < r"
+  for x r t :: hypreal
+  apply (frule lemma_st_part1a)
+     apply (frule_tac [4] lemma_st_part2a, auto)
+  apply (drule order_le_imp_less_or_eq)+
+  apply auto
+  using lemma_st_part_not_eq2 apply fastforce
+  using lemma_st_part_not_eq1 apply fastforce
+  done
 
 lemma lemma_st_part_major2:
-     "[| (x::hypreal) \<in> HFinite; isLub \<real> {s. s \<in> \<real> & s < x} t |]
-      ==> \<forall>r \<in> Reals. 0 < r --> \<bar>x - t\<bar> < r"
-by (blast dest!: lemma_st_part_major)
-
+  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> \<forall>r \<in> Reals. 0 < r \<longrightarrow> \<bar>x - t\<bar> < r"
+  for x t :: hypreal
+  by (blast dest!: lemma_st_part_major)
 
 
-text\<open>Existence of real and Standard Part Theorem\<close>
-lemma lemma_st_part_Ex:
-     "(x::hypreal) \<in> HFinite
-       ==> \<exists>t \<in> Reals. \<forall>r \<in> Reals. 0 < r --> \<bar>x - t\<bar> < r"
-apply (frule lemma_st_part_lub, safe)
-apply (frule isLubD1a)
-apply (blast dest: lemma_st_part_major2)
-done
+text\<open>Existence of real and Standard Part Theorem.\<close>
+
+lemma lemma_st_part_Ex: "x \<in> HFinite \<Longrightarrow> \<exists>t\<in>Reals. \<forall>r \<in> Reals. 0 < r \<longrightarrow> \<bar>x - t\<bar> < r"
+  for x :: hypreal
+  apply (frule lemma_st_part_lub, safe)
+  apply (frule isLubD1a)
+  apply (blast dest: lemma_st_part_major2)
+  done
 
-lemma st_part_Ex:
-     "(x::hypreal) \<in> HFinite ==> \<exists>t \<in> Reals. x \<approx> t"
-apply (simp add: approx_def Infinitesimal_def)
-apply (drule lemma_st_part_Ex, auto)
-done
+lemma st_part_Ex: "x \<in> HFinite \<Longrightarrow> \<exists>t\<in>Reals. x \<approx> t"
+  for x :: hypreal
+  apply (simp add: approx_def Infinitesimal_def)
+  apply (drule lemma_st_part_Ex, auto)
+  done
 
-text\<open>There is a unique real infinitely close\<close>
-lemma st_part_Ex1: "x \<in> HFinite ==> \<exists>!t::hypreal. t \<in> \<real> & x \<approx> t"
-apply (drule st_part_Ex, safe)
-apply (drule_tac [2] approx_sym, drule_tac [2] approx_sym, drule_tac [2] approx_sym)
-apply (auto intro!: approx_unique_real)
-done
+text \<open>There is a unique real infinitely close.\<close>
+lemma st_part_Ex1: "x \<in> HFinite \<Longrightarrow> \<exists>!t::hypreal. t \<in> \<real> \<and> x \<approx> t"
+  apply (drule st_part_Ex, safe)
+   apply (drule_tac [2] approx_sym, drule_tac [2] approx_sym, drule_tac [2] approx_sym)
+   apply (auto intro!: approx_unique_real)
+  done
 
-subsection\<open>Finite, Infinite and Infinitesimal\<close>
+
+subsection \<open>Finite, Infinite and Infinitesimal\<close>
 
 lemma HFinite_Int_HInfinite_empty [simp]: "HFinite Int HInfinite = {}"
-apply (simp add: HFinite_def HInfinite_def)
-apply (auto dest: order_less_trans)
-done
+  apply (simp add: HFinite_def HInfinite_def)
+  apply (auto dest: order_less_trans)
+  done
 
 lemma HFinite_not_HInfinite:
-  assumes x: "x \<in> HFinite" shows "x \<notin> HInfinite"
+  assumes x: "x \<in> HFinite"
+  shows "x \<notin> HInfinite"
 proof
   assume x': "x \<in> HInfinite"
   with x have "x \<in> HFinite \<inter> HInfinite" by blast
-  thus False by auto
+  then show False by auto
 qed
 
-lemma not_HFinite_HInfinite: "x\<notin> HFinite ==> x \<in> HInfinite"
-apply (simp add: HInfinite_def HFinite_def, auto)
-apply (drule_tac x = "r + 1" in bspec)
-apply (auto)
-done
+lemma not_HFinite_HInfinite: "x \<notin> HFinite \<Longrightarrow> x \<in> HInfinite"
+  apply (simp add: HInfinite_def HFinite_def, auto)
+  apply (drule_tac x = "r + 1" in bspec)
+   apply (auto)
+  done
 
-lemma HInfinite_HFinite_disj: "x \<in> HInfinite | x \<in> HFinite"
-by (blast intro: not_HFinite_HInfinite)
+lemma HInfinite_HFinite_disj: "x \<in> HInfinite \<or> x \<in> HFinite"
+  by (blast intro: not_HFinite_HInfinite)
 
-lemma HInfinite_HFinite_iff: "(x \<in> HInfinite) = (x \<notin> HFinite)"
-by (blast dest: HFinite_not_HInfinite not_HFinite_HInfinite)
+lemma HInfinite_HFinite_iff: "x \<in> HInfinite \<longleftrightarrow> x \<notin> HFinite"
+  by (blast dest: HFinite_not_HInfinite not_HFinite_HInfinite)
 
-lemma HFinite_HInfinite_iff: "(x \<in> HFinite) = (x \<notin> HInfinite)"
-by (simp add: HInfinite_HFinite_iff)
+lemma HFinite_HInfinite_iff: "x \<in> HFinite \<longleftrightarrow> x \<notin> HInfinite"
+  by (simp add: HInfinite_HFinite_iff)
 
 
 lemma HInfinite_diff_HFinite_Infinitesimal_disj:
-     "x \<notin> Infinitesimal ==> x \<in> HInfinite | x \<in> HFinite - Infinitesimal"
-by (fast intro: not_HFinite_HInfinite)
+  "x \<notin> Infinitesimal \<Longrightarrow> x \<in> HInfinite \<or> x \<in> HFinite - Infinitesimal"
+  by (fast intro: not_HFinite_HInfinite)
 
-lemma HFinite_inverse:
-  fixes x :: "'a::real_normed_div_algebra star"
-  shows "[| x \<in> HFinite; x \<notin> Infinitesimal |] ==> inverse x \<in> HFinite"
-apply (subgoal_tac "x \<noteq> 0")
-apply (cut_tac x = "inverse x" in HInfinite_HFinite_disj)
-apply (auto dest!: HInfinite_inverse_Infinitesimal
-            simp add: nonzero_inverse_inverse_eq)
-done
+lemma HFinite_inverse: "x \<in> HFinite \<Longrightarrow> x \<notin> Infinitesimal \<Longrightarrow> inverse x \<in> HFinite"
+  for x :: "'a::real_normed_div_algebra star"
+  apply (subgoal_tac "x \<noteq> 0")
+   apply (cut_tac x = "inverse x" in HInfinite_HFinite_disj)
+   apply (auto dest!: HInfinite_inverse_Infinitesimal simp: nonzero_inverse_inverse_eq)
+  done
 
-lemma HFinite_inverse2:
-  fixes x :: "'a::real_normed_div_algebra star"
-  shows "x \<in> HFinite - Infinitesimal ==> inverse x \<in> HFinite"
-by (blast intro: HFinite_inverse)
+lemma HFinite_inverse2: "x \<in> HFinite - Infinitesimal \<Longrightarrow> inverse x \<in> HFinite"
+  for x :: "'a::real_normed_div_algebra star"
+  by (blast intro: HFinite_inverse)
 
-(* stronger statement possible in fact *)
-lemma Infinitesimal_inverse_HFinite:
-  fixes x :: "'a::real_normed_div_algebra star"
-  shows "x \<notin> Infinitesimal ==> inverse(x) \<in> HFinite"
-apply (drule HInfinite_diff_HFinite_Infinitesimal_disj)
-apply (blast intro: HFinite_inverse HInfinite_inverse_Infinitesimal Infinitesimal_subset_HFinite [THEN subsetD])
-done
+text \<open>Stronger statement possible in fact.\<close>
+lemma Infinitesimal_inverse_HFinite: "x \<notin> Infinitesimal \<Longrightarrow> inverse x \<in> HFinite"
+  for x :: "'a::real_normed_div_algebra star"
+  apply (drule HInfinite_diff_HFinite_Infinitesimal_disj)
+  apply (blast intro: HFinite_inverse HInfinite_inverse_Infinitesimal Infinitesimal_subset_HFinite [THEN subsetD])
+  done
 
 lemma HFinite_not_Infinitesimal_inverse:
-  fixes x :: "'a::real_normed_div_algebra star"
-  shows "x \<in> HFinite - Infinitesimal ==> inverse x \<in> HFinite - Infinitesimal"
-apply (auto intro: Infinitesimal_inverse_HFinite)
-apply (drule Infinitesimal_HFinite_mult2, assumption)
-apply (simp add: not_Infinitesimal_not_zero)
-done
+  "x \<in> HFinite - Infinitesimal \<Longrightarrow> inverse x \<in> HFinite - Infinitesimal"
+  for x :: "'a::real_normed_div_algebra star"
+  apply (auto intro: Infinitesimal_inverse_HFinite)
+  apply (drule Infinitesimal_HFinite_mult2, assumption)
+  apply (simp add: not_Infinitesimal_not_zero)
+  done
 
-lemma approx_inverse:
-  fixes x y :: "'a::real_normed_div_algebra star"
-  shows
-     "[| x \<approx> y; y \<in>  HFinite - Infinitesimal |]
-      ==> inverse x \<approx> inverse y"
-apply (frule HFinite_diff_Infinitesimal_approx, assumption)
-apply (frule not_Infinitesimal_not_zero2)
-apply (frule_tac x = x in not_Infinitesimal_not_zero2)
-apply (drule HFinite_inverse2)+
-apply (drule approx_mult2, assumption, auto)
-apply (drule_tac c = "inverse x" in approx_mult1, assumption)
-apply (auto intro: approx_sym simp add: mult.assoc)
-done
+lemma approx_inverse: "x \<approx> y \<Longrightarrow> y \<in> HFinite - Infinitesimal \<Longrightarrow> inverse x \<approx> inverse y"
+  for x y :: "'a::real_normed_div_algebra star"
+  apply (frule HFinite_diff_Infinitesimal_approx, assumption)
+  apply (frule not_Infinitesimal_not_zero2)
+  apply (frule_tac x = x in not_Infinitesimal_not_zero2)
+  apply (drule HFinite_inverse2)+
+  apply (drule approx_mult2, assumption, auto)
+  apply (drule_tac c = "inverse x" in approx_mult1, assumption)
+  apply (auto intro: approx_sym simp add: mult.assoc)
+  done
 
 (*Used for NSLIM_inverse, NSLIMSEQ_inverse*)
 lemmas star_of_approx_inverse = star_of_HFinite_diff_Infinitesimal [THEN [2] approx_inverse]
 lemmas hypreal_of_real_approx_inverse =  hypreal_of_real_HFinite_diff_Infinitesimal [THEN [2] approx_inverse]
 
 lemma inverse_add_Infinitesimal_approx:
-  fixes x h :: "'a::real_normed_div_algebra star"
-  shows
-     "[| x \<in> HFinite - Infinitesimal;
-         h \<in> Infinitesimal |] ==> inverse(x + h) \<approx> inverse x"
-apply (auto intro: approx_inverse approx_sym Infinitesimal_add_approx_self)
-done
+  "x \<in> HFinite - Infinitesimal \<Longrightarrow> h \<in> Infinitesimal \<Longrightarrow> inverse (x + h) \<approx> inverse x"
+  for x h :: "'a::real_normed_div_algebra star"
+  by (auto intro: approx_inverse approx_sym Infinitesimal_add_approx_self)
 
 lemma inverse_add_Infinitesimal_approx2:
-  fixes x h :: "'a::real_normed_div_algebra star"
-  shows
-     "[| x \<in> HFinite - Infinitesimal;
-         h \<in> Infinitesimal |] ==> inverse(h + x) \<approx> inverse x"
-apply (rule add.commute [THEN subst])
-apply (blast intro: inverse_add_Infinitesimal_approx)
-done
+  "x \<in> HFinite - Infinitesimal \<Longrightarrow> h \<in> Infinitesimal \<Longrightarrow> inverse (h + x) \<approx> inverse x"
+  for x h :: "'a::real_normed_div_algebra star"
+  apply (rule add.commute [THEN subst])
+  apply (blast intro: inverse_add_Infinitesimal_approx)
+  done
 
 lemma inverse_add_Infinitesimal_approx_Infinitesimal:
-  fixes x h :: "'a::real_normed_div_algebra star"
-  shows
-     "[| x \<in> HFinite - Infinitesimal;
-         h \<in> Infinitesimal |] ==> inverse(x + h) - inverse x \<approx> h"
-apply (rule approx_trans2)
-apply (auto intro: inverse_add_Infinitesimal_approx
-            simp add: mem_infmal_iff approx_minus_iff [symmetric])
-done
+  "x \<in> HFinite - Infinitesimal \<Longrightarrow> h \<in> Infinitesimal \<Longrightarrow> inverse (x + h) - inverse x \<approx> h"
+  for x h :: "'a::real_normed_div_algebra star"
+  apply (rule approx_trans2)
+   apply (auto intro: inverse_add_Infinitesimal_approx
+      simp add: mem_infmal_iff approx_minus_iff [symmetric])
+  done
 
-lemma Infinitesimal_square_iff:
-  fixes x :: "'a::real_normed_div_algebra star"
-  shows "(x \<in> Infinitesimal) = (x*x \<in> Infinitesimal)"
-apply (auto intro: Infinitesimal_mult)
-apply (rule ccontr, frule Infinitesimal_inverse_HFinite)
-apply (frule not_Infinitesimal_not_zero)
-apply (auto dest: Infinitesimal_HFinite_mult simp add: mult.assoc)
-done
+lemma Infinitesimal_square_iff: "x \<in> Infinitesimal \<longleftrightarrow> x * x \<in> Infinitesimal"
+  for x :: "'a::real_normed_div_algebra star"
+  apply (auto intro: Infinitesimal_mult)
+  apply (rule ccontr, frule Infinitesimal_inverse_HFinite)
+  apply (frule not_Infinitesimal_not_zero)
+  apply (auto dest: Infinitesimal_HFinite_mult simp add: mult.assoc)
+  done
 declare Infinitesimal_square_iff [symmetric, simp]
 
-lemma HFinite_square_iff [simp]:
-  fixes x :: "'a::real_normed_div_algebra star"
-  shows "(x*x \<in> HFinite) = (x \<in> HFinite)"
-apply (auto intro: HFinite_mult)
-apply (auto dest: HInfinite_mult simp add: HFinite_HInfinite_iff)
-done
+lemma HFinite_square_iff [simp]: "x * x \<in> HFinite \<longleftrightarrow> x \<in> HFinite"
+  for x :: "'a::real_normed_div_algebra star"
+  apply (auto intro: HFinite_mult)
+  apply (auto dest: HInfinite_mult simp add: HFinite_HInfinite_iff)
+  done
 
-lemma HInfinite_square_iff [simp]:
-  fixes x :: "'a::real_normed_div_algebra star"
-  shows "(x*x \<in> HInfinite) = (x \<in> HInfinite)"
-by (auto simp add: HInfinite_HFinite_iff)
+lemma HInfinite_square_iff [simp]: "x * x \<in> HInfinite \<longleftrightarrow> x \<in> HInfinite"
+  for x :: "'a::real_normed_div_algebra star"
+  by (auto simp add: HInfinite_HFinite_iff)
 
-lemma approx_HFinite_mult_cancel:
-  fixes a w z :: "'a::real_normed_div_algebra star"
-  shows "[| a: HFinite-Infinitesimal; a* w \<approx> a*z |] ==> w \<approx> z"
-apply safe
-apply (frule HFinite_inverse, assumption)
-apply (drule not_Infinitesimal_not_zero)
-apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric])
-done
+lemma approx_HFinite_mult_cancel: "a \<in> HFinite - Infinitesimal \<Longrightarrow> a * w \<approx> a * z \<Longrightarrow> w \<approx> z"
+  for a w z :: "'a::real_normed_div_algebra star"
+  apply safe
+  apply (frule HFinite_inverse, assumption)
+  apply (drule not_Infinitesimal_not_zero)
+  apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric])
+  done
 
-lemma approx_HFinite_mult_cancel_iff1:
-  fixes a w z :: "'a::real_normed_div_algebra star"
-  shows "a: HFinite-Infinitesimal ==> (a * w \<approx> a * z) = (w \<approx> z)"
-by (auto intro: approx_mult2 approx_HFinite_mult_cancel)
+lemma approx_HFinite_mult_cancel_iff1: "a \<in> HFinite - Infinitesimal \<Longrightarrow> a * w \<approx> a * z \<longleftrightarrow> w \<approx> z"
+  for a w z :: "'a::real_normed_div_algebra star"
+  by (auto intro: approx_mult2 approx_HFinite_mult_cancel)
 
-lemma HInfinite_HFinite_add_cancel:
-     "[| x + y \<in> HInfinite; y \<in> HFinite |] ==> x \<in> HInfinite"
-apply (rule ccontr)
-apply (drule HFinite_HInfinite_iff [THEN iffD2])
-apply (auto dest: HFinite_add simp add: HInfinite_HFinite_iff)
-done
+lemma HInfinite_HFinite_add_cancel: "x + y \<in> HInfinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x \<in> HInfinite"
+  apply (rule ccontr)
+  apply (drule HFinite_HInfinite_iff [THEN iffD2])
+  apply (auto dest: HFinite_add simp add: HInfinite_HFinite_iff)
+  done
 
-lemma HInfinite_HFinite_add:
-     "[| x \<in> HInfinite; y \<in> HFinite |] ==> x + y \<in> HInfinite"
-apply (rule_tac y = "-y" in HInfinite_HFinite_add_cancel)
-apply (auto simp add: add.assoc HFinite_minus_iff)
-done
+lemma HInfinite_HFinite_add: "x \<in> HInfinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x + y \<in> HInfinite"
+  apply (rule_tac y = "-y" in HInfinite_HFinite_add_cancel)
+   apply (auto simp add: add.assoc HFinite_minus_iff)
+  done
 
-lemma HInfinite_ge_HInfinite:
-     "[| (x::hypreal) \<in> HInfinite; x \<le> y; 0 \<le> x |] ==> y \<in> HInfinite"
-by (auto intro: HFinite_bounded simp add: HInfinite_HFinite_iff)
+lemma HInfinite_ge_HInfinite: "x \<in> HInfinite \<Longrightarrow> x \<le> y \<Longrightarrow> 0 \<le> x \<Longrightarrow> y \<in> HInfinite"
+  for x y :: hypreal
+  by (auto intro: HFinite_bounded simp add: HInfinite_HFinite_iff)
 
-lemma Infinitesimal_inverse_HInfinite:
-  fixes x :: "'a::real_normed_div_algebra star"
-  shows "[| x \<in> Infinitesimal; x \<noteq> 0 |] ==> inverse x \<in> HInfinite"
-apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])
-apply (auto dest: Infinitesimal_HFinite_mult2)
-done
+lemma Infinitesimal_inverse_HInfinite: "x \<in> Infinitesimal \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> inverse x \<in> HInfinite"
+  for x :: "'a::real_normed_div_algebra star"
+  apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])
+  apply (auto dest: Infinitesimal_HFinite_mult2)
+  done
 
 lemma HInfinite_HFinite_not_Infinitesimal_mult:
-  fixes x y :: "'a::real_normed_div_algebra star"
-  shows "[| x \<in> HInfinite; y \<in> HFinite - Infinitesimal |]
-      ==> x * y \<in> HInfinite"
-apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])
-apply (frule HFinite_Infinitesimal_not_zero)
-apply (drule HFinite_not_Infinitesimal_inverse)
-apply (safe, drule HFinite_mult)
-apply (auto simp add: mult.assoc HFinite_HInfinite_iff)
-done
+  "x \<in> HInfinite \<Longrightarrow> y \<in> HFinite - Infinitesimal \<Longrightarrow> x * y \<in> HInfinite"
+  for x y :: "'a::real_normed_div_algebra star"
+  apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])
+  apply (frule HFinite_Infinitesimal_not_zero)
+  apply (drule HFinite_not_Infinitesimal_inverse)
+  apply (safe, drule HFinite_mult)
+   apply (auto simp add: mult.assoc HFinite_HInfinite_iff)
+  done
 
 lemma HInfinite_HFinite_not_Infinitesimal_mult2:
-  fixes x y :: "'a::real_normed_div_algebra star"
-  shows "[| x \<in> HInfinite; y \<in> HFinite - Infinitesimal |]
-      ==> y * x \<in> HInfinite"
-apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])
-apply (frule HFinite_Infinitesimal_not_zero)
-apply (drule HFinite_not_Infinitesimal_inverse)
-apply (safe, drule_tac x="inverse y" in HFinite_mult)
-apply assumption
-apply (auto simp add: mult.assoc [symmetric] HFinite_HInfinite_iff)
-done
+  "x \<in> HInfinite \<Longrightarrow> y \<in> HFinite - Infinitesimal \<Longrightarrow> y * x \<in> HInfinite"
+  for x y :: "'a::real_normed_div_algebra star"
+  apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])
+  apply (frule HFinite_Infinitesimal_not_zero)
+  apply (drule HFinite_not_Infinitesimal_inverse)
+  apply (safe, drule_tac x="inverse y" in HFinite_mult)
+   apply assumption
+  apply (auto simp add: mult.assoc [symmetric] HFinite_HInfinite_iff)
+  done
 
-lemma HInfinite_gt_SReal:
-  "[| (x::hypreal) \<in> HInfinite; 0 < x; y \<in> \<real> |] ==> y < x"
-by (auto dest!: bspec simp add: HInfinite_def abs_if order_less_imp_le)
+lemma HInfinite_gt_SReal: "x \<in> HInfinite \<Longrightarrow> 0 < x \<Longrightarrow> y \<in> \<real> \<Longrightarrow> y < x"
+  for x y :: hypreal
+  by (auto dest!: bspec simp add: HInfinite_def abs_if order_less_imp_le)
 
-lemma HInfinite_gt_zero_gt_one:
-  "[| (x::hypreal) \<in> HInfinite; 0 < x |] ==> 1 < x"
-by (auto intro: HInfinite_gt_SReal)
+lemma HInfinite_gt_zero_gt_one: "x \<in> HInfinite \<Longrightarrow> 0 < x \<Longrightarrow> 1 < x"
+  for x :: hypreal
+  by (auto intro: HInfinite_gt_SReal)
 
 
 lemma not_HInfinite_one [simp]: "1 \<notin> HInfinite"
-apply (simp (no_asm) add: HInfinite_HFinite_iff)
-done
+  by (simp add: HInfinite_HFinite_iff)
 
-lemma approx_hrabs_disj: "\<bar>x::hypreal\<bar> \<approx> x \<or> \<bar>x\<bar> \<approx> -x"
-by (cut_tac x = x in hrabs_disj, auto)
+lemma approx_hrabs_disj: "\<bar>x\<bar> \<approx> x \<or> \<bar>x\<bar> \<approx> -x"
+  for x :: hypreal
+  using hrabs_disj [of x] by auto
 
 
-subsection\<open>Theorems about Monads\<close>
-
-lemma monad_hrabs_Un_subset: "monad \<bar>x\<bar> \<le> monad(x::hypreal) Un monad(-x)"
-by (rule_tac x1 = x in hrabs_disj [THEN disjE], auto)
+subsection \<open>Theorems about Monads\<close>
 
-lemma Infinitesimal_monad_eq: "e \<in> Infinitesimal ==> monad (x+e) = monad x"
-by (fast intro!: Infinitesimal_add_approx_self [THEN approx_sym] approx_monad_iff [THEN iffD1])
+lemma monad_hrabs_Un_subset: "monad \<bar>x\<bar> \<le> monad x \<union> monad (- x)"
+  for x :: hypreal
+  by (rule hrabs_disj [of x, THEN disjE]) auto
 
-lemma mem_monad_iff: "(u \<in> monad x) = (-u \<in> monad (-x))"
-by (simp add: monad_def)
+lemma Infinitesimal_monad_eq: "e \<in> Infinitesimal \<Longrightarrow> monad (x + e) = monad x"
+  by (fast intro!: Infinitesimal_add_approx_self [THEN approx_sym] approx_monad_iff [THEN iffD1])
 
-lemma Infinitesimal_monad_zero_iff: "(x \<in> Infinitesimal) = (x \<in> monad 0)"
-by (auto intro: approx_sym simp add: monad_def mem_infmal_iff)
+lemma mem_monad_iff: "u \<in> monad x \<longleftrightarrow> - u \<in> monad (- x)"
+  by (simp add: monad_def)
+
+lemma Infinitesimal_monad_zero_iff: "x \<in> Infinitesimal \<longleftrightarrow> x \<in> monad 0"
+  by (auto intro: approx_sym simp add: monad_def mem_infmal_iff)
 
-lemma monad_zero_minus_iff: "(x \<in> monad 0) = (-x \<in> monad 0)"
-apply (simp (no_asm) add: Infinitesimal_monad_zero_iff [symmetric])
-done
+lemma monad_zero_minus_iff: "x \<in> monad 0 \<longleftrightarrow> - x \<in> monad 0"
+  by (simp add: Infinitesimal_monad_zero_iff [symmetric])
 
-lemma monad_zero_hrabs_iff: "((x::hypreal) \<in> monad 0) = (\<bar>x\<bar> \<in> monad 0)"
-apply (rule_tac x1 = x in hrabs_disj [THEN disjE])
-apply (auto simp add: monad_zero_minus_iff [symmetric])
-done
+lemma monad_zero_hrabs_iff: "x \<in> monad 0 \<longleftrightarrow> \<bar>x\<bar> \<in> monad 0"
+  for x :: hypreal
+  by (rule hrabs_disj [of x, THEN disjE]) (auto simp: monad_zero_minus_iff [symmetric])
 
 lemma mem_monad_self [simp]: "x \<in> monad x"
-by (simp add: monad_def)
+  by (simp add: monad_def)
 
 
-subsection\<open>Proof that @{term "x \<approx> y"} implies @{term"\<bar>x\<bar> \<approx> \<bar>y\<bar>"}\<close>
+subsection \<open>Proof that @{term "x \<approx> y"} implies @{term"\<bar>x\<bar> \<approx> \<bar>y\<bar>"}\<close>
 
-lemma approx_subset_monad: "x \<approx> y ==> {x,y} \<le> monad x"
-apply (simp (no_asm))
-apply (simp add: approx_monad_iff)
-done
+lemma approx_subset_monad: "x \<approx> y \<Longrightarrow> {x, y} \<le> monad x"
+  by (simp (no_asm)) (simp add: approx_monad_iff)
 
-lemma approx_subset_monad2: "x \<approx> y ==> {x,y} \<le> monad y"
-apply (drule approx_sym)
-apply (fast dest: approx_subset_monad)
-done
+lemma approx_subset_monad2: "x \<approx> y \<Longrightarrow> {x, y} \<le> monad y"
+  apply (drule approx_sym)
+  apply (fast dest: approx_subset_monad)
+  done
 
-lemma mem_monad_approx: "u \<in> monad x ==> x \<approx> u"
-by (simp add: monad_def)
+lemma mem_monad_approx: "u \<in> monad x \<Longrightarrow> x \<approx> u"
+  by (simp add: monad_def)
+
+lemma approx_mem_monad: "x \<approx> u \<Longrightarrow> u \<in> monad x"
+  by (simp add: monad_def)
 
-lemma approx_mem_monad: "x \<approx> u ==> u \<in> monad x"
-by (simp add: monad_def)
-
-lemma approx_mem_monad2: "x \<approx> u ==> x \<in> monad u"
-apply (simp add: monad_def)
-apply (blast intro!: approx_sym)
-done
+lemma approx_mem_monad2: "x \<approx> u \<Longrightarrow> x \<in> monad u"
+  apply (simp add: monad_def)
+  apply (blast intro!: approx_sym)
+  done
 
-lemma approx_mem_monad_zero: "[| x \<approx> y;x \<in> monad 0 |] ==> y \<in> monad 0"
-apply (drule mem_monad_approx)
-apply (fast intro: approx_mem_monad approx_trans)
-done
+lemma approx_mem_monad_zero: "x \<approx> y \<Longrightarrow> x \<in> monad 0 \<Longrightarrow> y \<in> monad 0"
+  apply (drule mem_monad_approx)
+  apply (fast intro: approx_mem_monad approx_trans)
+  done
 
-lemma Infinitesimal_approx_hrabs:
-     "[| x \<approx> y; (x::hypreal) \<in> Infinitesimal |] ==> \<bar>x\<bar> \<approx> \<bar>y\<bar>"
-apply (drule Infinitesimal_monad_zero_iff [THEN iffD1])
-apply (blast intro: approx_mem_monad_zero monad_zero_hrabs_iff [THEN iffD1] mem_monad_approx approx_trans3)
-done
+lemma Infinitesimal_approx_hrabs: "x \<approx> y \<Longrightarrow> x \<in> Infinitesimal \<Longrightarrow> \<bar>x\<bar> \<approx> \<bar>y\<bar>"
+  for x y :: hypreal
+  apply (drule Infinitesimal_monad_zero_iff [THEN iffD1])
+  apply (blast intro: approx_mem_monad_zero monad_zero_hrabs_iff [THEN iffD1]
+      mem_monad_approx approx_trans3)
+  done
 
-lemma less_Infinitesimal_less:
-     "[| 0 < x;  (x::hypreal) \<notin>Infinitesimal;  e :Infinitesimal |] ==> e < x"
-apply (rule ccontr)
-apply (auto intro: Infinitesimal_zero [THEN [2] Infinitesimal_interval]
-            dest!: order_le_imp_less_or_eq simp add: linorder_not_less)
-done
+lemma less_Infinitesimal_less: "0 < x \<Longrightarrow> x \<notin> Infinitesimal \<Longrightarrow> e \<in> Infinitesimal \<Longrightarrow> e < x"
+  for x :: hypreal
+  apply (rule ccontr)
+  apply (auto intro: Infinitesimal_zero [THEN [2] Infinitesimal_interval]
+      dest!: order_le_imp_less_or_eq simp add: linorder_not_less)
+  done
 
-lemma Ball_mem_monad_gt_zero:
-     "[| 0 < (x::hypreal);  x \<notin> Infinitesimal; u \<in> monad x |] ==> 0 < u"
-apply (drule mem_monad_approx [THEN approx_sym])
-apply (erule bex_Infinitesimal_iff2 [THEN iffD2, THEN bexE])
-apply (drule_tac e = "-xa" in less_Infinitesimal_less, auto)
-done
+lemma Ball_mem_monad_gt_zero: "0 < x \<Longrightarrow> x \<notin> Infinitesimal \<Longrightarrow> u \<in> monad x \<Longrightarrow> 0 < u"
+  for u x :: hypreal
+  apply (drule mem_monad_approx [THEN approx_sym])
+  apply (erule bex_Infinitesimal_iff2 [THEN iffD2, THEN bexE])
+  apply (drule_tac e = "-xa" in less_Infinitesimal_less, auto)
+  done
 
-lemma Ball_mem_monad_less_zero:
-     "[| (x::hypreal) < 0; x \<notin> Infinitesimal; u \<in> monad x |] ==> u < 0"
-apply (drule mem_monad_approx [THEN approx_sym])
-apply (erule bex_Infinitesimal_iff [THEN iffD2, THEN bexE])
-apply (cut_tac x = "-x" and e = xa in less_Infinitesimal_less, auto)
-done
+lemma Ball_mem_monad_less_zero: "x < 0 \<Longrightarrow> x \<notin> Infinitesimal \<Longrightarrow> u \<in> monad x \<Longrightarrow> u < 0"
+  for u x :: hypreal
+  apply (drule mem_monad_approx [THEN approx_sym])
+  apply (erule bex_Infinitesimal_iff [THEN iffD2, THEN bexE])
+  apply (cut_tac x = "-x" and e = xa in less_Infinitesimal_less, auto)
+  done
 
-lemma lemma_approx_gt_zero:
-     "[|0 < (x::hypreal); x \<notin> Infinitesimal; x \<approx> y|] ==> 0 < y"
-by (blast dest: Ball_mem_monad_gt_zero approx_subset_monad)
+lemma lemma_approx_gt_zero: "0 < x \<Longrightarrow> x \<notin> Infinitesimal \<Longrightarrow> x \<approx> y \<Longrightarrow> 0 < y"
+  for x y :: hypreal
+  by (blast dest: Ball_mem_monad_gt_zero approx_subset_monad)
 
-lemma lemma_approx_less_zero:
-     "[|(x::hypreal) < 0; x \<notin> Infinitesimal; x \<approx> y|] ==> y < 0"
-by (blast dest: Ball_mem_monad_less_zero approx_subset_monad)
+lemma lemma_approx_less_zero: "x < 0 \<Longrightarrow> x \<notin> Infinitesimal \<Longrightarrow> x \<approx> y \<Longrightarrow> y < 0"
+  for x y :: hypreal
+  by (blast dest: Ball_mem_monad_less_zero approx_subset_monad)
 
-theorem approx_hrabs: "(x::hypreal) \<approx> y ==> \<bar>x\<bar> \<approx> \<bar>y\<bar>"
-by (drule approx_hnorm, simp)
+lemma approx_hrabs: "x \<approx> y \<Longrightarrow> \<bar>x\<bar> \<approx> \<bar>y\<bar>"
+  for x y :: hypreal
+  by (drule approx_hnorm) simp
 
-lemma approx_hrabs_zero_cancel: "\<bar>x::hypreal\<bar> \<approx> 0 ==> x \<approx> 0"
-apply (cut_tac x = x in hrabs_disj)
-apply (auto dest: approx_minus)
-done
+lemma approx_hrabs_zero_cancel: "\<bar>x\<bar> \<approx> 0 \<Longrightarrow> x \<approx> 0"
+  for x :: hypreal
+  using hrabs_disj [of x] by (auto dest: approx_minus)
 
-lemma approx_hrabs_add_Infinitesimal:
-  "(e::hypreal) \<in> Infinitesimal ==> \<bar>x\<bar> \<approx> \<bar>x + e\<bar>"
-by (fast intro: approx_hrabs Infinitesimal_add_approx_self)
+lemma approx_hrabs_add_Infinitesimal: "e \<in> Infinitesimal \<Longrightarrow> \<bar>x\<bar> \<approx> \<bar>x + e\<bar>"
+  for e x :: hypreal
+  by (fast intro: approx_hrabs Infinitesimal_add_approx_self)
 
-lemma approx_hrabs_add_minus_Infinitesimal:
-     "(e::hypreal) \<in> Infinitesimal ==> \<bar>x\<bar> \<approx> \<bar>x + -e\<bar>"
-by (fast intro: approx_hrabs Infinitesimal_add_minus_approx_self)
+lemma approx_hrabs_add_minus_Infinitesimal: "e \<in> Infinitesimal ==> \<bar>x\<bar> \<approx> \<bar>x + -e\<bar>"
+  for e x :: hypreal
+  by (fast intro: approx_hrabs Infinitesimal_add_minus_approx_self)
 
 lemma hrabs_add_Infinitesimal_cancel:
-     "[| (e::hypreal) \<in> Infinitesimal; e' \<in> Infinitesimal;
-         \<bar>x + e\<bar> = \<bar>y + e'\<bar>|] ==> \<bar>x\<bar> \<approx> \<bar>y\<bar>"
-apply (drule_tac x = x in approx_hrabs_add_Infinitesimal)
-apply (drule_tac x = y in approx_hrabs_add_Infinitesimal)
-apply (auto intro: approx_trans2)
-done
+  "e \<in> Infinitesimal \<Longrightarrow> e' \<in> Infinitesimal \<Longrightarrow> \<bar>x + e\<bar> = \<bar>y + e'\<bar> \<Longrightarrow> \<bar>x\<bar> \<approx> \<bar>y\<bar>"
+  for e e' x y :: hypreal
+  apply (drule_tac x = x in approx_hrabs_add_Infinitesimal)
+  apply (drule_tac x = y in approx_hrabs_add_Infinitesimal)
+  apply (auto intro: approx_trans2)
+  done
 
 lemma hrabs_add_minus_Infinitesimal_cancel:
-     "[| (e::hypreal) \<in> Infinitesimal; e' \<in> Infinitesimal;
-         \<bar>x + -e\<bar> = \<bar>y + -e'\<bar>|] ==> \<bar>x\<bar> \<approx> \<bar>y\<bar>"
-apply (drule_tac x = x in approx_hrabs_add_minus_Infinitesimal)
-apply (drule_tac x = y in approx_hrabs_add_minus_Infinitesimal)
-apply (auto intro: approx_trans2)
-done
+  "e \<in> Infinitesimal \<Longrightarrow> e' \<in> Infinitesimal \<Longrightarrow> \<bar>x + -e\<bar> = \<bar>y + -e'\<bar> \<Longrightarrow> \<bar>x\<bar> \<approx> \<bar>y\<bar>"
+  for e e' x y :: hypreal
+  apply (drule_tac x = x in approx_hrabs_add_minus_Infinitesimal)
+  apply (drule_tac x = y in approx_hrabs_add_minus_Infinitesimal)
+  apply (auto intro: approx_trans2)
+  done
+
 
 subsection \<open>More @{term HFinite} and @{term Infinitesimal} Theorems\<close>
 
-(* interesting slightly counterintuitive theorem: necessary
-   for proving that an open interval is an NS open set
-*)
+text \<open>
+  Interesting slightly counterintuitive theorem: necessary
+  for proving that an open interval is an NS open set.
+\<close>
 lemma Infinitesimal_add_hypreal_of_real_less:
-     "[| x < y;  u \<in> Infinitesimal |]
-      ==> hypreal_of_real x + u < hypreal_of_real y"
-apply (simp add: Infinitesimal_def)
-apply (drule_tac x = "hypreal_of_real y + -hypreal_of_real x" in bspec, simp)
-apply (simp add: abs_less_iff)
-done
+  "x < y \<Longrightarrow> u \<in> Infinitesimal \<Longrightarrow> hypreal_of_real x + u < hypreal_of_real y"
+  apply (simp add: Infinitesimal_def)
+  apply (drule_tac x = "hypreal_of_real y + -hypreal_of_real x" in bspec, simp)
+  apply (simp add: abs_less_iff)
+  done
 
 lemma Infinitesimal_add_hrabs_hypreal_of_real_less:
-     "[| x \<in> Infinitesimal; \<bar>hypreal_of_real r\<bar> < hypreal_of_real y |]
-      ==> \<bar>hypreal_of_real r + x\<bar> < hypreal_of_real y"
-apply (drule_tac x = "hypreal_of_real r" in approx_hrabs_add_Infinitesimal)
-apply (drule approx_sym [THEN bex_Infinitesimal_iff2 [THEN iffD2]])
-apply (auto intro!: Infinitesimal_add_hypreal_of_real_less
-            simp del: star_of_abs
-            simp add: star_of_abs [symmetric])
-done
+  "x \<in> Infinitesimal \<Longrightarrow> \<bar>hypreal_of_real r\<bar> < hypreal_of_real y \<Longrightarrow>
+    \<bar>hypreal_of_real r + x\<bar> < hypreal_of_real y"
+  apply (drule_tac x = "hypreal_of_real r" in approx_hrabs_add_Infinitesimal)
+  apply (drule approx_sym [THEN bex_Infinitesimal_iff2 [THEN iffD2]])
+  apply (auto intro!: Infinitesimal_add_hypreal_of_real_less
+      simp del: star_of_abs simp add: star_of_abs [symmetric])
+  done
 
 lemma Infinitesimal_add_hrabs_hypreal_of_real_less2:
-     "[| x \<in> Infinitesimal;  \<bar>hypreal_of_real r\<bar> < hypreal_of_real y |]
-      ==> \<bar>x + hypreal_of_real r\<bar> < hypreal_of_real y"
-apply (rule add.commute [THEN subst])
-apply (erule Infinitesimal_add_hrabs_hypreal_of_real_less, assumption)
-done
+  "x \<in> Infinitesimal \<Longrightarrow> \<bar>hypreal_of_real r\<bar> < hypreal_of_real y \<Longrightarrow>
+    \<bar>x + hypreal_of_real r\<bar> < hypreal_of_real y"
+  apply (rule add.commute [THEN subst])
+  apply (erule Infinitesimal_add_hrabs_hypreal_of_real_less, assumption)
+  done
 
 lemma hypreal_of_real_le_add_Infininitesimal_cancel:
-     "[| u \<in> Infinitesimal; v \<in> Infinitesimal;
-         hypreal_of_real x + u \<le> hypreal_of_real y + v |]
-      ==> hypreal_of_real x \<le> hypreal_of_real y"
-apply (simp add: linorder_not_less [symmetric], auto)
-apply (drule_tac u = "v-u" in Infinitesimal_add_hypreal_of_real_less)
-apply (auto simp add: Infinitesimal_diff)
-done
+  "u \<in> Infinitesimal \<Longrightarrow> v \<in> Infinitesimal \<Longrightarrow>
+    hypreal_of_real x + u \<le> hypreal_of_real y + v \<Longrightarrow>
+    hypreal_of_real x \<le> hypreal_of_real y"
+  apply (simp add: linorder_not_less [symmetric], auto)
+  apply (drule_tac u = "v-u" in Infinitesimal_add_hypreal_of_real_less)
+   apply (auto simp add: Infinitesimal_diff)
+  done
 
 lemma hypreal_of_real_le_add_Infininitesimal_cancel2:
-     "[| u \<in> Infinitesimal; v \<in> Infinitesimal;
-         hypreal_of_real x + u \<le> hypreal_of_real y + v |]
-      ==> x \<le> y"
-by (blast intro: star_of_le [THEN iffD1]
-          intro!: hypreal_of_real_le_add_Infininitesimal_cancel)
+  "u \<in> Infinitesimal \<Longrightarrow> v \<in> Infinitesimal \<Longrightarrow>
+    hypreal_of_real x + u \<le> hypreal_of_real y + v \<Longrightarrow> x \<le> y"
+  by (blast intro: star_of_le [THEN iffD1] intro!: hypreal_of_real_le_add_Infininitesimal_cancel)
 
 lemma hypreal_of_real_less_Infinitesimal_le_zero:
-    "[| hypreal_of_real x < e; e \<in> Infinitesimal |] ==> hypreal_of_real x \<le> 0"
-apply (rule linorder_not_less [THEN iffD1], safe)
-apply (drule Infinitesimal_interval)
-apply (drule_tac [4] SReal_hypreal_of_real [THEN SReal_Infinitesimal_zero], auto)
-done
+  "hypreal_of_real x < e \<Longrightarrow> e \<in> Infinitesimal \<Longrightarrow> hypreal_of_real x \<le> 0"
+  apply (rule linorder_not_less [THEN iffD1], safe)
+  apply (drule Infinitesimal_interval)
+     apply (drule_tac [4] SReal_hypreal_of_real [THEN SReal_Infinitesimal_zero], auto)
+  done
 
 (*used once, in Lim/NSDERIV_inverse*)
-lemma Infinitesimal_add_not_zero:
-     "[| h \<in> Infinitesimal; x \<noteq> 0 |] ==> star_of x + h \<noteq> 0"
-apply auto
-apply (subgoal_tac "h = - star_of x", auto intro: minus_unique [symmetric])
-done
+lemma Infinitesimal_add_not_zero: "h \<in> Infinitesimal \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> star_of x + h \<noteq> 0"
+  apply auto
+  apply (subgoal_tac "h = - star_of x")
+   apply (auto intro: minus_unique [symmetric])
+  done
 
-lemma Infinitesimal_square_cancel [simp]:
-     "(x::hypreal)*x + y*y \<in> Infinitesimal ==> x*x \<in> Infinitesimal"
-apply (rule Infinitesimal_interval2)
-apply (rule_tac [3] zero_le_square, assumption)
-apply (auto)
-done
+lemma Infinitesimal_square_cancel [simp]: "x * x + y * y \<in> Infinitesimal \<Longrightarrow> x * x \<in> Infinitesimal"
+  for x y :: hypreal
+  apply (rule Infinitesimal_interval2)
+     apply (rule_tac [3] zero_le_square, assumption)
+   apply auto
+  done
 
-lemma HFinite_square_cancel [simp]:
-  "(x::hypreal)*x + y*y \<in> HFinite ==> x*x \<in> HFinite"
-apply (rule HFinite_bounded, assumption)
-apply (auto)
-done
+lemma HFinite_square_cancel [simp]: "x * x + y * y \<in> HFinite \<Longrightarrow> x * x \<in> HFinite"
+  for x y :: hypreal
+  apply (rule HFinite_bounded, assumption)
+   apply auto
+  done
 
-lemma Infinitesimal_square_cancel2 [simp]:
-     "(x::hypreal)*x + y*y \<in> Infinitesimal ==> y*y \<in> Infinitesimal"
-apply (rule Infinitesimal_square_cancel)
-apply (rule add.commute [THEN subst])
-apply (simp (no_asm))
-done
+lemma Infinitesimal_square_cancel2 [simp]: "x * x + y * y \<in> Infinitesimal \<Longrightarrow> y * y \<in> Infinitesimal"
+  for x y :: hypreal
+  apply (rule Infinitesimal_square_cancel)
+  apply (rule add.commute [THEN subst])
+  apply simp
+  done
 
-lemma HFinite_square_cancel2 [simp]:
-  "(x::hypreal)*x + y*y \<in> HFinite ==> y*y \<in> HFinite"
-apply (rule HFinite_square_cancel)
-apply (rule add.commute [THEN subst])
-apply (simp (no_asm))
-done
+lemma HFinite_square_cancel2 [simp]: "x * x + y * y \<in> HFinite \<Longrightarrow> y * y \<in> HFinite"
+  for x y :: hypreal
+  apply (rule HFinite_square_cancel)
+  apply (rule add.commute [THEN subst])
+  apply simp
+  done
 
 lemma Infinitesimal_sum_square_cancel [simp]:
-     "(x::hypreal)*x + y*y + z*z \<in> Infinitesimal ==> x*x \<in> Infinitesimal"
-apply (rule Infinitesimal_interval2, assumption)
-apply (rule_tac [2] zero_le_square, simp)
-apply (insert zero_le_square [of y])
-apply (insert zero_le_square [of z], simp del:zero_le_square)
-done
+  "x * x + y * y + z * z \<in> Infinitesimal \<Longrightarrow> x * x \<in> Infinitesimal"
+  for x y z :: hypreal
+  apply (rule Infinitesimal_interval2, assumption)
+    apply (rule_tac [2] zero_le_square, simp)
+  apply (insert zero_le_square [of y])
+  apply (insert zero_le_square [of z], simp del:zero_le_square)
+  done
 
-lemma HFinite_sum_square_cancel [simp]:
-     "(x::hypreal)*x + y*y + z*z \<in> HFinite ==> x*x \<in> HFinite"
-apply (rule HFinite_bounded, assumption)
-apply (rule_tac [2] zero_le_square)
-apply (insert zero_le_square [of y])
-apply (insert zero_le_square [of z], simp del:zero_le_square)
-done
+lemma HFinite_sum_square_cancel [simp]: "x * x + y * y + z * z \<in> HFinite \<Longrightarrow> x * x \<in> HFinite"
+  for x y z :: hypreal
+  apply (rule HFinite_bounded, assumption)
+   apply (rule_tac [2] zero_le_square)
+  apply (insert zero_le_square [of y])
+  apply (insert zero_le_square [of z], simp del:zero_le_square)
+  done
 
 lemma Infinitesimal_sum_square_cancel2 [simp]:
-     "(y::hypreal)*y + x*x + z*z \<in> Infinitesimal ==> x*x \<in> Infinitesimal"
-apply (rule Infinitesimal_sum_square_cancel)
-apply (simp add: ac_simps)
-done
+  "y * y + x * x + z * z \<in> Infinitesimal \<Longrightarrow> x * x \<in> Infinitesimal"
+  for x y z :: hypreal
+  apply (rule Infinitesimal_sum_square_cancel)
+  apply (simp add: ac_simps)
+  done
 
-lemma HFinite_sum_square_cancel2 [simp]:
-     "(y::hypreal)*y + x*x + z*z \<in> HFinite ==> x*x \<in> HFinite"
-apply (rule HFinite_sum_square_cancel)
-apply (simp add: ac_simps)
-done
+lemma HFinite_sum_square_cancel2 [simp]: "y * y + x * x + z * z \<in> HFinite \<Longrightarrow> x * x \<in> HFinite"
+  for x y z :: hypreal
+  apply (rule HFinite_sum_square_cancel)
+  apply (simp add: ac_simps)
+  done
 
 lemma Infinitesimal_sum_square_cancel3 [simp]:
-     "(z::hypreal)*z + y*y + x*x \<in> Infinitesimal ==> x*x \<in> Infinitesimal"
-apply (rule Infinitesimal_sum_square_cancel)
-apply (simp add: ac_simps)
-done
+  "z * z + y * y + x * x \<in> Infinitesimal \<Longrightarrow> x * x \<in> Infinitesimal"
+  for x y z :: hypreal
+  apply (rule Infinitesimal_sum_square_cancel)
+  apply (simp add: ac_simps)
+  done
 
-lemma HFinite_sum_square_cancel3 [simp]:
-     "(z::hypreal)*z + y*y + x*x \<in> HFinite ==> x*x \<in> HFinite"
-apply (rule HFinite_sum_square_cancel)
-apply (simp add: ac_simps)
-done
+lemma HFinite_sum_square_cancel3 [simp]: "z * z + y * y + x * x \<in> HFinite \<Longrightarrow> x * x \<in> HFinite"
+  for x y z :: hypreal
+  apply (rule HFinite_sum_square_cancel)
+  apply (simp add: ac_simps)
+  done
 
-lemma monad_hrabs_less:
-     "[| y \<in> monad x; 0 < hypreal_of_real e |]
-      ==> \<bar>y - x\<bar> < hypreal_of_real e"
-apply (drule mem_monad_approx [THEN approx_sym])
-apply (drule bex_Infinitesimal_iff [THEN iffD2])
-apply (auto dest!: InfinitesimalD)
-done
+lemma monad_hrabs_less: "y \<in> monad x \<Longrightarrow> 0 < hypreal_of_real e \<Longrightarrow> \<bar>y - x\<bar> < hypreal_of_real e"
+  apply (drule mem_monad_approx [THEN approx_sym])
+  apply (drule bex_Infinitesimal_iff [THEN iffD2])
+  apply (auto dest!: InfinitesimalD)
+  done
 
-lemma mem_monad_SReal_HFinite:
-     "x \<in> monad (hypreal_of_real  a) ==> x \<in> HFinite"
-apply (drule mem_monad_approx [THEN approx_sym])
-apply (drule bex_Infinitesimal_iff2 [THEN iffD2])
-apply (safe dest!: Infinitesimal_subset_HFinite [THEN subsetD])
-apply (erule SReal_hypreal_of_real [THEN SReal_subset_HFinite [THEN subsetD], THEN HFinite_add])
-done
+lemma mem_monad_SReal_HFinite: "x \<in> monad (hypreal_of_real  a) \<Longrightarrow> x \<in> HFinite"
+  apply (drule mem_monad_approx [THEN approx_sym])
+  apply (drule bex_Infinitesimal_iff2 [THEN iffD2])
+  apply (safe dest!: Infinitesimal_subset_HFinite [THEN subsetD])
+  apply (erule SReal_hypreal_of_real [THEN SReal_subset_HFinite [THEN subsetD], THEN HFinite_add])
+  done
 
 
-subsection\<open>Theorems about Standard Part\<close>
+subsection \<open>Theorems about Standard Part\<close>
 
-lemma st_approx_self: "x \<in> HFinite ==> st x \<approx> x"
-apply (simp add: st_def)
-apply (frule st_part_Ex, safe)
-apply (rule someI2)
-apply (auto intro: approx_sym)
-done
+lemma st_approx_self: "x \<in> HFinite \<Longrightarrow> st x \<approx> x"
+  apply (simp add: st_def)
+  apply (frule st_part_Ex, safe)
+  apply (rule someI2)
+   apply (auto intro: approx_sym)
+  done
 
-lemma st_SReal: "x \<in> HFinite ==> st x \<in> \<real>"
-apply (simp add: st_def)
-apply (frule st_part_Ex, safe)
-apply (rule someI2)
-apply (auto intro: approx_sym)
-done
+lemma st_SReal: "x \<in> HFinite \<Longrightarrow> st x \<in> \<real>"
+  apply (simp add: st_def)
+  apply (frule st_part_Ex, safe)
+  apply (rule someI2)
+   apply (auto intro: approx_sym)
+  done
 
-lemma st_HFinite: "x \<in> HFinite ==> st x \<in> HFinite"
-by (erule st_SReal [THEN SReal_subset_HFinite [THEN subsetD]])
+lemma st_HFinite: "x \<in> HFinite \<Longrightarrow> st x \<in> HFinite"
+  by (erule st_SReal [THEN SReal_subset_HFinite [THEN subsetD]])
 
-lemma st_unique: "\<lbrakk>r \<in> \<real>; r \<approx> x\<rbrakk> \<Longrightarrow> st x = r"
-apply (frule SReal_subset_HFinite [THEN subsetD])
-apply (drule (1) approx_HFinite)
-apply (unfold st_def)
-apply (rule some_equality)
-apply (auto intro: approx_unique_real)
-done
+lemma st_unique: "r \<in> \<real> \<Longrightarrow> r \<approx> x \<Longrightarrow> st x = r"
+  apply (frule SReal_subset_HFinite [THEN subsetD])
+  apply (drule (1) approx_HFinite)
+  apply (unfold st_def)
+  apply (rule some_equality)
+   apply (auto intro: approx_unique_real)
+  done
 
-lemma st_SReal_eq: "x \<in> \<real> ==> st x = x"
+lemma st_SReal_eq: "x \<in> \<real> \<Longrightarrow> st x = x"
   by (metis approx_refl st_unique)
 
 lemma st_hypreal_of_real [simp]: "st (hypreal_of_real x) = hypreal_of_real x"
-by (rule SReal_hypreal_of_real [THEN st_SReal_eq])
+  by (rule SReal_hypreal_of_real [THEN st_SReal_eq])
 
-lemma st_eq_approx: "[| x \<in> HFinite; y \<in> HFinite; st x = st y |] ==> x \<approx> y"
-by (auto dest!: st_approx_self elim!: approx_trans3)
+lemma st_eq_approx: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> st x = st y \<Longrightarrow> x \<approx> y"
+  by (auto dest!: st_approx_self elim!: approx_trans3)
 
 lemma approx_st_eq:
   assumes x: "x \<in> HFinite" and y: "y \<in> HFinite" and xy: "x \<approx> y"
@@ -1796,32 +1665,27 @@
     by (fast elim: approx_trans approx_trans2 SReal_approx_iff [THEN iffD1])
 qed
 
-lemma st_eq_approx_iff:
-     "[| x \<in> HFinite; y \<in> HFinite|]
-                   ==> (x \<approx> y) = (st x = st y)"
-by (blast intro: approx_st_eq st_eq_approx)
+lemma st_eq_approx_iff: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x \<approx> y \<longleftrightarrow> st x = st y"
+  by (blast intro: approx_st_eq st_eq_approx)
 
-lemma st_Infinitesimal_add_SReal:
-     "[| x \<in> \<real>; e \<in> Infinitesimal |] ==> st(x + e) = x"
-apply (erule st_unique)
-apply (erule Infinitesimal_add_approx_self)
-done
+lemma st_Infinitesimal_add_SReal: "x \<in> \<real> \<Longrightarrow> e \<in> Infinitesimal \<Longrightarrow> st (x + e) = x"
+  apply (erule st_unique)
+  apply (erule Infinitesimal_add_approx_self)
+  done
 
-lemma st_Infinitesimal_add_SReal2:
-     "[| x \<in> \<real>; e \<in> Infinitesimal |] ==> st(e + x) = x"
-apply (erule st_unique)
-apply (erule Infinitesimal_add_approx_self2)
-done
+lemma st_Infinitesimal_add_SReal2: "x \<in> \<real> \<Longrightarrow> e \<in> Infinitesimal \<Longrightarrow> st (e + x) = x"
+  apply (erule st_unique)
+  apply (erule Infinitesimal_add_approx_self2)
+  done
 
-lemma HFinite_st_Infinitesimal_add:
-     "x \<in> HFinite ==> \<exists>e \<in> Infinitesimal. x = st(x) + e"
-by (blast dest!: st_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2])
+lemma HFinite_st_Infinitesimal_add: "x \<in> HFinite \<Longrightarrow> \<exists>e \<in> Infinitesimal. x = st(x) + e"
+  by (blast dest!: st_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2])
 
-lemma st_add: "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> st (x + y) = st x + st y"
-by (simp add: st_unique st_SReal st_approx_self approx_add)
+lemma st_add: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> st (x + y) = st x + st y"
+  by (simp add: st_unique st_SReal st_approx_self approx_add)
 
 lemma st_numeral [simp]: "st (numeral w) = numeral w"
-by (rule Reals_numeral [THEN st_SReal_eq])
+  by (rule Reals_numeral [THEN st_SReal_eq])
 
 lemma st_neg_numeral [simp]: "st (- numeral w) = - numeral w"
 proof -
@@ -1831,79 +1695,71 @@
 qed
 
 lemma st_0 [simp]: "st 0 = 0"
-by (simp add: st_SReal_eq)
+  by (simp add: st_SReal_eq)
 
 lemma st_1 [simp]: "st 1 = 1"
-by (simp add: st_SReal_eq)
+  by (simp add: st_SReal_eq)
 
 lemma st_neg_1 [simp]: "st (- 1) = - 1"
-by (simp add: st_SReal_eq)
+  by (simp add: st_SReal_eq)
 
 lemma st_minus: "x \<in> HFinite \<Longrightarrow> st (- x) = - st x"
-by (simp add: st_unique st_SReal st_approx_self approx_minus)
+  by (simp add: st_unique st_SReal st_approx_self approx_minus)
 
 lemma st_diff: "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> st (x - y) = st x - st y"
-by (simp add: st_unique st_SReal st_approx_self approx_diff)
+  by (simp add: st_unique st_SReal st_approx_self approx_diff)
 
 lemma st_mult: "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> st (x * y) = st x * st y"
-by (simp add: st_unique st_SReal st_approx_self approx_mult_HFinite)
+  by (simp add: st_unique st_SReal st_approx_self approx_mult_HFinite)
 
-lemma st_Infinitesimal: "x \<in> Infinitesimal ==> st x = 0"
-by (simp add: st_unique mem_infmal_iff)
+lemma st_Infinitesimal: "x \<in> Infinitesimal \<Longrightarrow> st x = 0"
+  by (simp add: st_unique mem_infmal_iff)
 
-lemma st_not_Infinitesimal: "st(x) \<noteq> 0 ==> x \<notin> Infinitesimal"
+lemma st_not_Infinitesimal: "st(x) \<noteq> 0 \<Longrightarrow> x \<notin> Infinitesimal"
 by (fast intro: st_Infinitesimal)
 
-lemma st_inverse:
-     "[| x \<in> HFinite; st x \<noteq> 0 |]
-      ==> st(inverse x) = inverse (st x)"
-apply (rule_tac c1 = "st x" in mult_left_cancel [THEN iffD1])
-apply (auto simp add: st_mult [symmetric] st_not_Infinitesimal HFinite_inverse)
-apply (subst right_inverse, auto)
-done
+lemma st_inverse: "x \<in> HFinite \<Longrightarrow> st x \<noteq> 0 \<Longrightarrow> st (inverse x) = inverse (st x)"
+  apply (rule_tac c1 = "st x" in mult_left_cancel [THEN iffD1])
+   apply (auto simp add: st_mult [symmetric] st_not_Infinitesimal HFinite_inverse)
+  apply (subst right_inverse, auto)
+  done
 
-lemma st_divide [simp]:
-     "[| x \<in> HFinite; y \<in> HFinite; st y \<noteq> 0 |]
-      ==> st(x/y) = (st x) / (st y)"
-by (simp add: divide_inverse st_mult st_not_Infinitesimal HFinite_inverse st_inverse)
+lemma st_divide [simp]: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> st y \<noteq> 0 \<Longrightarrow> st (x / y) = st x / st y"
+  by (simp add: divide_inverse st_mult st_not_Infinitesimal HFinite_inverse st_inverse)
 
-lemma st_idempotent [simp]: "x \<in> HFinite ==> st(st(x)) = st(x)"
-by (blast intro: st_HFinite st_approx_self approx_st_eq)
+lemma st_idempotent [simp]: "x \<in> HFinite \<Longrightarrow> st (st x) = st x"
+  by (blast intro: st_HFinite st_approx_self approx_st_eq)
 
 lemma Infinitesimal_add_st_less:
-     "[| x \<in> HFinite; y \<in> HFinite; u \<in> Infinitesimal; st x < st y |]
-      ==> st x + u < st y"
-apply (drule st_SReal)+
-apply (auto intro!: Infinitesimal_add_hypreal_of_real_less simp add: SReal_iff)
-done
+  "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> u \<in> Infinitesimal \<Longrightarrow> st x < st y \<Longrightarrow> st x + u < st y"
+  apply (drule st_SReal)+
+  apply (auto intro!: Infinitesimal_add_hypreal_of_real_less simp add: SReal_iff)
+  done
 
 lemma Infinitesimal_add_st_le_cancel:
-     "[| x \<in> HFinite; y \<in> HFinite;
-         u \<in> Infinitesimal; st x \<le> st y + u
-      |] ==> st x \<le> st y"
-apply (simp add: linorder_not_less [symmetric])
-apply (auto dest: Infinitesimal_add_st_less)
-done
+  "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> u \<in> Infinitesimal \<Longrightarrow>
+    st x \<le> st y + u \<Longrightarrow> st x \<le> st y"
+  apply (simp add: linorder_not_less [symmetric])
+  apply (auto dest: Infinitesimal_add_st_less)
+  done
 
-lemma st_le: "[| x \<in> HFinite; y \<in> HFinite; x \<le> y |] ==> st(x) \<le> st(y)"
-by (metis approx_le_bound approx_sym linear st_SReal st_approx_self st_part_Ex1)
+lemma st_le: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x \<le> y \<Longrightarrow> st x \<le> st y"
+  by (metis approx_le_bound approx_sym linear st_SReal st_approx_self st_part_Ex1)
 
-lemma st_zero_le: "[| 0 \<le> x;  x \<in> HFinite |] ==> 0 \<le> st x"
-apply (subst st_0 [symmetric])
-apply (rule st_le, auto)
-done
+lemma st_zero_le: "0 \<le> x \<Longrightarrow> x \<in> HFinite \<Longrightarrow> 0 \<le> st x"
+  apply (subst st_0 [symmetric])
+  apply (rule st_le, auto)
+  done
 
-lemma st_zero_ge: "[| x \<le> 0;  x \<in> HFinite |] ==> st x \<le> 0"
-apply (subst st_0 [symmetric])
-apply (rule st_le, auto)
-done
+lemma st_zero_ge: "x \<le> 0 \<Longrightarrow> x \<in> HFinite \<Longrightarrow> st x \<le> 0"
+  apply (subst st_0 [symmetric])
+  apply (rule st_le, auto)
+  done
 
-lemma st_hrabs: "x \<in> HFinite ==> \<bar>st x\<bar> = st \<bar>x\<bar>"
-apply (simp add: linorder_not_le st_zero_le abs_if st_minus
-   linorder_not_less)
-apply (auto dest!: st_zero_ge [OF order_less_imp_le])
-done
-
+lemma st_hrabs: "x \<in> HFinite \<Longrightarrow> \<bar>st x\<bar> = st \<bar>x\<bar>"
+  apply (simp add: linorder_not_le st_zero_le abs_if st_minus linorder_not_less)
+  apply (auto dest!: st_zero_ge [OF order_less_imp_le])
+  done
 
 
 subsection \<open>Alternative Definitions using Free Ultrafilter\<close>
@@ -1911,78 +1767,73 @@
 subsubsection \<open>@{term HFinite}\<close>
 
 lemma HFinite_FreeUltrafilterNat:
-    "star_n X \<in> HFinite
-   ==> \<exists>u. eventually (\<lambda>n. norm (X n) < u) FreeUltrafilterNat"
-apply (auto simp add: HFinite_def SReal_def)
-apply (rule_tac x=r in exI)
-apply (simp add: hnorm_def star_of_def starfun_star_n)
-apply (simp add: star_less_def starP2_star_n)
-done
+  "star_n X \<in> HFinite \<Longrightarrow> \<exists>u. eventually (\<lambda>n. norm (X n) < u) FreeUltrafilterNat"
+  apply (auto simp add: HFinite_def SReal_def)
+  apply (rule_tac x=r in exI)
+  apply (simp add: hnorm_def star_of_def starfun_star_n)
+  apply (simp add: star_less_def starP2_star_n)
+  done
 
 lemma FreeUltrafilterNat_HFinite:
-     "\<exists>u. eventually (\<lambda>n. norm (X n) < u) FreeUltrafilterNat
-       ==>  star_n X \<in> HFinite"
-apply (auto simp add: HFinite_def mem_Rep_star_iff)
-apply (rule_tac x="star_of u" in bexI)
-apply (simp add: hnorm_def starfun_star_n star_of_def)
-apply (simp add: star_less_def starP2_star_n)
-apply (simp add: SReal_def)
-done
+  "\<exists>u. eventually (\<lambda>n. norm (X n) < u) FreeUltrafilterNat \<Longrightarrow> star_n X \<in> HFinite"
+  apply (auto simp add: HFinite_def mem_Rep_star_iff)
+  apply (rule_tac x="star_of u" in bexI)
+   apply (simp add: hnorm_def starfun_star_n star_of_def)
+   apply (simp add: star_less_def starP2_star_n)
+  apply (simp add: SReal_def)
+  done
 
 lemma HFinite_FreeUltrafilterNat_iff:
-     "(star_n X \<in> HFinite) = (\<exists>u. eventually (\<lambda>n. norm (X n) < u) FreeUltrafilterNat)"
-by (blast intro!: HFinite_FreeUltrafilterNat FreeUltrafilterNat_HFinite)
+  "star_n X \<in> HFinite \<longleftrightarrow> (\<exists>u. eventually (\<lambda>n. norm (X n) < u) FreeUltrafilterNat)"
+  by (blast intro!: HFinite_FreeUltrafilterNat FreeUltrafilterNat_HFinite)
+
 
 subsubsection \<open>@{term HInfinite}\<close>
 
 lemma lemma_Compl_eq: "- {n. u < norm (f n)} = {n. norm (f n) \<le> u}"
-by auto
+  by auto
 
 lemma lemma_Compl_eq2: "- {n. norm (f n) < u} = {n. u \<le> norm (f n)}"
-by auto
+  by auto
 
-lemma lemma_Int_eq1:
-     "{n. norm (f n) \<le> u} Int {n. u \<le> norm (f n)} = {n. norm(f n) = u}"
-by auto
+lemma lemma_Int_eq1: "{n. norm (f n) \<le> u} Int {n. u \<le> norm (f n)} = {n. norm(f n) = u}"
+  by auto
 
-lemma lemma_FreeUltrafilterNat_one:
-     "{n. norm (f n) = u} \<le> {n. norm (f n) < u + (1::real)}"
-by auto
+lemma lemma_FreeUltrafilterNat_one: "{n. norm (f n) = u} \<le> {n. norm (f n) < u + (1::real)}"
+  by auto
 
-(*-------------------------------------
-  Exclude this type of sets from free
-  ultrafilter for Infinite numbers!
- -------------------------------------*)
+text \<open>Exclude this type of sets from free ultrafilter for Infinite numbers!\<close>
 lemma FreeUltrafilterNat_const_Finite:
-     "eventually (\<lambda>n. norm (X n) = u) FreeUltrafilterNat ==> star_n X \<in> HFinite"
-apply (rule FreeUltrafilterNat_HFinite)
-apply (rule_tac x = "u + 1" in exI)
-apply (auto elim: eventually_mono)
-done
+  "eventually (\<lambda>n. norm (X n) = u) FreeUltrafilterNat \<Longrightarrow> star_n X \<in> HFinite"
+  apply (rule FreeUltrafilterNat_HFinite)
+  apply (rule_tac x = "u + 1" in exI)
+  apply (auto elim: eventually_mono)
+  done
 
 lemma HInfinite_FreeUltrafilterNat:
-     "star_n X \<in> HInfinite ==> \<forall>u. eventually (\<lambda>n. u < norm (X n)) FreeUltrafilterNat"
-apply (drule HInfinite_HFinite_iff [THEN iffD1])
-apply (simp add: HFinite_FreeUltrafilterNat_iff)
-apply (rule allI, drule_tac x="u + 1" in spec)
-apply (simp add: FreeUltrafilterNat.eventually_not_iff[symmetric])
-apply (auto elim: eventually_mono)
-done
+  "star_n X \<in> HInfinite \<Longrightarrow> \<forall>u. eventually (\<lambda>n. u < norm (X n)) FreeUltrafilterNat"
+  apply (drule HInfinite_HFinite_iff [THEN iffD1])
+  apply (simp add: HFinite_FreeUltrafilterNat_iff)
+  apply (rule allI, drule_tac x="u + 1" in spec)
+  apply (simp add: FreeUltrafilterNat.eventually_not_iff[symmetric])
+  apply (auto elim: eventually_mono)
+  done
 
-lemma lemma_Int_HI:
-     "{n. norm (Xa n) < u} Int {n. X n = Xa n} \<subseteq> {n. norm (X n) < (u::real)}"
-by auto
+lemma lemma_Int_HI: "{n. norm (Xa n) < u} \<inter> {n. X n = Xa n} \<subseteq> {n. norm (X n) < u}"
+  for u :: real
+  by auto
 
-lemma lemma_Int_HIa: "{n. u < norm (X n)} Int {n. norm (X n) < u} = {}"
-by (auto intro: order_less_asym)
+lemma lemma_Int_HIa: "{n. u < norm (X n)} \<inter> {n. norm (X n) < u} = {}"
+  by (auto intro: order_less_asym)
 
 lemma FreeUltrafilterNat_HInfinite:
-     "\<forall>u. eventually (\<lambda>n. u < norm (X n)) FreeUltrafilterNat ==> star_n X \<in> HInfinite"
-apply (rule HInfinite_HFinite_iff [THEN iffD2])
-apply (safe, drule HFinite_FreeUltrafilterNat, safe)
-apply (drule_tac x = u in spec)
+  "\<forall>u. eventually (\<lambda>n. u < norm (X n)) FreeUltrafilterNat \<Longrightarrow> star_n X \<in> HInfinite"
+  apply (rule HInfinite_HFinite_iff [THEN iffD2])
+  apply (safe, drule HFinite_FreeUltrafilterNat, safe)
+  apply (drule_tac x = u in spec)
 proof -
-  fix u assume "\<forall>\<^sub>Fn in \<U>. norm (X n) < u" "\<forall>\<^sub>Fn in \<U>. u < norm (X n)"
+  fix u
+  assume "\<forall>\<^sub>Fn in \<U>. norm (X n) < u" "\<forall>\<^sub>Fn in \<U>. u < norm (X n)"
   then have "\<forall>\<^sub>F x in \<U>. False"
     by eventually_elim auto
   then show False
@@ -1990,230 +1841,204 @@
 qed
 
 lemma HInfinite_FreeUltrafilterNat_iff:
-     "(star_n X \<in> HInfinite) = (\<forall>u. eventually (\<lambda>n. u < norm (X n)) FreeUltrafilterNat)"
-by (blast intro!: HInfinite_FreeUltrafilterNat FreeUltrafilterNat_HInfinite)
+  "star_n X \<in> HInfinite \<longleftrightarrow> (\<forall>u. eventually (\<lambda>n. u < norm (X n)) FreeUltrafilterNat)"
+  by (blast intro!: HInfinite_FreeUltrafilterNat FreeUltrafilterNat_HInfinite)
+
 
 subsubsection \<open>@{term Infinitesimal}\<close>
 
-lemma ball_SReal_eq: "(\<forall>x::hypreal \<in> Reals. P x) = (\<forall>x::real. P (star_of x))"
-by (unfold SReal_def, auto)
+lemma ball_SReal_eq: "(\<forall>x::hypreal \<in> Reals. P x) \<longleftrightarrow> (\<forall>x::real. P (star_of x))"
+  by (auto simp: SReal_def)
 
 lemma Infinitesimal_FreeUltrafilterNat:
-     "star_n X \<in> Infinitesimal ==> \<forall>u>0. eventually (\<lambda>n. norm (X n) < u) \<U>"
-apply (simp add: Infinitesimal_def ball_SReal_eq)
-apply (simp add: hnorm_def starfun_star_n star_of_def)
-apply (simp add: star_less_def starP2_star_n)
-done
+  "star_n X \<in> Infinitesimal \<Longrightarrow> \<forall>u>0. eventually (\<lambda>n. norm (X n) < u) \<U>"
+  apply (simp add: Infinitesimal_def ball_SReal_eq)
+  apply (simp add: hnorm_def starfun_star_n star_of_def)
+  apply (simp add: star_less_def starP2_star_n)
+  done
 
 lemma FreeUltrafilterNat_Infinitesimal:
-     "\<forall>u>0. eventually (\<lambda>n. norm (X n) < u) \<U> ==> star_n X \<in> Infinitesimal"
-apply (simp add: Infinitesimal_def ball_SReal_eq)
-apply (simp add: hnorm_def starfun_star_n star_of_def)
-apply (simp add: star_less_def starP2_star_n)
-done
+  "\<forall>u>0. eventually (\<lambda>n. norm (X n) < u) \<U> \<Longrightarrow> star_n X \<in> Infinitesimal"
+  apply (simp add: Infinitesimal_def ball_SReal_eq)
+  apply (simp add: hnorm_def starfun_star_n star_of_def)
+  apply (simp add: star_less_def starP2_star_n)
+  done
 
 lemma Infinitesimal_FreeUltrafilterNat_iff:
-     "(star_n X \<in> Infinitesimal) = (\<forall>u>0. eventually (\<lambda>n. norm (X n) < u) \<U>)"
-by (blast intro!: Infinitesimal_FreeUltrafilterNat FreeUltrafilterNat_Infinitesimal)
+  "(star_n X \<in> Infinitesimal) = (\<forall>u>0. eventually (\<lambda>n. norm (X n) < u) \<U>)"
+  by (blast intro!: Infinitesimal_FreeUltrafilterNat FreeUltrafilterNat_Infinitesimal)
+
 
-(*------------------------------------------------------------------------
-         Infinitesimals as smaller than 1/n for all n::nat (> 0)
- ------------------------------------------------------------------------*)
+text \<open>Infinitesimals as smaller than \<open>1/n\<close> for all \<open>n::nat (> 0)\<close>.\<close>
 
-lemma lemma_Infinitesimal:
-     "(\<forall>r. 0 < r --> x < r) = (\<forall>n. x < inverse(real (Suc n)))"
-apply (auto simp del: of_nat_Suc)
-apply (blast dest!: reals_Archimedean intro: order_less_trans)
-done
+lemma lemma_Infinitesimal: "(\<forall>r. 0 < r \<longrightarrow> x < r) \<longleftrightarrow> (\<forall>n. x < inverse (real (Suc n)))"
+  apply (auto simp del: of_nat_Suc)
+  apply (blast dest!: reals_Archimedean intro: order_less_trans)
+  done
 
 lemma lemma_Infinitesimal2:
-     "(\<forall>r \<in> Reals. 0 < r --> x < r) =
-      (\<forall>n. x < inverse(hypreal_of_nat (Suc n)))"
-apply safe
- apply (drule_tac x = "inverse (hypreal_of_real (real (Suc n))) " in bspec)
-  apply simp_all
+  "(\<forall>r \<in> Reals. 0 < r \<longrightarrow> x < r) \<longleftrightarrow> (\<forall>n. x < inverse(hypreal_of_nat (Suc n)))"
+  apply safe
+   apply (drule_tac x = "inverse (hypreal_of_real (real (Suc n))) " in bspec)
+    apply simp_all
   using less_imp_of_nat_less apply fastforce
-apply (auto dest!: reals_Archimedean simp add: SReal_iff simp del: of_nat_Suc)
-apply (drule star_of_less [THEN iffD2])
-apply simp
-apply (blast intro: order_less_trans)
-done
+  apply (auto dest!: reals_Archimedean simp add: SReal_iff simp del: of_nat_Suc)
+  apply (drule star_of_less [THEN iffD2])
+  apply simp
+  apply (blast intro: order_less_trans)
+  done
 
 
 lemma Infinitesimal_hypreal_of_nat_iff:
-     "Infinitesimal = {x. \<forall>n. hnorm x < inverse (hypreal_of_nat (Suc n))}"
-apply (simp add: Infinitesimal_def)
-apply (auto simp add: lemma_Infinitesimal2)
-done
+  "Infinitesimal = {x. \<forall>n. hnorm x < inverse (hypreal_of_nat (Suc n))}"
+  apply (simp add: Infinitesimal_def)
+  apply (auto simp add: lemma_Infinitesimal2)
+  done
 
 
-subsection\<open>Proof that \<open>\<omega>\<close> is an infinite number\<close>
+subsection \<open>Proof that \<open>\<omega>\<close> is an infinite number\<close>
 
-text\<open>It will follow that \<open>\<epsilon>\<close> is an infinitesimal number.\<close>
+text \<open>It will follow that \<open>\<epsilon>\<close> is an infinitesimal number.\<close>
 
 lemma Suc_Un_eq: "{n. n < Suc m} = {n. n < m} Un {n. n = m}"
-by (auto simp add: less_Suc_eq)
+  by (auto simp add: less_Suc_eq)
 
-(*-------------------------------------------
-  Prove that any segment is finite and hence cannot belong to FreeUltrafilterNat
- -------------------------------------------*)
+
+text \<open>Prove that any segment is finite and hence cannot belong to \<open>FreeUltrafilterNat\<close>.\<close>
 
 lemma finite_real_of_nat_segment: "finite {n::nat. real n < real (m::nat)}"
-  by (auto intro: finite_Collect_less_nat)
+  by auto
 
 lemma finite_real_of_nat_less_real: "finite {n::nat. real n < u}"
-apply (cut_tac x = u in reals_Archimedean2, safe)
-apply (rule finite_real_of_nat_segment [THEN [2] finite_subset])
-apply (auto dest: order_less_trans)
-done
+  apply (cut_tac x = u in reals_Archimedean2, safe)
+  apply (rule finite_real_of_nat_segment [THEN [2] finite_subset])
+  apply (auto dest: order_less_trans)
+  done
 
-lemma lemma_real_le_Un_eq:
-     "{n. f n \<le> u} = {n. f n < u} Un {n. u = (f n :: real)}"
-by (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le)
+lemma lemma_real_le_Un_eq: "{n. f n \<le> u} = {n. f n < u} \<union> {n. u = (f n :: real)}"
+  by (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le)
 
 lemma finite_real_of_nat_le_real: "finite {n::nat. real n \<le> u}"
-by (auto simp add: lemma_real_le_Un_eq lemma_finite_omega_set finite_real_of_nat_less_real)
+  by (auto simp add: lemma_real_le_Un_eq lemma_finite_omega_set finite_real_of_nat_less_real)
 
 lemma finite_rabs_real_of_nat_le_real: "finite {n::nat. \<bar>real n\<bar> \<le> u}"
-apply (simp (no_asm) add: finite_real_of_nat_le_real)
-done
+  by (simp add: finite_real_of_nat_le_real)
 
 lemma rabs_real_of_nat_le_real_FreeUltrafilterNat:
-     "\<not> eventually (\<lambda>n. \<bar>real n\<bar> \<le> u) FreeUltrafilterNat"
-by (blast intro!: FreeUltrafilterNat.finite finite_rabs_real_of_nat_le_real)
+  "\<not> eventually (\<lambda>n. \<bar>real n\<bar> \<le> u) FreeUltrafilterNat"
+  by (blast intro!: FreeUltrafilterNat.finite finite_rabs_real_of_nat_le_real)
 
 lemma FreeUltrafilterNat_nat_gt_real: "eventually (\<lambda>n. u < real n) FreeUltrafilterNat"
-apply (rule FreeUltrafilterNat.finite')
-apply (subgoal_tac "{n::nat. \<not> u < real n} = {n. real n \<le> u}")
-apply (auto simp add: finite_real_of_nat_le_real)
-done
+  apply (rule FreeUltrafilterNat.finite')
+  apply (subgoal_tac "{n::nat. \<not> u < real n} = {n. real n \<le> u}")
+   apply (auto simp add: finite_real_of_nat_le_real)
+  done
 
-(*--------------------------------------------------------------
- The complement of {n. \<bar>real n\<bar> \<le> u} =
- {n. u < \<bar>real n\<bar>} is in FreeUltrafilterNat
- by property of (free) ultrafilters
- --------------------------------------------------------------*)
+text \<open>The complement of \<open>{n. \<bar>real n\<bar> \<le> u} = {n. u < \<bar>real n\<bar>}\<close> is in
+  \<open>FreeUltrafilterNat\<close> by property of (free) ultrafilters.\<close>
 
 lemma Compl_real_le_eq: "- {n::nat. real n \<le> u} = {n. u < real n}"
-by (auto dest!: order_le_less_trans simp add: linorder_not_le)
+  by (auto dest!: order_le_less_trans simp add: linorder_not_le)
 
-text\<open>@{term \<omega>} is a member of @{term HInfinite}\<close>
-
+text \<open>@{term \<omega>} is a member of @{term HInfinite}.\<close>
 theorem HInfinite_omega [simp]: "\<omega> \<in> HInfinite"
-apply (simp add: omega_def)
-apply (rule FreeUltrafilterNat_HInfinite)
-apply clarify
-apply (rule_tac u1 = "u-1" in eventually_mono [OF FreeUltrafilterNat_nat_gt_real])
-apply auto
-done
+  apply (simp add: omega_def)
+  apply (rule FreeUltrafilterNat_HInfinite)
+  apply clarify
+  apply (rule_tac u1 = "u-1" in eventually_mono [OF FreeUltrafilterNat_nat_gt_real])
+  apply auto
+  done
 
-(*-----------------------------------------------
-       Epsilon is a member of Infinitesimal
- -----------------------------------------------*)
+
+text \<open>Epsilon is a member of Infinitesimal.\<close>
 
 lemma Infinitesimal_epsilon [simp]: "\<epsilon> \<in> Infinitesimal"
-by (auto intro!: HInfinite_inverse_Infinitesimal HInfinite_omega simp add: hypreal_epsilon_inverse_omega)
+  by (auto intro!: HInfinite_inverse_Infinitesimal HInfinite_omega
+      simp add: hypreal_epsilon_inverse_omega)
 
 lemma HFinite_epsilon [simp]: "\<epsilon> \<in> HFinite"
-by (auto intro: Infinitesimal_subset_HFinite [THEN subsetD])
+  by (auto intro: Infinitesimal_subset_HFinite [THEN subsetD])
 
 lemma epsilon_approx_zero [simp]: "\<epsilon> \<approx> 0"
-apply (simp (no_asm) add: mem_infmal_iff [symmetric])
-done
-
-(*------------------------------------------------------------------------
-  Needed for proof that we define a hyperreal [<X(n)] \<approx> hypreal_of_real a given
-  that \<forall>n. |X n - a| < 1/n. Used in proof of NSLIM => LIM.
- -----------------------------------------------------------------------*)
+  by (simp add: mem_infmal_iff [symmetric])
 
-lemma real_of_nat_less_inverse_iff:
-     "0 < u  ==> (u < inverse (real(Suc n))) = (real(Suc n) < inverse u)"
-apply (simp add: inverse_eq_divide)
-apply (subst pos_less_divide_eq, assumption)
-apply (subst pos_less_divide_eq)
- apply simp
-apply (simp add: mult.commute)
-done
+text \<open>Needed for proof that we define a hyperreal \<open>[<X(n)] \<approx> hypreal_of_real a\<close> given
+  that \<open>\<forall>n. |X n - a| < 1/n\<close>. Used in proof of \<open>NSLIM \<Rightarrow> LIM\<close>.\<close>
+lemma real_of_nat_less_inverse_iff: "0 < u \<Longrightarrow> u < inverse (real(Suc n)) \<longleftrightarrow> real(Suc n) < inverse u"
+  apply (simp add: inverse_eq_divide)
+  apply (subst pos_less_divide_eq, assumption)
+  apply (subst pos_less_divide_eq)
+   apply simp
+  apply (simp add: mult.commute)
+  done
 
-lemma finite_inverse_real_of_posnat_gt_real:
-     "0 < u ==> finite {n. u < inverse(real(Suc n))}"
+lemma finite_inverse_real_of_posnat_gt_real: "0 < u \<Longrightarrow> finite {n. u < inverse (real (Suc n))}"
 proof (simp only: real_of_nat_less_inverse_iff)
   have "{n. 1 + real n < inverse u} = {n. real n < inverse u - 1}"
     by fastforce
-  thus "finite {n. real (Suc n) < inverse u}"
-    using finite_real_of_nat_less_real [of "inverse u - 1"] by auto
+  then show "finite {n. real (Suc n) < inverse u}"
+    using finite_real_of_nat_less_real [of "inverse u - 1"]
+    by auto
 qed
 
 lemma lemma_real_le_Un_eq2:
-     "{n. u \<le> inverse(real(Suc n))} =
-     {n. u < inverse(real(Suc n))} Un {n. u = inverse(real(Suc n))}"
-by (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le)
+  "{n. u \<le> inverse(real(Suc n))} =
+    {n. u < inverse(real(Suc n))} \<union> {n. u = inverse(real(Suc n))}"
+  by (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le)
 
-lemma finite_inverse_real_of_posnat_ge_real:
-     "0 < u ==> finite {n. u \<le> inverse(real(Suc n))}"
-by (auto simp add: lemma_real_le_Un_eq2 lemma_finite_epsilon_set finite_inverse_real_of_posnat_gt_real
-            simp del: of_nat_Suc)
+lemma finite_inverse_real_of_posnat_ge_real: "0 < u \<Longrightarrow> finite {n. u \<le> inverse (real (Suc n))}"
+  by (auto simp add: lemma_real_le_Un_eq2 lemma_finite_epsilon_set finite_inverse_real_of_posnat_gt_real
+      simp del: of_nat_Suc)
 
 lemma inverse_real_of_posnat_ge_real_FreeUltrafilterNat:
-     "0 < u ==> \<not> eventually (\<lambda>n. u \<le> inverse(real(Suc n))) FreeUltrafilterNat"
-by (blast intro!: FreeUltrafilterNat.finite finite_inverse_real_of_posnat_ge_real)
+  "0 < u \<Longrightarrow> \<not> eventually (\<lambda>n. u \<le> inverse(real(Suc n))) FreeUltrafilterNat"
+  by (blast intro!: FreeUltrafilterNat.finite finite_inverse_real_of_posnat_ge_real)
 
-(*--------------------------------------------------------------
-    The complement of  {n. u \<le> inverse(real(Suc n))} =
-    {n. inverse(real(Suc n)) < u} is in FreeUltrafilterNat
-    by property of (free) ultrafilters
- --------------------------------------------------------------*)
-lemma Compl_le_inverse_eq:
-     "- {n. u \<le> inverse(real(Suc n))} = {n. inverse(real(Suc n)) < u}"
-by (auto dest!: order_le_less_trans simp add: linorder_not_le)
+text \<open>The complement of \<open>{n. u \<le> inverse(real(Suc n))} = {n. inverse (real (Suc n)) < u}\<close>
+  is in \<open>FreeUltrafilterNat\<close> by property of (free) ultrafilters.\<close>
+lemma Compl_le_inverse_eq: "- {n. u \<le> inverse(real(Suc n))} = {n. inverse(real(Suc n)) < u}"
+  by (auto dest!: order_le_less_trans simp add: linorder_not_le)
 
 
 lemma FreeUltrafilterNat_inverse_real_of_posnat:
-     "0 < u ==> eventually (\<lambda>n. inverse(real(Suc n)) < u) FreeUltrafilterNat"
-by (drule inverse_real_of_posnat_ge_real_FreeUltrafilterNat)
-   (simp add: FreeUltrafilterNat.eventually_not_iff not_le[symmetric])
+  "0 < u \<Longrightarrow> eventually (\<lambda>n. inverse(real(Suc n)) < u) FreeUltrafilterNat"
+  by (drule inverse_real_of_posnat_ge_real_FreeUltrafilterNat)
+    (simp add: FreeUltrafilterNat.eventually_not_iff not_le[symmetric])
 
-text\<open>Example of an hypersequence (i.e. an extended standard sequence)
-   whose term with an hypernatural suffix is an infinitesimal i.e.
-   the whn'nth term of the hypersequence is a member of Infinitesimal\<close>
+text \<open>Example of an hypersequence (i.e. an extended standard sequence)
+  whose term with an hypernatural suffix is an infinitesimal i.e.
+  the whn'nth term of the hypersequence is a member of Infinitesimal\<close>
 
-lemma SEQ_Infinitesimal:
-      "( *f* (%n::nat. inverse(real(Suc n)))) whn : Infinitesimal"
-by (simp add: hypnat_omega_def starfun_star_n star_n_inverse Infinitesimal_FreeUltrafilterNat_iff
-        FreeUltrafilterNat_inverse_real_of_posnat del: of_nat_Suc)
+lemma SEQ_Infinitesimal: "( *f* (\<lambda>n::nat. inverse(real(Suc n)))) whn \<in> Infinitesimal"
+  by (simp add: hypnat_omega_def starfun_star_n star_n_inverse Infinitesimal_FreeUltrafilterNat_iff
+      FreeUltrafilterNat_inverse_real_of_posnat del: of_nat_Suc)
 
-text\<open>Example where we get a hyperreal from a real sequence
-      for which a particular property holds. The theorem is
-      used in proofs about equivalence of nonstandard and
-      standard neighbourhoods. Also used for equivalence of
-      nonstandard ans standard definitions of pointwise
-      limit.\<close>
+text \<open>Example where we get a hyperreal from a real sequence
+  for which a particular property holds. The theorem is
+  used in proofs about equivalence of nonstandard and
+  standard neighbourhoods. Also used for equivalence of
+  nonstandard ans standard definitions of pointwise
+  limit.\<close>
 
-(*-----------------------------------------------------
-    |X(n) - x| < 1/n ==> [<X n>] - hypreal_of_real x| \<in> Infinitesimal
- -----------------------------------------------------*)
+text \<open>\<open>|X(n) - x| < 1/n \<Longrightarrow> [<X n>] - hypreal_of_real x| \<in> Infinitesimal\<close>\<close>
 lemma real_seq_to_hypreal_Infinitesimal:
-     "\<forall>n. norm(X n - x) < inverse(real(Suc n))
-     ==> star_n X - star_of x \<in> Infinitesimal"
-unfolding star_n_diff star_of_def Infinitesimal_FreeUltrafilterNat_iff star_n_inverse
-by (auto dest!: FreeUltrafilterNat_inverse_real_of_posnat
-         intro: order_less_trans elim!: eventually_mono)
+  "\<forall>n. norm (X n - x) < inverse (real (Suc n)) \<Longrightarrow> star_n X - star_of x \<in> Infinitesimal"
+  unfolding star_n_diff star_of_def Infinitesimal_FreeUltrafilterNat_iff star_n_inverse
+  by (auto dest!: FreeUltrafilterNat_inverse_real_of_posnat
+      intro: order_less_trans elim!: eventually_mono)
 
 lemma real_seq_to_hypreal_approx:
-     "\<forall>n. norm(X n - x) < inverse(real(Suc n))
-      ==> star_n X \<approx> star_of x"
-by (metis bex_Infinitesimal_iff real_seq_to_hypreal_Infinitesimal)
+  "\<forall>n. norm (X n - x) < inverse (real (Suc n)) \<Longrightarrow> star_n X \<approx> star_of x"
+  by (metis bex_Infinitesimal_iff real_seq_to_hypreal_Infinitesimal)
 
 lemma real_seq_to_hypreal_approx2:
-     "\<forall>n. norm(x - X n) < inverse(real(Suc n))
-               ==> star_n X \<approx> star_of x"
-by (metis norm_minus_commute real_seq_to_hypreal_approx)
+  "\<forall>n. norm (x - X n) < inverse(real(Suc n)) \<Longrightarrow> star_n X \<approx> star_of x"
+  by (metis norm_minus_commute real_seq_to_hypreal_approx)
 
 lemma real_seq_to_hypreal_Infinitesimal2:
-     "\<forall>n. norm(X n - Y n) < inverse(real(Suc n))
-      ==> star_n X - star_n Y \<in> Infinitesimal"
-unfolding Infinitesimal_FreeUltrafilterNat_iff star_n_diff
-by (auto dest!: FreeUltrafilterNat_inverse_real_of_posnat
-         intro: order_less_trans elim!: eventually_mono)
+  "\<forall>n. norm(X n - Y n) < inverse(real(Suc n)) \<Longrightarrow> star_n X - star_n Y \<in> Infinitesimal"
+  unfolding Infinitesimal_FreeUltrafilterNat_iff star_n_diff
+  by (auto dest!: FreeUltrafilterNat_inverse_real_of_posnat
+      intro: order_less_trans elim!: eventually_mono)
 
 end
--- a/src/HOL/Nonstandard_Analysis/NSComplex.thy	Mon Oct 31 16:26:36 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/NSComplex.thy	Tue Nov 01 00:44:24 2016 +0100
@@ -3,654 +3,604 @@
     Author:     Lawrence C Paulson
 *)
 
-section\<open>Nonstandard Complex Numbers\<close>
+section \<open>Nonstandard Complex Numbers\<close>
 
 theory NSComplex
-imports NSA
+  imports NSA
 begin
 
 type_synonym hcomplex = "complex star"
 
-abbreviation
-  hcomplex_of_complex :: "complex => complex star" where
-  "hcomplex_of_complex == star_of"
+abbreviation hcomplex_of_complex :: "complex \<Rightarrow> complex star"
+  where "hcomplex_of_complex \<equiv> star_of"
 
-abbreviation
-  hcmod :: "complex star => real star" where
-  "hcmod == hnorm"
+abbreviation hcmod :: "complex star \<Rightarrow> real star"
+  where "hcmod \<equiv> hnorm"
 
 
-  (*--- real and Imaginary parts ---*)
+subsubsection \<open>Real and Imaginary parts\<close>
+
+definition hRe :: "hcomplex \<Rightarrow> hypreal"
+  where "hRe = *f* Re"
+
+definition hIm :: "hcomplex \<Rightarrow> hypreal"
+  where "hIm = *f* Im"
+
 
-definition
-  hRe :: "hcomplex => hypreal" where
-  "hRe = *f* Re"
+subsubsection \<open>Imaginary unit\<close>
+
+definition iii :: hcomplex
+  where "iii = star_of \<i>"
 
-definition
-  hIm :: "hcomplex => hypreal" where
-  "hIm = *f* Im"
+
+subsubsection \<open>Complex conjugate\<close>
+
+definition hcnj :: "hcomplex \<Rightarrow> hcomplex"
+  where "hcnj = *f* cnj"
 
 
-  (*------ imaginary unit ----------*)
+subsubsection \<open>Argand\<close>
 
-definition
-  iii :: hcomplex where
-  "iii = star_of \<i>"
-
-  (*------- complex conjugate ------*)
+definition hsgn :: "hcomplex \<Rightarrow> hcomplex"
+  where "hsgn = *f* sgn"
 
-definition
-  hcnj :: "hcomplex => hcomplex" where
-  "hcnj = *f* cnj"
-
-  (*------------ Argand -------------*)
+definition harg :: "hcomplex \<Rightarrow> hypreal"
+  where "harg = *f* arg"
 
-definition
-  hsgn :: "hcomplex => hcomplex" where
-  "hsgn = *f* sgn"
+definition  \<comment> \<open>abbreviation for \<open>cos a + i sin a\<close>\<close>
+  hcis :: "hypreal \<Rightarrow> hcomplex"
+  where "hcis = *f* cis"
 
-definition
-  harg :: "hcomplex => hypreal" where
-  "harg = *f* arg"
+
+subsubsection \<open>Injection from hyperreals\<close>
 
-definition
-  (* abbreviation for (cos a + i sin a) *)
-  hcis :: "hypreal => hcomplex" where
-  "hcis = *f* cis"
+abbreviation hcomplex_of_hypreal :: "hypreal \<Rightarrow> hcomplex"
+  where "hcomplex_of_hypreal \<equiv> of_hypreal"
 
-  (*----- injection from hyperreals -----*)
+definition  \<comment> \<open>abbreviation for \<open>r * (cos a + i sin a)\<close>\<close>
+  hrcis :: "hypreal \<Rightarrow> hypreal \<Rightarrow> hcomplex"
+  where "hrcis = *f2* rcis"
 
-abbreviation
-  hcomplex_of_hypreal :: "hypreal \<Rightarrow> hcomplex" where
-  "hcomplex_of_hypreal \<equiv> of_hypreal"
 
-definition
-  (* abbreviation for r*(cos a + i sin a) *)
-  hrcis :: "[hypreal, hypreal] => hcomplex" where
-  "hrcis = *f2* rcis"
+subsubsection \<open>\<open>e ^ (x + iy)\<close>\<close>
 
-  (*------------ e ^ (x + iy) ------------*)
-definition
-  hExp :: "hcomplex => hcomplex" where
-  "hExp = *f* exp"
+definition hExp :: "hcomplex \<Rightarrow> hcomplex"
+  where "hExp = *f* exp"
 
-definition
-  HComplex :: "[hypreal,hypreal] => hcomplex" where
-  "HComplex = *f2* Complex"
+definition HComplex :: "hypreal \<Rightarrow> hypreal \<Rightarrow> hcomplex"
+  where "HComplex = *f2* Complex"
 
 lemmas hcomplex_defs [transfer_unfold] =
   hRe_def hIm_def iii_def hcnj_def hsgn_def harg_def hcis_def
   hrcis_def hExp_def HComplex_def
 
 lemma Standard_hRe [simp]: "x \<in> Standard \<Longrightarrow> hRe x \<in> Standard"
-by (simp add: hcomplex_defs)
+  by (simp add: hcomplex_defs)
 
 lemma Standard_hIm [simp]: "x \<in> Standard \<Longrightarrow> hIm x \<in> Standard"
-by (simp add: hcomplex_defs)
+  by (simp add: hcomplex_defs)
 
 lemma Standard_iii [simp]: "iii \<in> Standard"
-by (simp add: hcomplex_defs)
+  by (simp add: hcomplex_defs)
 
 lemma Standard_hcnj [simp]: "x \<in> Standard \<Longrightarrow> hcnj x \<in> Standard"
-by (simp add: hcomplex_defs)
+  by (simp add: hcomplex_defs)
 
 lemma Standard_hsgn [simp]: "x \<in> Standard \<Longrightarrow> hsgn x \<in> Standard"
-by (simp add: hcomplex_defs)
+  by (simp add: hcomplex_defs)
 
 lemma Standard_harg [simp]: "x \<in> Standard \<Longrightarrow> harg x \<in> Standard"
-by (simp add: hcomplex_defs)
+  by (simp add: hcomplex_defs)
 
 lemma Standard_hcis [simp]: "r \<in> Standard \<Longrightarrow> hcis r \<in> Standard"
-by (simp add: hcomplex_defs)
+  by (simp add: hcomplex_defs)
 
 lemma Standard_hExp [simp]: "x \<in> Standard \<Longrightarrow> hExp x \<in> Standard"
-by (simp add: hcomplex_defs)
+  by (simp add: hcomplex_defs)
 
-lemma Standard_hrcis [simp]:
-  "\<lbrakk>r \<in> Standard; s \<in> Standard\<rbrakk> \<Longrightarrow> hrcis r s \<in> Standard"
-by (simp add: hcomplex_defs)
+lemma Standard_hrcis [simp]: "r \<in> Standard \<Longrightarrow> s \<in> Standard \<Longrightarrow> hrcis r s \<in> Standard"
+  by (simp add: hcomplex_defs)
 
-lemma Standard_HComplex [simp]:
-  "\<lbrakk>r \<in> Standard; s \<in> Standard\<rbrakk> \<Longrightarrow> HComplex r s \<in> Standard"
-by (simp add: hcomplex_defs)
+lemma Standard_HComplex [simp]: "r \<in> Standard \<Longrightarrow> s \<in> Standard \<Longrightarrow> HComplex r s \<in> Standard"
+  by (simp add: hcomplex_defs)
 
 lemma hcmod_def: "hcmod = *f* cmod"
-by (rule hnorm_def)
+  by (rule hnorm_def)
 
 
-subsection\<open>Properties of Nonstandard Real and Imaginary Parts\<close>
+subsection \<open>Properties of Nonstandard Real and Imaginary Parts\<close>
 
-lemma hcomplex_hRe_hIm_cancel_iff:
-     "!!w z. (w=z) = (hRe(w) = hRe(z) & hIm(w) = hIm(z))"
-by transfer (rule complex_Re_Im_cancel_iff)
+lemma hcomplex_hRe_hIm_cancel_iff: "\<And>w z. w = z \<longleftrightarrow> hRe w = hRe z \<and> hIm w = hIm z"
+  by transfer (rule complex_Re_Im_cancel_iff)
 
-lemma hcomplex_equality [intro?]:
-  "!!z w. hRe z = hRe w ==> hIm z = hIm w ==> z = w"
-by transfer (rule complex_equality)
+lemma hcomplex_equality [intro?]: "\<And>z w. hRe z = hRe w \<Longrightarrow> hIm z = hIm w \<Longrightarrow> z = w"
+  by transfer (rule complex_equality)
 
 lemma hcomplex_hRe_zero [simp]: "hRe 0 = 0"
-by transfer simp
+  by transfer simp
 
 lemma hcomplex_hIm_zero [simp]: "hIm 0 = 0"
-by transfer simp
+  by transfer simp
 
 lemma hcomplex_hRe_one [simp]: "hRe 1 = 1"
-by transfer simp
+  by transfer simp
 
 lemma hcomplex_hIm_one [simp]: "hIm 1 = 0"
-by transfer simp
+  by transfer simp
+
+
+subsection \<open>Addition for Nonstandard Complex Numbers\<close>
+
+lemma hRe_add: "\<And>x y. hRe (x + y) = hRe x + hRe y"
+  by transfer simp
+
+lemma hIm_add: "\<And>x y. hIm (x + y) = hIm x + hIm y"
+  by transfer simp
 
 
-subsection\<open>Addition for Nonstandard Complex Numbers\<close>
-
-lemma hRe_add: "!!x y. hRe(x + y) = hRe(x) + hRe(y)"
-by transfer simp
+subsection \<open>More Minus Laws\<close>
 
-lemma hIm_add: "!!x y. hIm(x + y) = hIm(x) + hIm(y)"
-by transfer simp
+lemma hRe_minus: "\<And>z. hRe (- z) = - hRe z"
+  by transfer (rule uminus_complex.sel)
 
-subsection\<open>More Minus Laws\<close>
-
-lemma hRe_minus: "!!z. hRe(-z) = - hRe(z)"
-by transfer (rule uminus_complex.sel)
+lemma hIm_minus: "\<And>z. hIm (- z) = - hIm z"
+  by transfer (rule uminus_complex.sel)
 
-lemma hIm_minus: "!!z. hIm(-z) = - hIm(z)"
-by transfer (rule uminus_complex.sel)
+lemma hcomplex_add_minus_eq_minus: "x + y = 0 \<Longrightarrow> x = - y"
+  for x y :: hcomplex
+  apply (drule minus_unique)
+  apply (simp add: minus_equation_iff [of x y])
+  done
 
-lemma hcomplex_add_minus_eq_minus:
-      "x + y = (0::hcomplex) ==> x = -y"
-apply (drule minus_unique)
-apply (simp add: minus_equation_iff [of x y])
-done
+lemma hcomplex_i_mult_eq [simp]: "iii * iii = - 1"
+  by transfer (rule i_squared)
 
-lemma hcomplex_i_mult_eq [simp]: "iii * iii = -1"
-by transfer (rule i_squared)
-
-lemma hcomplex_i_mult_left [simp]: "!!z. iii * (iii * z) = -z"
-by transfer (rule complex_i_mult_minus)
+lemma hcomplex_i_mult_left [simp]: "\<And>z. iii * (iii * z) = - z"
+  by transfer (rule complex_i_mult_minus)
 
 lemma hcomplex_i_not_zero [simp]: "iii \<noteq> 0"
-by transfer (rule complex_i_not_zero)
+  by transfer (rule complex_i_not_zero)
 
 
-subsection\<open>More Multiplication Laws\<close>
+subsection \<open>More Multiplication Laws\<close>
 
-lemma hcomplex_mult_minus_one: "- 1 * (z::hcomplex) = -z"
-by simp
-
-lemma hcomplex_mult_minus_one_right: "(z::hcomplex) * - 1 = -z"
-by simp
+lemma hcomplex_mult_minus_one: "- 1 * z = - z"
+  for z :: hcomplex
+  by simp
 
-lemma hcomplex_mult_left_cancel:
-     "(c::hcomplex) \<noteq> (0::hcomplex) ==> (c*a=c*b) = (a=b)"
-by simp
+lemma hcomplex_mult_minus_one_right: "z * - 1 = - z"
+  for z :: hcomplex
+  by simp
 
-lemma hcomplex_mult_right_cancel:
-     "(c::hcomplex) \<noteq> (0::hcomplex) ==> (a*c=b*c) = (a=b)"
-by simp
+lemma hcomplex_mult_left_cancel: "c \<noteq> 0 \<Longrightarrow> c * a = c * b \<longleftrightarrow> a = b"
+  for a b c :: hcomplex
+  by simp
+
+lemma hcomplex_mult_right_cancel: "c \<noteq> 0 \<Longrightarrow> a * c = b * c \<longleftrightarrow> a = b"
+  for a b c :: hcomplex
+  by simp
 
 
-subsection\<open>Subraction and Division\<close>
+subsection \<open>Subtraction and Division\<close>
 
-lemma hcomplex_diff_eq_eq [simp]: "((x::hcomplex) - y = z) = (x = z + y)"
 (* TODO: delete *)
-by (rule diff_eq_eq)
+lemma hcomplex_diff_eq_eq [simp]: "x - y = z \<longleftrightarrow> x = z + y"
+  for x y z :: hcomplex
+  by (rule diff_eq_eq)
 
 
-subsection\<open>Embedding Properties for @{term hcomplex_of_hypreal} Map\<close>
+subsection \<open>Embedding Properties for @{term hcomplex_of_hypreal} Map\<close>
+
+lemma hRe_hcomplex_of_hypreal [simp]: "\<And>z. hRe (hcomplex_of_hypreal z) = z"
+  by transfer (rule Re_complex_of_real)
 
-lemma hRe_hcomplex_of_hypreal [simp]: "!!z. hRe(hcomplex_of_hypreal z) = z"
-by transfer (rule Re_complex_of_real)
+lemma hIm_hcomplex_of_hypreal [simp]: "\<And>z. hIm (hcomplex_of_hypreal z) = 0"
+  by transfer (rule Im_complex_of_real)
 
-lemma hIm_hcomplex_of_hypreal [simp]: "!!z. hIm(hcomplex_of_hypreal z) = 0"
-by transfer (rule Im_complex_of_real)
+lemma hcomplex_of_hypreal_epsilon_not_zero [simp]: "hcomplex_of_hypreal \<epsilon> \<noteq> 0"
+  by (simp add: hypreal_epsilon_not_zero)
 
-lemma hcomplex_of_hypreal_epsilon_not_zero [simp]:
-     "hcomplex_of_hypreal \<epsilon> \<noteq> 0"
-by (simp add: hypreal_epsilon_not_zero)
+
+subsection \<open>\<open>HComplex\<close> theorems\<close>
 
-subsection\<open>HComplex theorems\<close>
+lemma hRe_HComplex [simp]: "\<And>x y. hRe (HComplex x y) = x"
+  by transfer simp
 
-lemma hRe_HComplex [simp]: "!!x y. hRe (HComplex x y) = x"
-by transfer simp
+lemma hIm_HComplex [simp]: "\<And>x y. hIm (HComplex x y) = y"
+  by transfer simp
 
-lemma hIm_HComplex [simp]: "!!x y. hIm (HComplex x y) = y"
-by transfer simp
-
-lemma hcomplex_surj [simp]: "!!z. HComplex (hRe z) (hIm z) = z"
-by transfer (rule complex_surj)
+lemma hcomplex_surj [simp]: "\<And>z. HComplex (hRe z) (hIm z) = z"
+  by transfer (rule complex_surj)
 
 lemma hcomplex_induct [case_names rect(*, induct type: hcomplex*)]:
-     "(\<And>x y. P (HComplex x y)) ==> P z"
-by (rule hcomplex_surj [THEN subst], blast)
+  "(\<And>x y. P (HComplex x y)) \<Longrightarrow> P z"
+  by (rule hcomplex_surj [THEN subst]) blast
 
 
-subsection\<open>Modulus (Absolute Value) of Nonstandard Complex Number\<close>
+subsection \<open>Modulus (Absolute Value) of Nonstandard Complex Number\<close>
 
 lemma hcomplex_of_hypreal_abs:
-     "hcomplex_of_hypreal \<bar>x\<bar> =
-      hcomplex_of_hypreal (hcmod (hcomplex_of_hypreal x))"
-by simp
+  "hcomplex_of_hypreal \<bar>x\<bar> = hcomplex_of_hypreal (hcmod (hcomplex_of_hypreal x))"
+  by simp
 
-lemma HComplex_inject [simp]:
-  "!!x y x' y'. HComplex x y = HComplex x' y' = (x=x' & y=y')"
-by transfer (rule complex.inject)
+lemma HComplex_inject [simp]: "\<And>x y x' y'. HComplex x y = HComplex x' y' \<longleftrightarrow> x = x' \<and> y = y'"
+  by transfer (rule complex.inject)
 
 lemma HComplex_add [simp]:
-  "!!x1 y1 x2 y2. HComplex x1 y1 + HComplex x2 y2 = HComplex (x1+x2) (y1+y2)"
-by transfer (rule complex_add)
+  "\<And>x1 y1 x2 y2. HComplex x1 y1 + HComplex x2 y2 = HComplex (x1 + x2) (y1 + y2)"
+  by transfer (rule complex_add)
 
-lemma HComplex_minus [simp]: "!!x y. - HComplex x y = HComplex (-x) (-y)"
-by transfer (rule complex_minus)
+lemma HComplex_minus [simp]: "\<And>x y. - HComplex x y = HComplex (- x) (- y)"
+  by transfer (rule complex_minus)
 
 lemma HComplex_diff [simp]:
-  "!!x1 y1 x2 y2. HComplex x1 y1 - HComplex x2 y2 = HComplex (x1-x2) (y1-y2)"
-by transfer (rule complex_diff)
+  "\<And>x1 y1 x2 y2. HComplex x1 y1 - HComplex x2 y2 = HComplex (x1 - x2) (y1 - y2)"
+  by transfer (rule complex_diff)
 
 lemma HComplex_mult [simp]:
-  "!!x1 y1 x2 y2. HComplex x1 y1 * HComplex x2 y2 =
-   HComplex (x1*x2 - y1*y2) (x1*y2 + y1*x2)"
-by transfer (rule complex_mult)
+  "\<And>x1 y1 x2 y2. HComplex x1 y1 * HComplex x2 y2 = HComplex (x1*x2 - y1*y2) (x1*y2 + y1*x2)"
+  by transfer (rule complex_mult)
 
-(*HComplex_inverse is proved below*)
+text \<open>\<open>HComplex_inverse\<close> is proved below.\<close>
 
-lemma hcomplex_of_hypreal_eq: "!!r. hcomplex_of_hypreal r = HComplex r 0"
-by transfer (rule complex_of_real_def)
+lemma hcomplex_of_hypreal_eq: "\<And>r. hcomplex_of_hypreal r = HComplex r 0"
+  by transfer (rule complex_of_real_def)
 
 lemma HComplex_add_hcomplex_of_hypreal [simp]:
-     "!!x y r. HComplex x y + hcomplex_of_hypreal r = HComplex (x+r) y"
-by transfer (rule Complex_add_complex_of_real)
+  "\<And>x y r. HComplex x y + hcomplex_of_hypreal r = HComplex (x + r) y"
+  by transfer (rule Complex_add_complex_of_real)
 
 lemma hcomplex_of_hypreal_add_HComplex [simp]:
-     "!!r x y. hcomplex_of_hypreal r + HComplex x y = HComplex (r+x) y"
-by transfer (rule complex_of_real_add_Complex)
+  "\<And>r x y. hcomplex_of_hypreal r + HComplex x y = HComplex (r + x) y"
+  by transfer (rule complex_of_real_add_Complex)
 
 lemma HComplex_mult_hcomplex_of_hypreal:
-     "!!x y r. HComplex x y * hcomplex_of_hypreal r = HComplex (x*r) (y*r)"
-by transfer (rule Complex_mult_complex_of_real)
+  "\<And>x y r. HComplex x y * hcomplex_of_hypreal r = HComplex (x * r) (y * r)"
+  by transfer (rule Complex_mult_complex_of_real)
 
 lemma hcomplex_of_hypreal_mult_HComplex:
-     "!!r x y. hcomplex_of_hypreal r * HComplex x y = HComplex (r*x) (r*y)"
-by transfer (rule complex_of_real_mult_Complex)
+  "\<And>r x y. hcomplex_of_hypreal r * HComplex x y = HComplex (r * x) (r * y)"
+  by transfer (rule complex_of_real_mult_Complex)
 
-lemma i_hcomplex_of_hypreal [simp]:
-     "!!r. iii * hcomplex_of_hypreal r = HComplex 0 r"
-by transfer (rule i_complex_of_real)
+lemma i_hcomplex_of_hypreal [simp]: "\<And>r. iii * hcomplex_of_hypreal r = HComplex 0 r"
+  by transfer (rule i_complex_of_real)
 
-lemma hcomplex_of_hypreal_i [simp]:
-     "!!r. hcomplex_of_hypreal r * iii = HComplex 0 r"
-by transfer (rule complex_of_real_i)
+lemma hcomplex_of_hypreal_i [simp]: "\<And>r. hcomplex_of_hypreal r * iii = HComplex 0 r"
+  by transfer (rule complex_of_real_i)
 
 
-subsection\<open>Conjugation\<close>
+subsection \<open>Conjugation\<close>
 
-lemma hcomplex_hcnj_cancel_iff [iff]: "!!x y. (hcnj x = hcnj y) = (x = y)"
-by transfer (rule complex_cnj_cancel_iff)
+lemma hcomplex_hcnj_cancel_iff [iff]: "\<And>x y. hcnj x = hcnj y \<longleftrightarrow> x = y"
+  by transfer (rule complex_cnj_cancel_iff)
 
-lemma hcomplex_hcnj_hcnj [simp]: "!!z. hcnj (hcnj z) = z"
-by transfer (rule complex_cnj_cnj)
+lemma hcomplex_hcnj_hcnj [simp]: "\<And>z. hcnj (hcnj z) = z"
+  by transfer (rule complex_cnj_cnj)
 
 lemma hcomplex_hcnj_hcomplex_of_hypreal [simp]:
-     "!!x. hcnj (hcomplex_of_hypreal x) = hcomplex_of_hypreal x"
-by transfer (rule complex_cnj_complex_of_real)
+  "\<And>x. hcnj (hcomplex_of_hypreal x) = hcomplex_of_hypreal x"
+  by transfer (rule complex_cnj_complex_of_real)
 
-lemma hcomplex_hmod_hcnj [simp]: "!!z. hcmod (hcnj z) = hcmod z"
-by transfer (rule complex_mod_cnj)
+lemma hcomplex_hmod_hcnj [simp]: "\<And>z. hcmod (hcnj z) = hcmod z"
+  by transfer (rule complex_mod_cnj)
 
-lemma hcomplex_hcnj_minus: "!!z. hcnj (-z) = - hcnj z"
-by transfer (rule complex_cnj_minus)
+lemma hcomplex_hcnj_minus: "\<And>z. hcnj (- z) = - hcnj z"
+  by transfer (rule complex_cnj_minus)
 
-lemma hcomplex_hcnj_inverse: "!!z. hcnj(inverse z) = inverse(hcnj z)"
-by transfer (rule complex_cnj_inverse)
+lemma hcomplex_hcnj_inverse: "\<And>z. hcnj (inverse z) = inverse (hcnj z)"
+  by transfer (rule complex_cnj_inverse)
 
-lemma hcomplex_hcnj_add: "!!w z. hcnj(w + z) = hcnj(w) + hcnj(z)"
-by transfer (rule complex_cnj_add)
+lemma hcomplex_hcnj_add: "\<And>w z. hcnj (w + z) = hcnj w + hcnj z"
+  by transfer (rule complex_cnj_add)
 
-lemma hcomplex_hcnj_diff: "!!w z. hcnj(w - z) = hcnj(w) - hcnj(z)"
-by transfer (rule complex_cnj_diff)
+lemma hcomplex_hcnj_diff: "\<And>w z. hcnj (w - z) = hcnj w - hcnj z"
+  by transfer (rule complex_cnj_diff)
 
-lemma hcomplex_hcnj_mult: "!!w z. hcnj(w * z) = hcnj(w) * hcnj(z)"
-by transfer (rule complex_cnj_mult)
+lemma hcomplex_hcnj_mult: "\<And>w z. hcnj (w * z) = hcnj w * hcnj z"
+  by transfer (rule complex_cnj_mult)
 
-lemma hcomplex_hcnj_divide: "!!w z. hcnj(w / z) = (hcnj w)/(hcnj z)"
-by transfer (rule complex_cnj_divide)
+lemma hcomplex_hcnj_divide: "\<And>w z. hcnj (w / z) = hcnj w / hcnj z"
+  by transfer (rule complex_cnj_divide)
 
 lemma hcnj_one [simp]: "hcnj 1 = 1"
-by transfer (rule complex_cnj_one)
+  by transfer (rule complex_cnj_one)
 
 lemma hcomplex_hcnj_zero [simp]: "hcnj 0 = 0"
-by transfer (rule complex_cnj_zero)
-
-lemma hcomplex_hcnj_zero_iff [iff]: "!!z. (hcnj z = 0) = (z = 0)"
-by transfer (rule complex_cnj_zero_iff)
-
-lemma hcomplex_mult_hcnj:
-     "!!z. z * hcnj z = hcomplex_of_hypreal ((hRe z)\<^sup>2 + (hIm z)\<^sup>2)"
-by transfer (rule complex_mult_cnj)
-
-
-subsection\<open>More Theorems about the Function @{term hcmod}\<close>
+  by transfer (rule complex_cnj_zero)
 
-lemma hcmod_hcomplex_of_hypreal_of_nat [simp]:
-     "hcmod (hcomplex_of_hypreal(hypreal_of_nat n)) = hypreal_of_nat n"
-by simp
-
-lemma hcmod_hcomplex_of_hypreal_of_hypnat [simp]:
-     "hcmod (hcomplex_of_hypreal(hypreal_of_hypnat n)) = hypreal_of_hypnat n"
-by simp
+lemma hcomplex_hcnj_zero_iff [iff]: "\<And>z. hcnj z = 0 \<longleftrightarrow> z = 0"
+  by transfer (rule complex_cnj_zero_iff)
 
-lemma hcmod_mult_hcnj: "!!z. hcmod(z * hcnj(z)) = (hcmod z)\<^sup>2"
-by transfer (rule complex_mod_mult_cnj)
-
-lemma hcmod_triangle_ineq2 [simp]:
-  "!!a b. hcmod(b + a) - hcmod b \<le> hcmod a"
-by transfer (rule complex_mod_triangle_ineq2)
-
-lemma hcmod_diff_ineq [simp]: "!!a b. hcmod(a) - hcmod(b) \<le> hcmod(a + b)"
-by transfer (rule norm_diff_ineq)
+lemma hcomplex_mult_hcnj: "\<And>z. z * hcnj z = hcomplex_of_hypreal ((hRe z)\<^sup>2 + (hIm z)\<^sup>2)"
+  by transfer (rule complex_mult_cnj)
 
 
-subsection\<open>Exponentiation\<close>
+subsection \<open>More Theorems about the Function @{term hcmod}\<close>
+
+lemma hcmod_hcomplex_of_hypreal_of_nat [simp]:
+  "hcmod (hcomplex_of_hypreal (hypreal_of_nat n)) = hypreal_of_nat n"
+  by simp
+
+lemma hcmod_hcomplex_of_hypreal_of_hypnat [simp]:
+  "hcmod (hcomplex_of_hypreal(hypreal_of_hypnat n)) = hypreal_of_hypnat n"
+  by simp
+
+lemma hcmod_mult_hcnj: "\<And>z. hcmod (z * hcnj z) = (hcmod z)\<^sup>2"
+  by transfer (rule complex_mod_mult_cnj)
 
-lemma hcomplexpow_0 [simp]:   "z ^ 0       = (1::hcomplex)"
-by (rule power_0)
+lemma hcmod_triangle_ineq2 [simp]: "\<And>a b. hcmod (b + a) - hcmod b \<le> hcmod a"
+  by transfer (rule complex_mod_triangle_ineq2)
+
+lemma hcmod_diff_ineq [simp]: "\<And>a b. hcmod a - hcmod b \<le> hcmod (a + b)"
+  by transfer (rule norm_diff_ineq)
+
 
-lemma hcomplexpow_Suc [simp]: "z ^ (Suc n) = (z::hcomplex) * (z ^ n)"
-by (rule power_Suc)
+subsection \<open>Exponentiation\<close>
+
+lemma hcomplexpow_0 [simp]: "z ^ 0 = 1"
+  for z :: hcomplex
+  by (rule power_0)
+
+lemma hcomplexpow_Suc [simp]: "z ^ (Suc n) = z * (z ^ n)"
+  for z :: hcomplex
+  by (rule power_Suc)
 
 lemma hcomplexpow_i_squared [simp]: "iii\<^sup>2 = -1"
-by transfer (rule power2_i)
+  by transfer (rule power2_i)
 
-lemma hcomplex_of_hypreal_pow:
-     "!!x. hcomplex_of_hypreal (x ^ n) = (hcomplex_of_hypreal x) ^ n"
-by transfer (rule of_real_power)
+lemma hcomplex_of_hypreal_pow: "\<And>x. hcomplex_of_hypreal (x ^ n) = hcomplex_of_hypreal x ^ n"
+  by transfer (rule of_real_power)
 
-lemma hcomplex_hcnj_pow: "!!z. hcnj(z ^ n) = hcnj(z) ^ n"
-by transfer (rule complex_cnj_power)
+lemma hcomplex_hcnj_pow: "\<And>z. hcnj (z ^ n) = hcnj z ^ n"
+  by transfer (rule complex_cnj_power)
 
-lemma hcmod_hcomplexpow: "!!x. hcmod(x ^ n) = hcmod(x) ^ n"
-by transfer (rule norm_power)
+lemma hcmod_hcomplexpow: "\<And>x. hcmod (x ^ n) = hcmod x ^ n"
+  by transfer (rule norm_power)
 
 lemma hcpow_minus:
-     "!!x n. (-x::hcomplex) pow n =
-      (if ( *p* even) n then (x pow n) else -(x pow n))"
-by transfer simp
+  "\<And>x n. (- x :: hcomplex) pow n = (if ( *p* even) n then (x pow n) else - (x pow n))"
+  by transfer simp
 
-lemma hcpow_mult:
-  "((r::hcomplex) * s) pow n = (r pow n) * (s pow n)"
+lemma hcpow_mult: "(r * s) pow n = (r pow n) * (s pow n)"
+  for r s :: hcomplex
   by (fact hyperpow_mult)
 
-lemma hcpow_zero2 [simp]:
-  "\<And>n. 0 pow (hSuc n) = (0::'a::semiring_1 star)"
+lemma hcpow_zero2 [simp]: "\<And>n. 0 pow (hSuc n) = (0::'a::semiring_1 star)"
   by transfer (rule power_0_Suc)
 
-lemma hcpow_not_zero [simp,intro]:
-  "!!r n. r \<noteq> 0 ==> r pow n \<noteq> (0::hcomplex)"
+lemma hcpow_not_zero [simp,intro]: "\<And>r n. r \<noteq> 0 \<Longrightarrow> r pow n \<noteq> (0::hcomplex)"
   by (fact hyperpow_not_zero)
 
-lemma hcpow_zero_zero:
-  "r pow n = (0::hcomplex) ==> r = 0"
+lemma hcpow_zero_zero: "r pow n = 0 \<Longrightarrow> r = 0"
+  for r :: hcomplex
   by (blast intro: ccontr dest: hcpow_not_zero)
 
-subsection\<open>The Function @{term hsgn}\<close>
+
+subsection \<open>The Function @{term hsgn}\<close>
 
 lemma hsgn_zero [simp]: "hsgn 0 = 0"
-by transfer (rule sgn_zero)
+  by transfer (rule sgn_zero)
 
 lemma hsgn_one [simp]: "hsgn 1 = 1"
-by transfer (rule sgn_one)
+  by transfer (rule sgn_one)
 
-lemma hsgn_minus: "!!z. hsgn (-z) = - hsgn(z)"
-by transfer (rule sgn_minus)
+lemma hsgn_minus: "\<And>z. hsgn (- z) = - hsgn z"
+  by transfer (rule sgn_minus)
 
-lemma hsgn_eq: "!!z. hsgn z = z / hcomplex_of_hypreal (hcmod z)"
-by transfer (rule sgn_eq)
+lemma hsgn_eq: "\<And>z. hsgn z = z / hcomplex_of_hypreal (hcmod z)"
+  by transfer (rule sgn_eq)
 
-lemma hcmod_i: "!!x y. hcmod (HComplex x y) = ( *f* sqrt) (x\<^sup>2 + y\<^sup>2)"
-by transfer (rule complex_norm)
+lemma hcmod_i: "\<And>x y. hcmod (HComplex x y) = ( *f* sqrt) (x\<^sup>2 + y\<^sup>2)"
+  by transfer (rule complex_norm)
 
 lemma hcomplex_eq_cancel_iff1 [simp]:
-     "(hcomplex_of_hypreal xa = HComplex x y) = (xa = x & y = 0)"
-by (simp add: hcomplex_of_hypreal_eq)
+  "hcomplex_of_hypreal xa = HComplex x y \<longleftrightarrow> xa = x \<and> y = 0"
+  by (simp add: hcomplex_of_hypreal_eq)
 
 lemma hcomplex_eq_cancel_iff2 [simp]:
-     "(HComplex x y = hcomplex_of_hypreal xa) = (x = xa & y = 0)"
-by (simp add: hcomplex_of_hypreal_eq)
+  "HComplex x y = hcomplex_of_hypreal xa \<longleftrightarrow> x = xa \<and> y = 0"
+  by (simp add: hcomplex_of_hypreal_eq)
 
-lemma HComplex_eq_0 [simp]: "!!x y. (HComplex x y = 0) = (x = 0 & y = 0)"
-by transfer (rule Complex_eq_0)
+lemma HComplex_eq_0 [simp]: "\<And>x y. HComplex x y = 0 \<longleftrightarrow> x = 0 \<and> y = 0"
+  by transfer (rule Complex_eq_0)
 
-lemma HComplex_eq_1 [simp]: "!!x y. (HComplex x y = 1) = (x = 1 & y = 0)"
-by transfer (rule Complex_eq_1)
+lemma HComplex_eq_1 [simp]: "\<And>x y. HComplex x y = 1 \<longleftrightarrow> x = 1 \<and> y = 0"
+  by transfer (rule Complex_eq_1)
 
 lemma i_eq_HComplex_0_1: "iii = HComplex 0 1"
-by transfer (simp add: complex_eq_iff)
+  by transfer (simp add: complex_eq_iff)
 
-lemma HComplex_eq_i [simp]: "!!x y. (HComplex x y = iii) = (x = 0 & y = 1)"
-by transfer (rule Complex_eq_i)
-
-lemma hRe_hsgn [simp]: "!!z. hRe(hsgn z) = hRe(z)/hcmod z"
-by transfer (rule Re_sgn)
+lemma HComplex_eq_i [simp]: "\<And>x y. HComplex x y = iii \<longleftrightarrow> x = 0 \<and> y = 1"
+  by transfer (rule Complex_eq_i)
 
-lemma hIm_hsgn [simp]: "!!z. hIm(hsgn z) = hIm(z)/hcmod z"
-by transfer (rule Im_sgn)
+lemma hRe_hsgn [simp]: "\<And>z. hRe (hsgn z) = hRe z / hcmod z"
+  by transfer (rule Re_sgn)
 
-lemma HComplex_inverse:
-     "!!x y. inverse (HComplex x y) = HComplex (x/(x\<^sup>2 + y\<^sup>2)) (-y/(x\<^sup>2 + y\<^sup>2))"
-by transfer (rule complex_inverse)
-
-lemma hRe_mult_i_eq[simp]:
-    "!!y. hRe (iii * hcomplex_of_hypreal y) = 0"
-by transfer simp
+lemma hIm_hsgn [simp]: "\<And>z. hIm (hsgn z) = hIm z / hcmod z"
+  by transfer (rule Im_sgn)
 
-lemma hIm_mult_i_eq [simp]:
-    "!!y. hIm (iii * hcomplex_of_hypreal y) = y"
-by transfer simp
+lemma HComplex_inverse: "\<And>x y. inverse (HComplex x y) = HComplex (x / (x\<^sup>2 + y\<^sup>2)) (- y / (x\<^sup>2 + y\<^sup>2))"
+  by transfer (rule complex_inverse)
 
-lemma hcmod_mult_i [simp]: "!!y. hcmod (iii * hcomplex_of_hypreal y) = \<bar>y\<bar>"
-by transfer (simp add: norm_complex_def)
-
-lemma hcmod_mult_i2 [simp]: "!!y. hcmod (hcomplex_of_hypreal y * iii) = \<bar>y\<bar>"
-by transfer (simp add: norm_complex_def)
+lemma hRe_mult_i_eq[simp]: "\<And>y. hRe (iii * hcomplex_of_hypreal y) = 0"
+  by transfer simp
 
-(*---------------------------------------------------------------------------*)
-(*  harg                                                                     *)
-(*---------------------------------------------------------------------------*)
+lemma hIm_mult_i_eq [simp]: "\<And>y. hIm (iii * hcomplex_of_hypreal y) = y"
+  by transfer simp
 
-lemma cos_harg_i_mult_zero [simp]:
-     "!!y. y \<noteq> 0 ==> ( *f* cos) (harg(HComplex 0 y)) = 0"
-by transfer simp
+lemma hcmod_mult_i [simp]: "\<And>y. hcmod (iii * hcomplex_of_hypreal y) = \<bar>y\<bar>"
+  by transfer (simp add: norm_complex_def)
 
-lemma hcomplex_of_hypreal_zero_iff [simp]:
-     "!!y. (hcomplex_of_hypreal y = 0) = (y = 0)"
-by transfer (rule of_real_eq_0_iff)
+lemma hcmod_mult_i2 [simp]: "\<And>y. hcmod (hcomplex_of_hypreal y * iii) = \<bar>y\<bar>"
+  by transfer (simp add: norm_complex_def)
 
 
-subsection\<open>Polar Form for Nonstandard Complex Numbers\<close>
+subsubsection \<open>\<open>harg\<close>\<close>
+
+lemma cos_harg_i_mult_zero [simp]: "\<And>y. y \<noteq> 0 \<Longrightarrow> ( *f* cos) (harg (HComplex 0 y)) = 0"
+  by transfer simp
 
-lemma complex_split_polar2:
-     "\<forall>n. \<exists>r a. (z n) =  complex_of_real r * (Complex (cos a) (sin a))"
-by (auto intro: complex_split_polar)
+lemma hcomplex_of_hypreal_zero_iff [simp]: "\<And>y. hcomplex_of_hypreal y = 0 \<longleftrightarrow> y = 0"
+  by transfer (rule of_real_eq_0_iff)
+
+
+subsection \<open>Polar Form for Nonstandard Complex Numbers\<close>
+
+lemma complex_split_polar2: "\<forall>n. \<exists>r a. (z n) = complex_of_real r * Complex (cos a) (sin a)"
+  by (auto intro: complex_split_polar)
 
 lemma hcomplex_split_polar:
-  "!!z. \<exists>r a. z = hcomplex_of_hypreal r * (HComplex(( *f* cos) a)(( *f* sin) a))"
-by transfer (simp add: complex_split_polar)
+  "\<And>z. \<exists>r a. z = hcomplex_of_hypreal r * (HComplex (( *f* cos) a) (( *f* sin) a))"
+  by transfer (simp add: complex_split_polar)
 
 lemma hcis_eq:
-   "!!a. hcis a =
-    (hcomplex_of_hypreal(( *f* cos) a) +
-    iii * hcomplex_of_hypreal(( *f* sin) a))"
-by transfer (simp add: complex_eq_iff)
+  "\<And>a. hcis a = hcomplex_of_hypreal (( *f* cos) a) + iii * hcomplex_of_hypreal (( *f* sin) a)"
+  by transfer (simp add: complex_eq_iff)
 
-lemma hrcis_Ex: "!!z. \<exists>r a. z = hrcis r a"
-by transfer (rule rcis_Ex)
+lemma hrcis_Ex: "\<And>z. \<exists>r a. z = hrcis r a"
+  by transfer (rule rcis_Ex)
 
 lemma hRe_hcomplex_polar [simp]:
-  "!!r a. hRe (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) =
-      r * ( *f* cos) a"
-by transfer simp
+  "\<And>r a. hRe (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = r * ( *f* cos) a"
+  by transfer simp
 
-lemma hRe_hrcis [simp]: "!!r a. hRe(hrcis r a) = r * ( *f* cos) a"
-by transfer (rule Re_rcis)
+lemma hRe_hrcis [simp]: "\<And>r a. hRe (hrcis r a) = r * ( *f* cos) a"
+  by transfer (rule Re_rcis)
 
 lemma hIm_hcomplex_polar [simp]:
-  "!!r a. hIm (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) =
-      r * ( *f* sin) a"
-by transfer simp
+  "\<And>r a. hIm (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = r * ( *f* sin) a"
+  by transfer simp
 
-lemma hIm_hrcis [simp]: "!!r a. hIm(hrcis r a) = r * ( *f* sin) a"
-by transfer (rule Im_rcis)
+lemma hIm_hrcis [simp]: "\<And>r a. hIm (hrcis r a) = r * ( *f* sin) a"
+  by transfer (rule Im_rcis)
 
-lemma hcmod_unit_one [simp]:
-     "!!a. hcmod (HComplex (( *f* cos) a) (( *f* sin) a)) = 1"
-by transfer (simp add: cmod_unit_one)
+lemma hcmod_unit_one [simp]: "\<And>a. hcmod (HComplex (( *f* cos) a) (( *f* sin) a)) = 1"
+  by transfer (simp add: cmod_unit_one)
 
 lemma hcmod_complex_polar [simp]:
-  "!!r a. hcmod (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = \<bar>r\<bar>"
-by transfer (simp add: cmod_complex_polar)
-
-lemma hcmod_hrcis [simp]: "!!r a. hcmod(hrcis r a) = \<bar>r\<bar>"
-by transfer (rule complex_mod_rcis)
+  "\<And>r a. hcmod (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = \<bar>r\<bar>"
+  by transfer (simp add: cmod_complex_polar)
 
-(*---------------------------------------------------------------------------*)
-(*  (r1 * hrcis a) * (r2 * hrcis b) = r1 * r2 * hrcis (a + b)                *)
-(*---------------------------------------------------------------------------*)
+lemma hcmod_hrcis [simp]: "\<And>r a. hcmod(hrcis r a) = \<bar>r\<bar>"
+  by transfer (rule complex_mod_rcis)
 
-lemma hcis_hrcis_eq: "!!a. hcis a = hrcis 1 a"
-by transfer (rule cis_rcis_eq)
+text \<open>\<open>(r1 * hrcis a) * (r2 * hrcis b) = r1 * r2 * hrcis (a + b)\<close>\<close>
+
+lemma hcis_hrcis_eq: "\<And>a. hcis a = hrcis 1 a"
+  by transfer (rule cis_rcis_eq)
 declare hcis_hrcis_eq [symmetric, simp]
 
-lemma hrcis_mult:
-  "!!a b r1 r2. hrcis r1 a * hrcis r2 b = hrcis (r1*r2) (a + b)"
-by transfer (rule rcis_mult)
+lemma hrcis_mult: "\<And>a b r1 r2. hrcis r1 a * hrcis r2 b = hrcis (r1 * r2) (a + b)"
+  by transfer (rule rcis_mult)
 
-lemma hcis_mult: "!!a b. hcis a * hcis b = hcis (a + b)"
-by transfer (rule cis_mult)
+lemma hcis_mult: "\<And>a b. hcis a * hcis b = hcis (a + b)"
+  by transfer (rule cis_mult)
 
 lemma hcis_zero [simp]: "hcis 0 = 1"
-by transfer (rule cis_zero)
+  by transfer (rule cis_zero)
 
-lemma hrcis_zero_mod [simp]: "!!a. hrcis 0 a = 0"
-by transfer (rule rcis_zero_mod)
+lemma hrcis_zero_mod [simp]: "\<And>a. hrcis 0 a = 0"
+  by transfer (rule rcis_zero_mod)
 
-lemma hrcis_zero_arg [simp]: "!!r. hrcis r 0 = hcomplex_of_hypreal r"
-by transfer (rule rcis_zero_arg)
+lemma hrcis_zero_arg [simp]: "\<And>r. hrcis r 0 = hcomplex_of_hypreal r"
+  by transfer (rule rcis_zero_arg)
 
-lemma hcomplex_i_mult_minus [simp]: "!!x. iii * (iii * x) = - x"
-by transfer (rule complex_i_mult_minus)
+lemma hcomplex_i_mult_minus [simp]: "\<And>x. iii * (iii * x) = - x"
+  by transfer (rule complex_i_mult_minus)
 
 lemma hcomplex_i_mult_minus2 [simp]: "iii * iii * x = - x"
-by simp
+  by simp
 
 lemma hcis_hypreal_of_nat_Suc_mult:
-   "!!a. hcis (hypreal_of_nat (Suc n) * a) =
-     hcis a * hcis (hypreal_of_nat n * a)"
-apply transfer
-apply (simp add: distrib_right cis_mult)
-done
+  "\<And>a. hcis (hypreal_of_nat (Suc n) * a) = hcis a * hcis (hypreal_of_nat n * a)"
+  by transfer (simp add: distrib_right cis_mult)
 
-lemma NSDeMoivre: "!!a. (hcis a) ^ n = hcis (hypreal_of_nat n * a)"
-apply transfer
-apply (rule DeMoivre)
-done
+lemma NSDeMoivre: "\<And>a. (hcis a) ^ n = hcis (hypreal_of_nat n * a)"
+  by transfer (rule DeMoivre)
 
 lemma hcis_hypreal_of_hypnat_Suc_mult:
-     "!! a n. hcis (hypreal_of_hypnat (n + 1) * a) =
-      hcis a * hcis (hypreal_of_hypnat n * a)"
-by transfer (simp add: distrib_right cis_mult)
+  "\<And>a n. hcis (hypreal_of_hypnat (n + 1) * a) = hcis a * hcis (hypreal_of_hypnat n * a)"
+  by transfer (simp add: distrib_right cis_mult)
 
-lemma NSDeMoivre_ext:
-  "!!a n. (hcis a) pow n = hcis (hypreal_of_hypnat n * a)"
-by transfer (rule DeMoivre)
+lemma NSDeMoivre_ext: "\<And>a n. (hcis a) pow n = hcis (hypreal_of_hypnat n * a)"
+  by transfer (rule DeMoivre)
+
+lemma NSDeMoivre2: "\<And>a r. (hrcis r a) ^ n = hrcis (r ^ n) (hypreal_of_nat n * a)"
+  by transfer (rule DeMoivre2)
 
-lemma NSDeMoivre2:
-  "!!a r. (hrcis r a) ^ n = hrcis (r ^ n) (hypreal_of_nat n * a)"
-by transfer (rule DeMoivre2)
+lemma DeMoivre2_ext: "\<And>a r n. (hrcis r a) pow n = hrcis (r pow n) (hypreal_of_hypnat n * a)"
+  by transfer (rule DeMoivre2)
 
-lemma DeMoivre2_ext:
-  "!! a r n. (hrcis r a) pow n = hrcis (r pow n) (hypreal_of_hypnat n * a)"
-by transfer (rule DeMoivre2)
+lemma hcis_inverse [simp]: "\<And>a. inverse (hcis a) = hcis (- a)"
+  by transfer (rule cis_inverse)
 
-lemma hcis_inverse [simp]: "!!a. inverse(hcis a) = hcis (-a)"
-by transfer (rule cis_inverse)
+lemma hrcis_inverse: "\<And>a r. inverse (hrcis r a) = hrcis (inverse r) (- a)"
+  by transfer (simp add: rcis_inverse inverse_eq_divide [symmetric])
 
-lemma hrcis_inverse: "!!a r. inverse(hrcis r a) = hrcis (inverse r) (-a)"
-by transfer (simp add: rcis_inverse inverse_eq_divide [symmetric])
-
-lemma hRe_hcis [simp]: "!!a. hRe(hcis a) = ( *f* cos) a"
-by transfer simp
+lemma hRe_hcis [simp]: "\<And>a. hRe (hcis a) = ( *f* cos) a"
+  by transfer simp
 
-lemma hIm_hcis [simp]: "!!a. hIm(hcis a) = ( *f* sin) a"
-by transfer simp
+lemma hIm_hcis [simp]: "\<And>a. hIm (hcis a) = ( *f* sin) a"
+  by transfer simp
 
-lemma cos_n_hRe_hcis_pow_n: "( *f* cos) (hypreal_of_nat n * a) = hRe(hcis a ^ n)"
-by (simp add: NSDeMoivre)
+lemma cos_n_hRe_hcis_pow_n: "( *f* cos) (hypreal_of_nat n * a) = hRe (hcis a ^ n)"
+  by (simp add: NSDeMoivre)
 
-lemma sin_n_hIm_hcis_pow_n: "( *f* sin) (hypreal_of_nat n * a) = hIm(hcis a ^ n)"
-by (simp add: NSDeMoivre)
+lemma sin_n_hIm_hcis_pow_n: "( *f* sin) (hypreal_of_nat n * a) = hIm (hcis a ^ n)"
+  by (simp add: NSDeMoivre)
 
-lemma cos_n_hRe_hcis_hcpow_n: "( *f* cos) (hypreal_of_hypnat n * a) = hRe(hcis a pow n)"
-by (simp add: NSDeMoivre_ext)
+lemma cos_n_hRe_hcis_hcpow_n: "( *f* cos) (hypreal_of_hypnat n * a) = hRe (hcis a pow n)"
+  by (simp add: NSDeMoivre_ext)
 
-lemma sin_n_hIm_hcis_hcpow_n: "( *f* sin) (hypreal_of_hypnat n * a) = hIm(hcis a pow n)"
-by (simp add: NSDeMoivre_ext)
+lemma sin_n_hIm_hcis_hcpow_n: "( *f* sin) (hypreal_of_hypnat n * a) = hIm (hcis a pow n)"
+  by (simp add: NSDeMoivre_ext)
 
-lemma hExp_add: "!!a b. hExp(a + b) = hExp(a) * hExp(b)"
-by transfer (rule exp_add)
+lemma hExp_add: "\<And>a b. hExp (a + b) = hExp a * hExp b"
+  by transfer (rule exp_add)
 
 
-subsection\<open>@{term hcomplex_of_complex}: the Injection from
-  type @{typ complex} to to @{typ hcomplex}\<close>
+subsection \<open>@{term hcomplex_of_complex}: the Injection from type @{typ complex} to to @{typ hcomplex}\<close>
 
 lemma hcomplex_of_complex_i: "iii = hcomplex_of_complex \<i>"
-by (rule iii_def)
+  by (rule iii_def)
 
-lemma hRe_hcomplex_of_complex:
-   "hRe (hcomplex_of_complex z) = hypreal_of_real (Re z)"
-by transfer (rule refl)
+lemma hRe_hcomplex_of_complex: "hRe (hcomplex_of_complex z) = hypreal_of_real (Re z)"
+  by transfer (rule refl)
 
-lemma hIm_hcomplex_of_complex:
-   "hIm (hcomplex_of_complex z) = hypreal_of_real (Im z)"
-by transfer (rule refl)
+lemma hIm_hcomplex_of_complex: "hIm (hcomplex_of_complex z) = hypreal_of_real (Im z)"
+  by transfer (rule refl)
 
-lemma hcmod_hcomplex_of_complex:
-     "hcmod (hcomplex_of_complex x) = hypreal_of_real (cmod x)"
-by transfer (rule refl)
+lemma hcmod_hcomplex_of_complex: "hcmod (hcomplex_of_complex x) = hypreal_of_real (cmod x)"
+  by transfer (rule refl)
 
 
-subsection\<open>Numerals and Arithmetic\<close>
+subsection \<open>Numerals and Arithmetic\<close>
 
 lemma hcomplex_of_hypreal_eq_hcomplex_of_complex:
-     "hcomplex_of_hypreal (hypreal_of_real x) =
-      hcomplex_of_complex (complex_of_real x)"
-by transfer (rule refl)
+  "hcomplex_of_hypreal (hypreal_of_real x) = hcomplex_of_complex (complex_of_real x)"
+  by transfer (rule refl)
 
 lemma hcomplex_hypreal_numeral:
   "hcomplex_of_complex (numeral w) = hcomplex_of_hypreal(numeral w)"
-by transfer (rule of_real_numeral [symmetric])
+  by transfer (rule of_real_numeral [symmetric])
 
 lemma hcomplex_hypreal_neg_numeral:
   "hcomplex_of_complex (- numeral w) = hcomplex_of_hypreal(- numeral w)"
-by transfer (rule of_real_neg_numeral [symmetric])
+  by transfer (rule of_real_neg_numeral [symmetric])
 
-lemma hcomplex_numeral_hcnj [simp]:
-     "hcnj (numeral v :: hcomplex) = numeral v"
-by transfer (rule complex_cnj_numeral)
+lemma hcomplex_numeral_hcnj [simp]: "hcnj (numeral v :: hcomplex) = numeral v"
+  by transfer (rule complex_cnj_numeral)
 
-lemma hcomplex_numeral_hcmod [simp]:
-      "hcmod(numeral v :: hcomplex) = (numeral v :: hypreal)"
-by transfer (rule norm_numeral)
+lemma hcomplex_numeral_hcmod [simp]: "hcmod (numeral v :: hcomplex) = (numeral v :: hypreal)"
+  by transfer (rule norm_numeral)
 
-lemma hcomplex_neg_numeral_hcmod [simp]:
-      "hcmod(- numeral v :: hcomplex) = (numeral v :: hypreal)"
-by transfer (rule norm_neg_numeral)
+lemma hcomplex_neg_numeral_hcmod [simp]: "hcmod (- numeral v :: hcomplex) = (numeral v :: hypreal)"
+  by transfer (rule norm_neg_numeral)
 
-lemma hcomplex_numeral_hRe [simp]:
-      "hRe(numeral v :: hcomplex) = numeral v"
-by transfer (rule complex_Re_numeral)
+lemma hcomplex_numeral_hRe [simp]: "hRe (numeral v :: hcomplex) = numeral v"
+  by transfer (rule complex_Re_numeral)
 
-lemma hcomplex_numeral_hIm [simp]:
-      "hIm(numeral v :: hcomplex) = 0"
-by transfer (rule complex_Im_numeral)
+lemma hcomplex_numeral_hIm [simp]: "hIm (numeral v :: hcomplex) = 0"
+  by transfer (rule complex_Im_numeral)
 
 end
--- a/src/HOL/Nonstandard_Analysis/NatStar.thy	Mon Oct 31 16:26:36 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/NatStar.thy	Tue Nov 01 00:44:24 2016 +0100
@@ -5,237 +5,198 @@
 Converted to Isar and polished by lcp
 *)
 
-section\<open>Star-transforms for the Hypernaturals\<close>
+section \<open>Star-transforms for the Hypernaturals\<close>
 
 theory NatStar
-imports Star
+  imports Star
 begin
 
 lemma star_n_eq_starfun_whn: "star_n X = ( *f* X) whn"
-by (simp add: hypnat_omega_def starfun_def star_of_def Ifun_star_n)
+  by (simp add: hypnat_omega_def starfun_def star_of_def Ifun_star_n)
 
-lemma starset_n_Un: "*sn* (%n. (A n) Un (B n)) = *sn* A Un *sn* B"
-apply (simp add: starset_n_def star_n_eq_starfun_whn Un_def)
-apply (rule_tac x=whn in spec, transfer, simp)
-done
+lemma starset_n_Un: "*sn* (\<lambda>n. (A n) \<union> (B n)) = *sn* A \<union> *sn* B"
+  apply (simp add: starset_n_def star_n_eq_starfun_whn Un_def)
+  apply (rule_tac x=whn in spec, transfer, simp)
+  done
 
-lemma InternalSets_Un:
-     "[| X \<in> InternalSets; Y \<in> InternalSets |]
-      ==> (X Un Y) \<in> InternalSets"
-by (auto simp add: InternalSets_def starset_n_Un [symmetric])
+lemma InternalSets_Un: "X \<in> InternalSets \<Longrightarrow> Y \<in> InternalSets \<Longrightarrow> X \<union> Y \<in> InternalSets"
+  by (auto simp add: InternalSets_def starset_n_Un [symmetric])
 
-lemma starset_n_Int:
-      "*sn* (%n. (A n) Int (B n)) = *sn* A Int *sn* B"
-apply (simp add: starset_n_def star_n_eq_starfun_whn Int_def)
-apply (rule_tac x=whn in spec, transfer, simp)
-done
+lemma starset_n_Int: "*sn* (\<lambda>n. A n \<inter> B n) = *sn* A \<inter> *sn* B"
+  apply (simp add: starset_n_def star_n_eq_starfun_whn Int_def)
+  apply (rule_tac x=whn in spec, transfer, simp)
+  done
 
-lemma InternalSets_Int:
-     "[| X \<in> InternalSets; Y \<in> InternalSets |]
-      ==> (X Int Y) \<in> InternalSets"
-by (auto simp add: InternalSets_def starset_n_Int [symmetric])
+lemma InternalSets_Int: "X \<in> InternalSets \<Longrightarrow> Y \<in> InternalSets \<Longrightarrow> X \<inter> Y \<in> InternalSets"
+  by (auto simp add: InternalSets_def starset_n_Int [symmetric])
 
-lemma starset_n_Compl: "*sn* ((%n. - A n)) = -( *sn* A)"
-apply (simp add: starset_n_def star_n_eq_starfun_whn Compl_eq)
-apply (rule_tac x=whn in spec, transfer, simp)
-done
+lemma starset_n_Compl: "*sn* ((\<lambda>n. - A n)) = - ( *sn* A)"
+  apply (simp add: starset_n_def star_n_eq_starfun_whn Compl_eq)
+  apply (rule_tac x=whn in spec, transfer, simp)
+  done
 
-lemma InternalSets_Compl: "X \<in> InternalSets ==> -X \<in> InternalSets"
-by (auto simp add: InternalSets_def starset_n_Compl [symmetric])
+lemma InternalSets_Compl: "X \<in> InternalSets \<Longrightarrow> - X \<in> InternalSets"
+  by (auto simp add: InternalSets_def starset_n_Compl [symmetric])
 
-lemma starset_n_diff: "*sn* (%n. (A n) - (B n)) = *sn* A - *sn* B"
-apply (simp add: starset_n_def star_n_eq_starfun_whn set_diff_eq)
-apply (rule_tac x=whn in spec, transfer, simp)
-done
+lemma starset_n_diff: "*sn* (\<lambda>n. (A n) - (B n)) = *sn* A - *sn* B"
+  apply (simp add: starset_n_def star_n_eq_starfun_whn set_diff_eq)
+  apply (rule_tac x=whn in spec, transfer, simp)
+  done
 
-lemma InternalSets_diff:
-     "[| X \<in> InternalSets; Y \<in> InternalSets |]
-      ==> (X - Y) \<in> InternalSets"
-by (auto simp add: InternalSets_def starset_n_diff [symmetric])
+lemma InternalSets_diff: "X \<in> InternalSets \<Longrightarrow> Y \<in> InternalSets \<Longrightarrow> X - Y \<in> InternalSets"
+  by (auto simp add: InternalSets_def starset_n_diff [symmetric])
 
 lemma NatStar_SHNat_subset: "Nats \<le> *s* (UNIV:: nat set)"
-by simp
+  by simp
 
-lemma NatStar_hypreal_of_real_Int:
-     "*s* X Int Nats = hypnat_of_nat ` X"
-by (auto simp add: SHNat_eq)
+lemma NatStar_hypreal_of_real_Int: "*s* X Int Nats = hypnat_of_nat ` X"
+  by (auto simp add: SHNat_eq)
 
-lemma starset_starset_n_eq: "*s* X = *sn* (%n. X)"
-by (simp add: starset_n_starset)
+lemma starset_starset_n_eq: "*s* X = *sn* (\<lambda>n. X)"
+  by (simp add: starset_n_starset)
 
 lemma InternalSets_starset_n [simp]: "( *s* X) \<in> InternalSets"
-by (auto simp add: InternalSets_def starset_starset_n_eq)
+  by (auto simp add: InternalSets_def starset_starset_n_eq)
 
-lemma InternalSets_UNIV_diff:
-     "X \<in> InternalSets ==> UNIV - X \<in> InternalSets"
-apply (subgoal_tac "UNIV - X = - X")
-by (auto intro: InternalSets_Compl)
+lemma InternalSets_UNIV_diff: "X \<in> InternalSets \<Longrightarrow> UNIV - X \<in> InternalSets"
+  apply (subgoal_tac "UNIV - X = - X")
+   apply (auto intro: InternalSets_Compl)
+  done
 
 
-subsection\<open>Nonstandard Extensions of Functions\<close>
-
-text\<open>Example of transfer of a property from reals to hyperreals
-    --- used for limit comparison of sequences\<close>
+subsection \<open>Nonstandard Extensions of Functions\<close>
 
-lemma starfun_le_mono:
-     "\<forall>n. N \<le> n --> f n \<le> g n
-      ==> \<forall>n. hypnat_of_nat N \<le> n --> ( *f* f) n \<le> ( *f* g) n"
-by transfer
+text \<open>Example of transfer of a property from reals to hyperreals
+  --- used for limit comparison of sequences.\<close>
+
+lemma starfun_le_mono: "\<forall>n. N \<le> n \<longrightarrow> f n \<le> g n \<Longrightarrow>
+  \<forall>n. hypnat_of_nat N \<le> n \<longrightarrow> ( *f* f) n \<le> ( *f* g) n"
+  by transfer
 
-(*****----- and another -----*****)
+text \<open>And another:\<close>
 lemma starfun_less_mono:
-     "\<forall>n. N \<le> n --> f n < g n
-      ==> \<forall>n. hypnat_of_nat N \<le> n --> ( *f* f) n < ( *f* g) n"
-by transfer
+  "\<forall>n. N \<le> n \<longrightarrow> f n < g n \<Longrightarrow> \<forall>n. hypnat_of_nat N \<le> n \<longrightarrow> ( *f* f) n < ( *f* g) n"
+  by transfer
 
-text\<open>Nonstandard extension when we increment the argument by one\<close>
+text \<open>Nonstandard extension when we increment the argument by one.\<close>
 
-lemma starfun_shift_one:
-     "!!N. ( *f* (%n. f (Suc n))) N = ( *f* f) (N + (1::hypnat))"
-by (transfer, simp)
+lemma starfun_shift_one: "\<And>N. ( *f* (\<lambda>n. f (Suc n))) N = ( *f* f) (N + (1::hypnat))"
+  by transfer simp
 
-text\<open>Nonstandard extension with absolute value\<close>
-
-lemma starfun_abs: "!!N. ( *f* (%n. \<bar>f n\<bar>)) N = \<bar>( *f* f) N\<bar>"
-by (transfer, rule refl)
+text \<open>Nonstandard extension with absolute value.\<close>
+lemma starfun_abs: "\<And>N. ( *f* (\<lambda>n. \<bar>f n\<bar>)) N = \<bar>( *f* f) N\<bar>"
+  by transfer (rule refl)
 
-text\<open>The hyperpow function as a nonstandard extension of realpow\<close>
-
-lemma starfun_pow: "!!N. ( *f* (%n. r ^ n)) N = (hypreal_of_real r) pow N"
-by (transfer, rule refl)
+text \<open>The \<open>hyperpow\<close> function as a nonstandard extension of \<open>realpow\<close>.\<close>
+lemma starfun_pow: "\<And>N. ( *f* (\<lambda>n. r ^ n)) N = hypreal_of_real r pow N"
+  by transfer (rule refl)
 
-lemma starfun_pow2:
-     "!!N. ( *f* (%n. (X n) ^ m)) N = ( *f* X) N pow hypnat_of_nat m"
-by (transfer, rule refl)
+lemma starfun_pow2: "\<And>N. ( *f* (\<lambda>n. X n ^ m)) N = ( *f* X) N pow hypnat_of_nat m"
+  by transfer (rule refl)
 
-lemma starfun_pow3: "!!R. ( *f* (%r. r ^ n)) R = (R) pow hypnat_of_nat n"
-by (transfer, rule refl)
+lemma starfun_pow3: "\<And>R. ( *f* (\<lambda>r. r ^ n)) R = R pow hypnat_of_nat n"
+  by transfer (rule refl)
 
-text\<open>The @{term hypreal_of_hypnat} function as a nonstandard extension of
-  @{term real_of_nat}\<close>
-
+text \<open>The @{term hypreal_of_hypnat} function as a nonstandard extension of
+  @{term real_of_nat}.\<close>
 lemma starfunNat_real_of_nat: "( *f* real) = hypreal_of_hypnat"
-by transfer (simp add: fun_eq_iff)
+  by transfer (simp add: fun_eq_iff)
 
 lemma starfun_inverse_real_of_nat_eq:
-     "N \<in> HNatInfinite
-   ==> ( *f* (%x::nat. inverse(real x))) N = inverse(hypreal_of_hypnat N)"
-apply (rule_tac f1 = inverse in starfun_o2 [THEN subst])
-apply (subgoal_tac "hypreal_of_hypnat N ~= 0")
-apply (simp_all add: zero_less_HNatInfinite starfunNat_real_of_nat starfun_inverse_inverse)
-done
+  "N \<in> HNatInfinite \<Longrightarrow> ( *f* (\<lambda>x::nat. inverse (real x))) N = inverse (hypreal_of_hypnat N)"
+  apply (rule_tac f1 = inverse in starfun_o2 [THEN subst])
+  apply (subgoal_tac "hypreal_of_hypnat N \<noteq> 0")
+   apply (simp_all add: zero_less_HNatInfinite starfunNat_real_of_nat)
+  done
 
-text\<open>Internal functions - some redundancy with *f* now\<close>
+text \<open>Internal functions -- some redundancy with \<open>*f*\<close> now.\<close>
 
-lemma starfun_n: "( *fn* f) (star_n X) = star_n (%n. f n (X n))"
-by (simp add: starfun_n_def Ifun_star_n)
-
-text\<open>Multiplication: \<open>( *fn) x ( *gn) = *(fn x gn)\<close>\<close>
+lemma starfun_n: "( *fn* f) (star_n X) = star_n (\<lambda>n. f n (X n))"
+  by (simp add: starfun_n_def Ifun_star_n)
 
-lemma starfun_n_mult:
-     "( *fn* f) z * ( *fn* g) z = ( *fn* (% i x. f i x * g i x)) z"
-apply (cases z)
-apply (simp add: starfun_n star_n_mult)
-done
+text \<open>Multiplication: \<open>( *fn) x ( *gn) = *(fn x gn)\<close>\<close>
 
-text\<open>Addition: \<open>( *fn) + ( *gn) = *(fn + gn)\<close>\<close>
+lemma starfun_n_mult: "( *fn* f) z * ( *fn* g) z = ( *fn* (\<lambda>i x. f i x * g i x)) z"
+  by (cases z) (simp add: starfun_n star_n_mult)
 
-lemma starfun_n_add:
-     "( *fn* f) z + ( *fn* g) z = ( *fn* (%i x. f i x + g i x)) z"
-apply (cases z)
-apply (simp add: starfun_n star_n_add)
-done
+text \<open>Addition: \<open>( *fn) + ( *gn) = *(fn + gn)\<close>\<close>
+lemma starfun_n_add: "( *fn* f) z + ( *fn* g) z = ( *fn* (\<lambda>i x. f i x + g i x)) z"
+  by (cases z) (simp add: starfun_n star_n_add)
 
-text\<open>Subtraction: \<open>( *fn) - ( *gn) = *(fn + - gn)\<close>\<close>
-
-lemma starfun_n_add_minus:
-     "( *fn* f) z + -( *fn* g) z = ( *fn* (%i x. f i x + -g i x)) z"
-apply (cases z)
-apply (simp add: starfun_n star_n_minus star_n_add)
-done
+text \<open>Subtraction: \<open>( *fn) - ( *gn) = *(fn + - gn)\<close>\<close>
+lemma starfun_n_add_minus: "( *fn* f) z + -( *fn* g) z = ( *fn* (\<lambda>i x. f i x + -g i x)) z"
+  by (cases z) (simp add: starfun_n star_n_minus star_n_add)
 
 
-text\<open>Composition: \<open>( *fn) o ( *gn) = *(fn o gn)\<close>\<close>
+text \<open>Composition: \<open>( *fn) \<circ> ( *gn) = *(fn \<circ> gn)\<close>\<close>
 
-lemma starfun_n_const_fun [simp]:
-     "( *fn* (%i x. k)) z = star_of k"
-apply (cases z)
-apply (simp add: starfun_n star_of_def)
-done
+lemma starfun_n_const_fun [simp]: "( *fn* (\<lambda>i x. k)) z = star_of k"
+  by (cases z) (simp add: starfun_n star_of_def)
 
-lemma starfun_n_minus: "- ( *fn* f) x = ( *fn* (%i x. - (f i) x)) x"
-apply (cases x)
-apply (simp add: starfun_n star_n_minus)
-done
+lemma starfun_n_minus: "- ( *fn* f) x = ( *fn* (\<lambda>i x. - (f i) x)) x"
+  by (cases x) (simp add: starfun_n star_n_minus)
 
-lemma starfun_n_eq [simp]:
-     "( *fn* f) (star_of n) = star_n (%i. f i n)"
-by (simp add: starfun_n star_of_def)
+lemma starfun_n_eq [simp]: "( *fn* f) (star_of n) = star_n (\<lambda>i. f i n)"
+  by (simp add: starfun_n star_of_def)
 
-lemma starfun_eq_iff: "(( *f* f) = ( *f* g)) = (f = g)"
-by (transfer, rule refl)
+lemma starfun_eq_iff: "(( *f* f) = ( *f* g)) \<longleftrightarrow> f = g"
+  by transfer (rule refl)
 
 lemma starfunNat_inverse_real_of_nat_Infinitesimal [simp]:
-     "N \<in> HNatInfinite ==> ( *f* (%x. inverse (real x))) N \<in> Infinitesimal"
-apply (rule_tac f1 = inverse in starfun_o2 [THEN subst])
-apply (subgoal_tac "hypreal_of_hypnat N ~= 0")
-apply (simp_all add: zero_less_HNatInfinite starfunNat_real_of_nat)
-done
+  "N \<in> HNatInfinite \<Longrightarrow> ( *f* (%x. inverse (real x))) N \<in> Infinitesimal"
+  apply (rule_tac f1 = inverse in starfun_o2 [THEN subst])
+  apply (subgoal_tac "hypreal_of_hypnat N ~= 0")
+   apply (simp_all add: zero_less_HNatInfinite starfunNat_real_of_nat)
+  done
 
 
-subsection\<open>Nonstandard Characterization of Induction\<close>
+subsection \<open>Nonstandard Characterization of Induction\<close>
 
 lemma hypnat_induct_obj:
-    "!!n. (( *p* P) (0::hypnat) &
-            (\<forall>n. ( *p* P)(n) --> ( *p* P)(n + 1)))
-       --> ( *p* P)(n)"
-by (transfer, induct_tac n, auto)
+  "\<And>n. (( *p* P) (0::hypnat) \<and> (\<forall>n. ( *p* P) n \<longrightarrow> ( *p* P) (n + 1))) \<longrightarrow> ( *p* P) n"
+  by transfer (induct_tac n, auto)
 
 lemma hypnat_induct:
-  "!!n. [| ( *p* P) (0::hypnat);
-      !!n. ( *p* P)(n) ==> ( *p* P)(n + 1)|]
-     ==> ( *p* P)(n)"
-by (transfer, induct_tac n, auto)
+  "\<And>n. ( *p* P) (0::hypnat) \<Longrightarrow> (\<And>n. ( *p* P) n \<Longrightarrow> ( *p* P) (n + 1)) \<Longrightarrow> ( *p* P) n"
+  by transfer (induct_tac n, auto)
 
 lemma starP2_eq_iff: "( *p2* (op =)) = (op =)"
-by transfer (rule refl)
+  by transfer (rule refl)
 
-lemma starP2_eq_iff2: "( *p2* (%x y. x = y)) X Y = (X = Y)"
-by (simp add: starP2_eq_iff)
+lemma starP2_eq_iff2: "( *p2* (\<lambda>x y. x = y)) X Y \<longleftrightarrow> X = Y"
+  by (simp add: starP2_eq_iff)
 
-lemma nonempty_nat_set_Least_mem:
-  "c \<in> (S :: nat set) ==> (LEAST n. n \<in> S) \<in> S"
-by (erule LeastI)
+lemma nonempty_nat_set_Least_mem: "c \<in> S \<Longrightarrow> (LEAST n. n \<in> S) \<in> S"
+  for S :: "nat set"
+  by (erule LeastI)
 
 lemma nonempty_set_star_has_least:
-    "!!S::nat set star. Iset S \<noteq> {} ==> \<exists>n \<in> Iset S. \<forall>m \<in> Iset S. n \<le> m"
-apply (transfer empty_def)
-apply (rule_tac x="LEAST n. n \<in> S" in bexI)
-apply (simp add: Least_le)
-apply (rule LeastI_ex, auto)
-done
+  "\<And>S::nat set star. Iset S \<noteq> {} \<Longrightarrow> \<exists>n \<in> Iset S. \<forall>m \<in> Iset S. n \<le> m"
+  apply (transfer empty_def)
+  apply (rule_tac x="LEAST n. n \<in> S" in bexI)
+   apply (simp add: Least_le)
+  apply (rule LeastI_ex, auto)
+  done
 
-lemma nonempty_InternalNatSet_has_least:
-    "[| (S::hypnat set) \<in> InternalSets; S \<noteq> {} |] ==> \<exists>n \<in> S. \<forall>m \<in> S. n \<le> m"
-apply (clarsimp simp add: InternalSets_def starset_n_def)
-apply (erule nonempty_set_star_has_least)
-done
+lemma nonempty_InternalNatSet_has_least: "S \<in> InternalSets \<Longrightarrow> S \<noteq> {} \<Longrightarrow> \<exists>n \<in> S. \<forall>m \<in> S. n \<le> m"
+  for S :: "hypnat set"
+  apply (clarsimp simp add: InternalSets_def starset_n_def)
+  apply (erule nonempty_set_star_has_least)
+  done
 
-text\<open>Goldblatt page 129 Thm 11.3.2\<close>
+text \<open>Goldblatt, page 129 Thm 11.3.2.\<close>
 lemma internal_induct_lemma:
-     "!!X::nat set star. [| (0::hypnat) \<in> Iset X; \<forall>n. n \<in> Iset X --> n + 1 \<in> Iset X |]
-      ==> Iset X = (UNIV:: hypnat set)"
-apply (transfer UNIV_def)
-apply (rule equalityI [OF subset_UNIV subsetI])
-apply (induct_tac x, auto)
-done
+  "\<And>X::nat set star.
+    (0::hypnat) \<in> Iset X \<Longrightarrow> \<forall>n. n \<in> Iset X \<longrightarrow> n + 1 \<in> Iset X \<Longrightarrow> Iset X = (UNIV:: hypnat set)"
+  apply (transfer UNIV_def)
+  apply (rule equalityI [OF subset_UNIV subsetI])
+  apply (induct_tac x, auto)
+  done
 
 lemma internal_induct:
-     "[| X \<in> InternalSets; (0::hypnat) \<in> X; \<forall>n. n \<in> X --> n + 1 \<in> X |]
-      ==> X = (UNIV:: hypnat set)"
-apply (clarsimp simp add: InternalSets_def starset_n_def)
-apply (erule (1) internal_induct_lemma)
-done
-
+  "X \<in> InternalSets \<Longrightarrow> (0::hypnat) \<in> X \<Longrightarrow> \<forall>n. n \<in> X \<longrightarrow> n + 1 \<in> X \<Longrightarrow> X = (UNIV:: hypnat set)"
+  apply (clarsimp simp add: InternalSets_def starset_n_def)
+  apply (erule (1) internal_induct_lemma)
+  done
 
 end
--- a/src/HOL/Nonstandard_Analysis/Star.thy	Mon Oct 31 16:26:36 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/Star.thy	Tue Nov 01 00:44:24 2016 +0100
@@ -4,340 +4,290 @@
     Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
 *)
 
-section\<open>Star-Transforms in Non-Standard Analysis\<close>
+section \<open>Star-Transforms in Non-Standard Analysis\<close>
 
 theory Star
-imports NSA
+  imports NSA
 begin
 
-definition
-    (* internal sets *)
-  starset_n :: "(nat => 'a set) => 'a star set" ("*sn* _" [80] 80) where
-  "*sn* As = Iset (star_n As)"
+definition  \<comment> \<open>internal sets\<close>
+  starset_n :: "(nat \<Rightarrow> 'a set) \<Rightarrow> 'a star set"  ("*sn* _" [80] 80)
+  where "*sn* As = Iset (star_n As)"
 
-definition
-  InternalSets :: "'a star set set" where
-  "InternalSets = {X. \<exists>As. X = *sn* As}"
+definition InternalSets :: "'a star set set"
+  where "InternalSets = {X. \<exists>As. X = *sn* As}"
 
-definition
-  (* nonstandard extension of function *)
-  is_starext  :: "['a star => 'a star, 'a => 'a] => bool" where
-  "is_starext F f =
-    (\<forall>x y. \<exists>X \<in> Rep_star(x). \<exists>Y \<in> Rep_star(y). ((y = (F x)) = (eventually (\<lambda>n. Y n = f(X n)) \<U>)))"
+definition  \<comment> \<open>nonstandard extension of function\<close>
+  is_starext :: "('a star \<Rightarrow> 'a star) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
+  where "is_starext F f \<longleftrightarrow>
+    (\<forall>x y. \<exists>X \<in> Rep_star x. \<exists>Y \<in> Rep_star y. y = F x \<longleftrightarrow> eventually (\<lambda>n. Y n = f(X n)) \<U>)"
 
-definition
-  (* internal functions *)
-  starfun_n :: "(nat => ('a => 'b)) => 'a star => 'b star"   ("*fn* _" [80] 80) where
-  "*fn* F = Ifun (star_n F)"
+definition  \<comment> \<open>internal functions\<close>
+  starfun_n :: "(nat \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a star \<Rightarrow> 'b star"  ("*fn* _" [80] 80)
+  where "*fn* F = Ifun (star_n F)"
 
-definition
-  InternalFuns :: "('a star => 'b star) set" where
-  "InternalFuns = {X. \<exists>F. X = *fn* F}"
+definition InternalFuns :: "('a star => 'b star) set"
+  where "InternalFuns = {X. \<exists>F. X = *fn* F}"
 
 
-(*--------------------------------------------------------
-   Preamble - Pulling "EX" over "ALL"
- ---------------------------------------------------------*)
+subsection \<open>Preamble - Pulling \<open>\<exists>\<close> over \<open>\<forall>\<close>\<close>
 
-(* This proof does not need AC and was suggested by the
-   referee for the JCM Paper: let f(x) be least y such
-   that  Q(x,y)
-*)
-lemma no_choice: "\<forall>x. \<exists>y. Q x y ==> \<exists>(f :: 'a => nat). \<forall>x. Q x (f x)"
-apply (rule_tac x = "%x. LEAST y. Q x y" in exI)
-apply (blast intro: LeastI)
-done
+text \<open>This proof does not need AC and was suggested by the
+   referee for the JCM Paper: let \<open>f x\<close> be least \<open>y\<close> such
+   that \<open>Q x y\<close>.\<close>
+lemma no_choice: "\<forall>x. \<exists>y. Q x y \<Longrightarrow> \<exists>f :: 'a \<Rightarrow> nat. \<forall>x. Q x (f x)"
+  by (rule exI [where x = "\<lambda>x. LEAST y. Q x y"]) (blast intro: LeastI)
 
-subsection\<open>Properties of the Star-transform Applied to Sets of Reals\<close>
+
+subsection \<open>Properties of the Star-transform Applied to Sets of Reals\<close>
 
-lemma STAR_star_of_image_subset: "star_of ` A <= *s* A"
-by auto
+lemma STAR_star_of_image_subset: "star_of ` A \<subseteq> *s* A"
+  by auto
 
-lemma STAR_hypreal_of_real_Int: "*s* X Int \<real> = hypreal_of_real ` X"
-by (auto simp add: SReal_def)
+lemma STAR_hypreal_of_real_Int: "*s* X \<inter> \<real> = hypreal_of_real ` X"
+  by (auto simp add: SReal_def)
 
-lemma STAR_star_of_Int: "*s* X Int Standard = star_of ` X"
-by (auto simp add: Standard_def)
+lemma STAR_star_of_Int: "*s* X \<inter> Standard = star_of ` X"
+  by (auto simp add: Standard_def)
 
-lemma lemma_not_hyprealA: "x \<notin> hypreal_of_real ` A ==> \<forall>y \<in> A. x \<noteq> hypreal_of_real y"
-by auto
+lemma lemma_not_hyprealA: "x \<notin> hypreal_of_real ` A \<Longrightarrow> \<forall>y \<in> A. x \<noteq> hypreal_of_real y"
+  by auto
 
-lemma lemma_not_starA: "x \<notin> star_of ` A ==> \<forall>y \<in> A. x \<noteq> star_of y"
-by auto
+lemma lemma_not_starA: "x \<notin> star_of ` A \<Longrightarrow> \<forall>y \<in> A. x \<noteq> star_of y"
+  by auto
 
 lemma lemma_Compl_eq: "- {n. X n = xa} = {n. X n \<noteq> xa}"
-by auto
+  by auto
 
-lemma STAR_real_seq_to_hypreal:
-    "\<forall>n. (X n) \<notin> M ==> star_n X \<notin> *s* M"
-apply (unfold starset_def star_of_def)
-apply (simp add: Iset_star_n FreeUltrafilterNat.proper)
-done
+lemma STAR_real_seq_to_hypreal: "\<forall>n. (X n) \<notin> M \<Longrightarrow> star_n X \<notin> *s* M"
+  by (simp add: starset_def star_of_def Iset_star_n FreeUltrafilterNat.proper)
 
 lemma STAR_singleton: "*s* {x} = {star_of x}"
-by simp
+  by simp
 
-lemma STAR_not_mem: "x \<notin> F ==> star_of x \<notin> *s* F"
-by transfer
-
-lemma STAR_subset_closed: "[| x : *s* A; A <= B |] ==> x : *s* B"
-by (erule rev_subsetD, simp)
+lemma STAR_not_mem: "x \<notin> F \<Longrightarrow> star_of x \<notin> *s* F"
+  by transfer
 
-text\<open>Nonstandard extension of a set (defined using a constant
-   sequence) as a special case of an internal set\<close>
+lemma STAR_subset_closed: "x \<in> *s* A \<Longrightarrow> A \<subseteq> B \<Longrightarrow> x \<in> *s* B"
+  by (erule rev_subsetD) simp
 
-lemma starset_n_starset: "\<forall>n. (As n = A) ==> *sn* As = *s* A"
-apply (drule fun_eq_iff [THEN iffD2])
-apply (simp add: starset_n_def starset_def star_of_def)
-done
+text \<open>Nonstandard extension of a set (defined using a constant
+   sequence) as a special case of an internal set.\<close>
+lemma starset_n_starset: "\<forall>n. As n = A \<Longrightarrow> *sn* As = *s* A"
+  by (drule fun_eq_iff [THEN iffD2]) (simp add: starset_n_def starset_def star_of_def)
 
 
-(*----------------------------------------------------------------*)
-(* Theorems about nonstandard extensions of functions             *)
-(*----------------------------------------------------------------*)
+subsection \<open>Theorems about nonstandard extensions of functions\<close>
 
-(*----------------------------------------------------------------*)
-(* Nonstandard extension of a function (defined using a           *)
-(* constant sequence) as a special case of an internal function   *)
-(*----------------------------------------------------------------*)
+text \<open>Nonstandard extension of a function (defined using a
+  constant sequence) as a special case of an internal function.\<close>
 
-lemma starfun_n_starfun: "\<forall>n. (F n = f) ==> *fn* F = *f* f"
-apply (drule fun_eq_iff [THEN iffD2])
-apply (simp add: starfun_n_def starfun_def star_of_def)
-done
-
+lemma starfun_n_starfun: "\<forall>n. F n = f \<Longrightarrow> *fn* F = *f* f"
+  apply (drule fun_eq_iff [THEN iffD2])
+  apply (simp add: starfun_n_def starfun_def star_of_def)
+  done
 
-(*
-   Prove that abs for hypreal is a nonstandard extension of abs for real w/o
-   use of congruence property (proved after this for general
-   nonstandard extensions of real valued functions).
+text \<open>Prove that \<open>abs\<close> for hypreal is a nonstandard extension of abs for real w/o
+  use of congruence property (proved after this for general
+  nonstandard extensions of real valued functions).
 
-   Proof now Uses the ultrafilter tactic!
-*)
+  Proof now Uses the ultrafilter tactic!\<close>
 
 lemma hrabs_is_starext_rabs: "is_starext abs abs"
-apply (simp add: is_starext_def, safe)
-apply (rule_tac x=x in star_cases)
-apply (rule_tac x=y in star_cases)
-apply (unfold star_n_def, auto)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl)
-apply (fold star_n_def)
-apply (unfold star_abs_def starfun_def star_of_def)
-apply (simp add: Ifun_star_n star_n_eq_iff)
-done
+  apply (simp add: is_starext_def, safe)
+  apply (rule_tac x=x in star_cases)
+  apply (rule_tac x=y in star_cases)
+  apply (unfold star_n_def, auto)
+  apply (rule bexI, rule_tac [2] lemma_starrel_refl)
+  apply (rule bexI, rule_tac [2] lemma_starrel_refl)
+  apply (fold star_n_def)
+  apply (unfold star_abs_def starfun_def star_of_def)
+  apply (simp add: Ifun_star_n star_n_eq_iff)
+  done
 
-text\<open>Nonstandard extension of functions\<close>
 
-lemma starfun:
-      "( *f* f) (star_n X) = star_n (%n. f (X n))"
-by (rule starfun_star_n)
+text \<open>Nonstandard extension of functions.\<close>
+
+lemma starfun: "( *f* f) (star_n X) = star_n (\<lambda>n. f (X n))"
+  by (rule starfun_star_n)
 
-lemma starfun_if_eq:
-     "!!w. w \<noteq> star_of x
-       ==> ( *f* (\<lambda>z. if z = x then a else g z)) w = ( *f* g) w"
-by (transfer, simp)
+lemma starfun_if_eq: "\<And>w. w \<noteq> star_of x \<Longrightarrow> ( *f* (\<lambda>z. if z = x then a else g z)) w = ( *f* g) w"
+  by transfer simp
 
-(*-------------------------------------------
-  multiplication: ( *f) x ( *g) = *(f x g)
- ------------------------------------------*)
-lemma starfun_mult: "!!x. ( *f* f) x * ( *f* g) x = ( *f* (%x. f x * g x)) x"
-by (transfer, rule refl)
+text \<open>Multiplication: \<open>( *f) x ( *g) = *(f x g)\<close>\<close>
+lemma starfun_mult: "\<And>x. ( *f* f) x * ( *f* g) x = ( *f* (\<lambda>x. f x * g x)) x"
+  by transfer (rule refl)
 declare starfun_mult [symmetric, simp]
 
-(*---------------------------------------
-  addition: ( *f) + ( *g) = *(f + g)
- ---------------------------------------*)
-lemma starfun_add: "!!x. ( *f* f) x + ( *f* g) x = ( *f* (%x. f x + g x)) x"
-by (transfer, rule refl)
+text \<open>Addition: \<open>( *f) + ( *g) = *(f + g)\<close>\<close>
+lemma starfun_add: "\<And>x. ( *f* f) x + ( *f* g) x = ( *f* (\<lambda>x. f x + g x)) x"
+  by transfer (rule refl)
 declare starfun_add [symmetric, simp]
 
-(*--------------------------------------------
-  subtraction: ( *f) + -( *g) = *(f + -g)
- -------------------------------------------*)
-lemma starfun_minus: "!!x. - ( *f* f) x = ( *f* (%x. - f x)) x"
-by (transfer, rule refl)
+text \<open>Subtraction: \<open>( *f) + -( *g) = *(f + -g)\<close>\<close>
+lemma starfun_minus: "\<And>x. - ( *f* f) x = ( *f* (\<lambda>x. - f x)) x"
+  by transfer (rule refl)
 declare starfun_minus [symmetric, simp]
 
 (*FIXME: delete*)
-lemma starfun_add_minus: "!!x. ( *f* f) x + -( *f* g) x = ( *f* (%x. f x + -g x)) x"
-by (transfer, rule refl)
+lemma starfun_add_minus: "\<And>x. ( *f* f) x + -( *f* g) x = ( *f* (\<lambda>x. f x + -g x)) x"
+  by transfer (rule refl)
 declare starfun_add_minus [symmetric, simp]
 
-lemma starfun_diff: "!!x. ( *f* f) x  - ( *f* g) x = ( *f* (%x. f x - g x)) x"
-by (transfer, rule refl)
+lemma starfun_diff: "\<And>x. ( *f* f) x  - ( *f* g) x = ( *f* (\<lambda>x. f x - g x)) x"
+  by transfer (rule refl)
 declare starfun_diff [symmetric, simp]
 
-(*--------------------------------------
-  composition: ( *f) o ( *g) = *(f o g)
- ---------------------------------------*)
+text \<open>Composition: \<open>( *f) \<circ> ( *g) = *(f \<circ> g)\<close>\<close>
+lemma starfun_o2: "(\<lambda>x. ( *f* f) (( *f* g) x)) = *f* (\<lambda>x. f (g x))"
+  by transfer (rule refl)
 
-lemma starfun_o2: "(%x. ( *f* f) (( *f* g) x)) = *f* (%x. f (g x))"
-by (transfer, rule refl)
-
-lemma starfun_o: "( *f* f) o ( *f* g) = ( *f* (f o g))"
-by (transfer o_def, rule refl)
+lemma starfun_o: "( *f* f) \<circ> ( *f* g) = ( *f* (f \<circ> g))"
+  by (transfer o_def) (rule refl)
 
-text\<open>NS extension of constant function\<close>
-lemma starfun_const_fun [simp]: "!!x. ( *f* (%x. k)) x = star_of k"
-by (transfer, rule refl)
-
-text\<open>the NS extension of the identity function\<close>
+text \<open>NS extension of constant function.\<close>
+lemma starfun_const_fun [simp]: "\<And>x. ( *f* (\<lambda>x. k)) x = star_of k"
+  by transfer (rule refl)
 
-lemma starfun_Id [simp]: "!!x. ( *f* (%x. x)) x = x"
-by (transfer, rule refl)
+text \<open>The NS extension of the identity function.\<close>
+lemma starfun_Id [simp]: "\<And>x. ( *f* (\<lambda>x. x)) x = x"
+  by transfer (rule refl)
 
-(* this is trivial, given starfun_Id *)
-lemma starfun_Idfun_approx:
-  "x \<approx> star_of a ==> ( *f* (%x. x)) x \<approx> star_of a"
-by (simp only: starfun_Id)
+text \<open>This is trivial, given \<open>starfun_Id\<close>.\<close>
+lemma starfun_Idfun_approx: "x \<approx> star_of a \<Longrightarrow> ( *f* (\<lambda>x. x)) x \<approx> star_of a"
+  by (simp only: starfun_Id)
 
-text\<open>The Star-function is a (nonstandard) extension of the function\<close>
-
+text \<open>The Star-function is a (nonstandard) extension of the function.\<close>
 lemma is_starext_starfun: "is_starext ( *f* f) f"
-apply (simp add: is_starext_def, auto)
-apply (rule_tac x = x in star_cases)
-apply (rule_tac x = y in star_cases)
-apply (auto intro!: bexI [OF _ Rep_star_star_n]
-            simp add: starfun star_n_eq_iff)
-done
-
-text\<open>Any nonstandard extension is in fact the Star-function\<close>
+  apply (auto simp: is_starext_def)
+  apply (rule_tac x = x in star_cases)
+  apply (rule_tac x = y in star_cases)
+  apply (auto intro!: bexI [OF _ Rep_star_star_n] simp: starfun star_n_eq_iff)
+  done
 
-lemma is_starfun_starext: "is_starext F f ==> F = *f* f"
-apply (simp add: is_starext_def)
-apply (rule ext)
-apply (rule_tac x = x in star_cases)
-apply (drule_tac x = x in spec)
-apply (drule_tac x = "( *f* f) x" in spec)
-apply (auto simp add: starfun_star_n)
-apply (simp add: star_n_eq_iff [symmetric])
-apply (simp add: starfun_star_n [of f, symmetric])
-done
+text \<open>Any nonstandard extension is in fact the Star-function.\<close>
+lemma is_starfun_starext: "is_starext F f \<Longrightarrow> F = *f* f"
+  apply (simp add: is_starext_def)
+  apply (rule ext)
+  apply (rule_tac x = x in star_cases)
+  apply (drule_tac x = x in spec)
+  apply (drule_tac x = "( *f* f) x" in spec)
+  apply (auto simp add: starfun_star_n)
+  apply (simp add: star_n_eq_iff [symmetric])
+  apply (simp add: starfun_star_n [of f, symmetric])
+  done
 
-lemma is_starext_starfun_iff: "(is_starext F f) = (F = *f* f)"
-by (blast intro: is_starfun_starext is_starext_starfun)
+lemma is_starext_starfun_iff: "is_starext F f \<longleftrightarrow> F = *f* f"
+  by (blast intro: is_starfun_starext is_starext_starfun)
 
-text\<open>extented function has same solution as its standard
-   version for real arguments. i.e they are the same
-   for all real arguments\<close>
+text \<open>Extented function has same solution as its standard version
+  for real arguments. i.e they are the same for all real arguments.\<close>
 lemma starfun_eq: "( *f* f) (star_of a) = star_of (f a)"
-by (rule starfun_star_of)
+  by (rule starfun_star_of)
 
 lemma starfun_approx: "( *f* f) (star_of a) \<approx> star_of (f a)"
-by simp
+  by simp
 
-(* useful for NS definition of derivatives *)
-lemma starfun_lambda_cancel:
-  "!!x'. ( *f* (%h. f (x + h))) x'  = ( *f* f) (star_of x + x')"
-by (transfer, rule refl)
+text \<open>Useful for NS definition of derivatives.\<close>
+lemma starfun_lambda_cancel: "\<And>x'. ( *f* (\<lambda>h. f (x + h))) x'  = ( *f* f) (star_of x + x')"
+  by transfer (rule refl)
 
-lemma starfun_lambda_cancel2:
-  "( *f* (%h. f(g(x + h)))) x' = ( *f* (f o g)) (star_of x + x')"
-by (unfold o_def, rule starfun_lambda_cancel)
+lemma starfun_lambda_cancel2: "( *f* (\<lambda>h. f (g (x + h)))) x' = ( *f* (f \<circ> g)) (star_of x + x')"
+  unfolding o_def by (rule starfun_lambda_cancel)
 
 lemma starfun_mult_HFinite_approx:
-  fixes l m :: "'a::real_normed_algebra star"
-  shows "[| ( *f* f) x \<approx> l; ( *f* g) x \<approx> m;
-                  l: HFinite; m: HFinite
-               |] ==>  ( *f* (%x. f x * g x)) x \<approx> l * m"
-apply (drule (3) approx_mult_HFinite)
-apply (auto intro: approx_HFinite [OF _ approx_sym])
-done
+  "( *f* f) x \<approx> l \<Longrightarrow> ( *f* g) x \<approx> m \<Longrightarrow> l \<in> HFinite \<Longrightarrow> m \<in> HFinite \<Longrightarrow>
+    ( *f* (\<lambda>x. f x * g x)) x \<approx> l * m"
+  for l m :: "'a::real_normed_algebra star"
+  apply (drule (3) approx_mult_HFinite)
+  apply (auto intro: approx_HFinite [OF _ approx_sym])
+  done
 
-lemma starfun_add_approx: "[| ( *f* f) x \<approx> l; ( *f* g) x \<approx> m
-               |] ==>  ( *f* (%x. f x + g x)) x \<approx> l + m"
-by (auto intro: approx_add)
+lemma starfun_add_approx: "( *f* f) x \<approx> l \<Longrightarrow> ( *f* g) x \<approx> m \<Longrightarrow> ( *f* (%x. f x + g x)) x \<approx> l + m"
+  by (auto intro: approx_add)
 
-text\<open>Examples: hrabs is nonstandard extension of rabs
-              inverse is nonstandard extension of inverse\<close>
+text \<open>Examples: \<open>hrabs\<close> is nonstandard extension of \<open>rabs\<close>,
+  \<open>inverse\<close> is nonstandard extension of \<open>inverse\<close>.\<close>
 
-(* can be proved easily using theorem "starfun" and *)
-(* properties of ultrafilter as for inverse below we  *)
-(* use the theorem we proved above instead          *)
+text \<open>Can be proved easily using theorem \<open>starfun\<close> and
+  properties of ultrafilter as for inverse below we
+  use the theorem we proved above instead.\<close>
 
 lemma starfun_rabs_hrabs: "*f* abs = abs"
-by (simp only: star_abs_def)
+  by (simp only: star_abs_def)
 
-lemma starfun_inverse_inverse [simp]: "( *f* inverse) x = inverse(x)"
-by (simp only: star_inverse_def)
+lemma starfun_inverse_inverse [simp]: "( *f* inverse) x = inverse x"
+  by (simp only: star_inverse_def)
 
-lemma starfun_inverse: "!!x. inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x"
-by (transfer, rule refl)
+lemma starfun_inverse: "\<And>x. inverse (( *f* f) x) = ( *f* (\<lambda>x. inverse (f x))) x"
+  by transfer (rule refl)
 declare starfun_inverse [symmetric, simp]
 
-lemma starfun_divide: "!!x. ( *f* f) x / ( *f* g) x = ( *f* (%x. f x / g x)) x"
-by (transfer, rule refl)
+lemma starfun_divide: "\<And>x. ( *f* f) x / ( *f* g) x = ( *f* (\<lambda>x. f x / g x)) x"
+  by transfer (rule refl)
 declare starfun_divide [symmetric, simp]
 
-lemma starfun_inverse2: "!!x. inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x"
-by (transfer, rule refl)
+lemma starfun_inverse2: "\<And>x. inverse (( *f* f) x) = ( *f* (\<lambda>x. inverse (f x))) x"
+  by transfer (rule refl)
 
-text\<open>General lemma/theorem needed for proofs in elementary
-    topology of the reals\<close>
-lemma starfun_mem_starset:
-      "!!x. ( *f* f) x : *s* A ==> x : *s* {x. f x  \<in> A}"
-by (transfer, simp)
+text \<open>General lemma/theorem needed for proofs in elementary topology of the reals.\<close>
+lemma starfun_mem_starset: "\<And>x. ( *f* f) x : *s* A \<Longrightarrow> x \<in> *s* {x. f x \<in> A}"
+  by transfer simp
 
-text\<open>Alternative definition for hrabs with rabs function
-   applied entrywise to equivalence class representative.
-   This is easily proved using starfun and ns extension thm\<close>
-lemma hypreal_hrabs: "\<bar>star_n X\<bar> = star_n (%n. \<bar>X n\<bar>)"
-by (simp only: starfun_rabs_hrabs [symmetric] starfun)
+text \<open>Alternative definition for \<open>hrabs\<close> with \<open>rabs\<close> function applied
+  entrywise to equivalence class representative.
+  This is easily proved using @{thm [source] starfun} and ns extension thm.\<close>
+lemma hypreal_hrabs: "\<bar>star_n X\<bar> = star_n (\<lambda>n. \<bar>X n\<bar>)"
+  by (simp only: starfun_rabs_hrabs [symmetric] starfun)
 
-text\<open>nonstandard extension of set through nonstandard extension
-   of rabs function i.e hrabs. A more general result should be
-   where we replace rabs by some arbitrary function f and hrabs
+text \<open>Nonstandard extension of set through nonstandard extension
+   of \<open>rabs\<close> function i.e. \<open>hrabs\<close>. A more general result should be
+   where we replace \<open>rabs\<close> by some arbitrary function \<open>f\<close> and \<open>hrabs\<close>
    by its NS extenson. See second NS set extension below.\<close>
-lemma STAR_rabs_add_minus:
-   "*s* {x. \<bar>x + - y\<bar> < r} = {x. \<bar>x + -star_of y\<bar> < star_of r}"
-by (transfer, rule refl)
+lemma STAR_rabs_add_minus: "*s* {x. \<bar>x + - y\<bar> < r} = {x. \<bar>x + -star_of y\<bar> < star_of r}"
+  by transfer (rule refl)
 
 lemma STAR_starfun_rabs_add_minus:
-  "*s* {x. \<bar>f x + - y\<bar> < r} =
-       {x. \<bar>( *f* f) x + -star_of y\<bar> < star_of r}"
-by (transfer, rule refl)
+  "*s* {x. \<bar>f x + - y\<bar> < r} = {x. \<bar>( *f* f) x + -star_of y\<bar> < star_of r}"
+  by transfer (rule refl)
 
-text\<open>Another characterization of Infinitesimal and one of \<open>\<approx>\<close> relation.
-   In this theory since \<open>hypreal_hrabs\<close> proved here. Maybe
-   move both theorems??\<close>
+text \<open>Another characterization of Infinitesimal and one of \<open>\<approx>\<close> relation.
+  In this theory since \<open>hypreal_hrabs\<close> proved here. Maybe move both theorems??\<close>
 lemma Infinitesimal_FreeUltrafilterNat_iff2:
-     "(star_n X \<in> Infinitesimal) = (\<forall>m. eventually (\<lambda>n. norm(X n) < inverse(real(Suc m))) \<U>)"
-by (simp add: Infinitesimal_hypreal_of_nat_iff star_of_def
-     hnorm_def star_of_nat_def starfun_star_n
-     star_n_inverse star_n_less)
+  "star_n X \<in> Infinitesimal \<longleftrightarrow> (\<forall>m. eventually (\<lambda>n. norm (X n) < inverse (real (Suc m))) \<U>)"
+  by (simp add: Infinitesimal_hypreal_of_nat_iff star_of_def hnorm_def
+      star_of_nat_def starfun_star_n star_n_inverse star_n_less)
 
 lemma HNatInfinite_inverse_Infinitesimal [simp]:
-     "n \<in> HNatInfinite ==> inverse (hypreal_of_hypnat n) \<in> Infinitesimal"
-apply (cases n)
-apply (auto simp add: of_hypnat_def starfun_star_n star_n_inverse real_norm_def
-      HNatInfinite_FreeUltrafilterNat_iff
-      Infinitesimal_FreeUltrafilterNat_iff2)
-apply (drule_tac x="Suc m" in spec)
-apply (auto elim!: eventually_mono)
-done
+  "n \<in> HNatInfinite \<Longrightarrow> inverse (hypreal_of_hypnat n) \<in> Infinitesimal"
+  apply (cases n)
+  apply (auto simp: of_hypnat_def starfun_star_n star_n_inverse
+    HNatInfinite_FreeUltrafilterNat_iff Infinitesimal_FreeUltrafilterNat_iff2)
+  apply (drule_tac x = "Suc m" in spec)
+  apply (auto elim!: eventually_mono)
+  done
 
-lemma approx_FreeUltrafilterNat_iff: "star_n X \<approx> star_n Y =
-      (\<forall>r>0. eventually (\<lambda>n. norm (X n - Y n) < r) \<U>)"
-apply (subst approx_minus_iff)
-apply (rule mem_infmal_iff [THEN subst])
-apply (simp add: star_n_diff)
-apply (simp add: Infinitesimal_FreeUltrafilterNat_iff)
-done
+lemma approx_FreeUltrafilterNat_iff:
+  "star_n X \<approx> star_n Y \<longleftrightarrow> (\<forall>r>0. eventually (\<lambda>n. norm (X n - Y n) < r) \<U>)"
+  apply (subst approx_minus_iff)
+  apply (rule mem_infmal_iff [THEN subst])
+  apply (simp add: star_n_diff)
+  apply (simp add: Infinitesimal_FreeUltrafilterNat_iff)
+  done
 
-lemma approx_FreeUltrafilterNat_iff2: "star_n X \<approx> star_n Y =
-      (\<forall>m. eventually (\<lambda>n. norm (X n - Y n) < inverse(real(Suc m))) \<U>)"
-apply (subst approx_minus_iff)
-apply (rule mem_infmal_iff [THEN subst])
-apply (simp add: star_n_diff)
-apply (simp add: Infinitesimal_FreeUltrafilterNat_iff2)
-done
+lemma approx_FreeUltrafilterNat_iff2:
+  "star_n X \<approx> star_n Y \<longleftrightarrow> (\<forall>m. eventually (\<lambda>n. norm (X n - Y n) < inverse (real (Suc m))) \<U>)"
+  apply (subst approx_minus_iff)
+  apply (rule mem_infmal_iff [THEN subst])
+  apply (simp add: star_n_diff)
+  apply (simp add: Infinitesimal_FreeUltrafilterNat_iff2)
+  done
 
 lemma inj_starfun: "inj starfun"
-apply (rule inj_onI)
-apply (rule ext, rule ccontr)
-apply (drule_tac x = "star_n (%n. xa)" in fun_cong)
-apply (auto simp add: starfun star_n_eq_iff FreeUltrafilterNat.proper)
-done
+  apply (rule inj_onI)
+  apply (rule ext, rule ccontr)
+  apply (drule_tac x = "star_n (\<lambda>n. xa)" in fun_cong)
+  apply (auto simp add: starfun star_n_eq_iff FreeUltrafilterNat.proper)
+  done
 
 end
--- a/src/HOL/Nonstandard_Analysis/StarDef.thy	Mon Oct 31 16:26:36 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/StarDef.thy	Tue Nov 01 00:44:24 2016 +0100
@@ -5,54 +5,58 @@
 section \<open>Construction of Star Types Using Ultrafilters\<close>
 
 theory StarDef
-imports Free_Ultrafilter
+  imports Free_Ultrafilter
 begin
 
 subsection \<open>A Free Ultrafilter over the Naturals\<close>
 
-definition
-  FreeUltrafilterNat :: "nat filter"  ("\<U>") where
-  "\<U> = (SOME U. freeultrafilter U)"
+definition FreeUltrafilterNat :: "nat filter"  ("\<U>")
+  where "\<U> = (SOME U. freeultrafilter U)"
 
 lemma freeultrafilter_FreeUltrafilterNat: "freeultrafilter \<U>"
-apply (unfold FreeUltrafilterNat_def)
-apply (rule someI_ex)
-apply (rule freeultrafilter_Ex)
-apply (rule infinite_UNIV_nat)
-done
+  apply (unfold FreeUltrafilterNat_def)
+  apply (rule someI_ex)
+  apply (rule freeultrafilter_Ex)
+  apply (rule infinite_UNIV_nat)
+  done
 
 interpretation FreeUltrafilterNat: freeultrafilter FreeUltrafilterNat
-by (rule freeultrafilter_FreeUltrafilterNat)
+  by (rule freeultrafilter_FreeUltrafilterNat)
+
 
 subsection \<open>Definition of \<open>star\<close> type constructor\<close>
 
-definition
-  starrel :: "((nat \<Rightarrow> 'a) \<times> (nat \<Rightarrow> 'a)) set" where
-  "starrel = {(X,Y). eventually (\<lambda>n. X n = Y n) \<U>}"
+definition starrel :: "((nat \<Rightarrow> 'a) \<times> (nat \<Rightarrow> 'a)) set"
+  where "starrel = {(X, Y). eventually (\<lambda>n. X n = Y n) \<U>}"
 
 definition "star = (UNIV :: (nat \<Rightarrow> 'a) set) // starrel"
 
 typedef 'a star = "star :: (nat \<Rightarrow> 'a) set set"
-  unfolding star_def by (auto intro: quotientI)
+  by (auto simp: star_def intro: quotientI)
 
-definition
-  star_n :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a star" where
-  "star_n X = Abs_star (starrel `` {X})"
+definition star_n :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a star"
+  where "star_n X = Abs_star (starrel `` {X})"
 
 theorem star_cases [case_names star_n, cases type: star]:
-  "(\<And>X. x = star_n X \<Longrightarrow> P) \<Longrightarrow> P"
-by (cases x, unfold star_n_def star_def, erule quotientE, fast)
+  obtains X where "x = star_n X"
+  by (cases x) (auto simp: star_n_def star_def elim: quotientE)
 
-lemma all_star_eq: "(\<forall>x. P x) = (\<forall>X. P (star_n X))"
-by (auto, rule_tac x=x in star_cases, simp)
+lemma all_star_eq: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>X. P (star_n X))"
+  apply auto
+  apply (rule_tac x = x in star_cases)
+  apply simp
+  done
 
-lemma ex_star_eq: "(\<exists>x. P x) = (\<exists>X. P (star_n X))"
-by (auto, rule_tac x=x in star_cases, auto)
+lemma ex_star_eq: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>X. P (star_n X))"
+  apply auto
+  apply (rule_tac x=x in star_cases)
+  apply auto
+  done
 
-text \<open>Proving that @{term starrel} is an equivalence relation\<close>
+text \<open>Proving that @{term starrel} is an equivalence relation.\<close>
 
-lemma starrel_iff [iff]: "((X,Y) \<in> starrel) = (eventually (\<lambda>n. X n = Y n) \<U>)"
-by (simp add: starrel_def)
+lemma starrel_iff [iff]: "(X, Y) \<in> starrel \<longleftrightarrow> eventually (\<lambda>n. X n = Y n) \<U>"
+  by (simp add: starrel_def)
 
 lemma equiv_starrel: "equiv UNIV starrel"
 proof (rule equivI)
@@ -61,356 +65,322 @@
   show "trans starrel" by (intro transI) (auto elim: eventually_elim2)
 qed
 
-lemmas equiv_starrel_iff =
-  eq_equiv_class_iff [OF equiv_starrel UNIV_I UNIV_I]
+lemmas equiv_starrel_iff = eq_equiv_class_iff [OF equiv_starrel UNIV_I UNIV_I]
 
 lemma starrel_in_star: "starrel``{x} \<in> star"
-by (simp add: star_def quotientI)
+  by (simp add: star_def quotientI)
 
-lemma star_n_eq_iff: "(star_n X = star_n Y) = (eventually (\<lambda>n. X n = Y n) \<U>)"
-by (simp add: star_n_def Abs_star_inject starrel_in_star equiv_starrel_iff)
+lemma star_n_eq_iff: "star_n X = star_n Y \<longleftrightarrow> eventually (\<lambda>n. X n = Y n) \<U>"
+  by (simp add: star_n_def Abs_star_inject starrel_in_star equiv_starrel_iff)
 
 
 subsection \<open>Transfer principle\<close>
 
 text \<open>This introduction rule starts each transfer proof.\<close>
-lemma transfer_start:
-  "P \<equiv> eventually (\<lambda>n. Q) \<U> \<Longrightarrow> Trueprop P \<equiv> Trueprop Q"
+lemma transfer_start: "P \<equiv> eventually (\<lambda>n. Q) \<U> \<Longrightarrow> Trueprop P \<equiv> Trueprop Q"
   by (simp add: FreeUltrafilterNat.proper)
 
 text \<open>Standard principles that play a central role in the transfer tactic.\<close>
-definition
-  Ifun :: "('a \<Rightarrow> 'b) star \<Rightarrow> 'a star \<Rightarrow> 'b star" ("_ \<star> _" [300,301] 300) where
-  "Ifun f \<equiv> \<lambda>x. Abs_star
-       (\<Union>F\<in>Rep_star f. \<Union>X\<in>Rep_star x. starrel``{\<lambda>n. F n (X n)})"
+definition Ifun :: "('a \<Rightarrow> 'b) star \<Rightarrow> 'a star \<Rightarrow> 'b star" ("_ \<star> _" [300,301] 300)
+  where "Ifun f \<equiv>
+    \<lambda>x. Abs_star (\<Union>F\<in>Rep_star f. \<Union>X\<in>Rep_star x. starrel``{\<lambda>n. F n (X n)})"
 
-lemma Ifun_congruent2:
-  "congruent2 starrel starrel (\<lambda>F X. starrel``{\<lambda>n. F n (X n)})"
-by (auto simp add: congruent2_def equiv_starrel_iff elim!: eventually_rev_mp)
+lemma Ifun_congruent2: "congruent2 starrel starrel (\<lambda>F X. starrel``{\<lambda>n. F n (X n)})"
+  by (auto simp add: congruent2_def equiv_starrel_iff elim!: eventually_rev_mp)
 
 lemma Ifun_star_n: "star_n F \<star> star_n X = star_n (\<lambda>n. F n (X n))"
-by (simp add: Ifun_def star_n_def Abs_star_inverse starrel_in_star
-    UN_equiv_class2 [OF equiv_starrel equiv_starrel Ifun_congruent2])
+  by (simp add: Ifun_def star_n_def Abs_star_inverse starrel_in_star
+      UN_equiv_class2 [OF equiv_starrel equiv_starrel Ifun_congruent2])
 
-lemma transfer_Ifun:
-  "\<lbrakk>f \<equiv> star_n F; x \<equiv> star_n X\<rbrakk> \<Longrightarrow> f \<star> x \<equiv> star_n (\<lambda>n. F n (X n))"
-by (simp only: Ifun_star_n)
+lemma transfer_Ifun: "f \<equiv> star_n F \<Longrightarrow> x \<equiv> star_n X \<Longrightarrow> f \<star> x \<equiv> star_n (\<lambda>n. F n (X n))"
+  by (simp only: Ifun_star_n)
 
-definition
-  star_of :: "'a \<Rightarrow> 'a star" where
-  "star_of x == star_n (\<lambda>n. x)"
+definition star_of :: "'a \<Rightarrow> 'a star"
+  where "star_of x \<equiv> star_n (\<lambda>n. x)"
 
 text \<open>Initialize transfer tactic.\<close>
 ML_file "transfer.ML"
 
-method_setup transfer = \<open>
-  Attrib.thms >> (fn ths => fn ctxt =>
-    SIMPLE_METHOD' (Transfer_Principle.transfer_tac ctxt ths))
-\<close> "transfer principle"
+method_setup transfer =
+  \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (Transfer_Principle.transfer_tac ctxt ths))\<close>
+  "transfer principle"
 
 
 text \<open>Transfer introduction rules.\<close>
 
 lemma transfer_ex [transfer_intro]:
-  "\<lbrakk>\<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>\<rbrakk>
-    \<Longrightarrow> \<exists>x::'a star. p x \<equiv> eventually (\<lambda>n. \<exists>x. P n x) \<U>"
-by (simp only: ex_star_eq eventually_ex)
+  "(\<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>) \<Longrightarrow>
+    \<exists>x::'a star. p x \<equiv> eventually (\<lambda>n. \<exists>x. P n x) \<U>"
+  by (simp only: ex_star_eq eventually_ex)
 
 lemma transfer_all [transfer_intro]:
-  "\<lbrakk>\<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>\<rbrakk>
-    \<Longrightarrow> \<forall>x::'a star. p x \<equiv> eventually (\<lambda>n. \<forall>x. P n x) \<U>"
-by (simp only: all_star_eq FreeUltrafilterNat.eventually_all_iff)
+  "(\<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>) \<Longrightarrow>
+    \<forall>x::'a star. p x \<equiv> eventually (\<lambda>n. \<forall>x. P n x) \<U>"
+  by (simp only: all_star_eq FreeUltrafilterNat.eventually_all_iff)
 
-lemma transfer_not [transfer_intro]:
-  "\<lbrakk>p \<equiv> eventually P \<U>\<rbrakk> \<Longrightarrow> \<not> p \<equiv> eventually (\<lambda>n. \<not> P n) \<U>"
-by (simp only: FreeUltrafilterNat.eventually_not_iff)
+lemma transfer_not [transfer_intro]: "p \<equiv> eventually P \<U> \<Longrightarrow> \<not> p \<equiv> eventually (\<lambda>n. \<not> P n) \<U>"
+  by (simp only: FreeUltrafilterNat.eventually_not_iff)
 
 lemma transfer_conj [transfer_intro]:
-  "\<lbrakk>p \<equiv> eventually P \<U>; q \<equiv> eventually Q \<U>\<rbrakk>
-    \<Longrightarrow> p \<and> q \<equiv> eventually (\<lambda>n. P n \<and> Q n) \<U>"
-by (simp only: eventually_conj_iff)
+  "p \<equiv> eventually P \<U> \<Longrightarrow> q \<equiv> eventually Q \<U> \<Longrightarrow> p \<and> q \<equiv> eventually (\<lambda>n. P n \<and> Q n) \<U>"
+  by (simp only: eventually_conj_iff)
 
 lemma transfer_disj [transfer_intro]:
-  "\<lbrakk>p \<equiv> eventually P \<U>; q \<equiv> eventually Q \<U>\<rbrakk>
-    \<Longrightarrow> p \<or> q \<equiv> eventually (\<lambda>n. P n \<or> Q n) \<U>"
-by (simp only: FreeUltrafilterNat.eventually_disj_iff)
+  "p \<equiv> eventually P \<U> \<Longrightarrow> q \<equiv> eventually Q \<U> \<Longrightarrow> p \<or> q \<equiv> eventually (\<lambda>n. P n \<or> Q n) \<U>"
+  by (simp only: FreeUltrafilterNat.eventually_disj_iff)
 
 lemma transfer_imp [transfer_intro]:
-  "\<lbrakk>p \<equiv> eventually P \<U>; q \<equiv> eventually Q \<U>\<rbrakk>
-    \<Longrightarrow> p \<longrightarrow> q \<equiv> eventually (\<lambda>n. P n \<longrightarrow> Q n) \<U>"
-by (simp only: FreeUltrafilterNat.eventually_imp_iff)
+  "p \<equiv> eventually P \<U> \<Longrightarrow> q \<equiv> eventually Q \<U> \<Longrightarrow> p \<longrightarrow> q \<equiv> eventually (\<lambda>n. P n \<longrightarrow> Q n) \<U>"
+  by (simp only: FreeUltrafilterNat.eventually_imp_iff)
 
 lemma transfer_iff [transfer_intro]:
-  "\<lbrakk>p \<equiv> eventually P \<U>; q \<equiv> eventually Q \<U>\<rbrakk>
-    \<Longrightarrow> p = q \<equiv> eventually (\<lambda>n. P n = Q n) \<U>"
-by (simp only: FreeUltrafilterNat.eventually_iff_iff)
+  "p \<equiv> eventually P \<U> \<Longrightarrow> q \<equiv> eventually Q \<U> \<Longrightarrow> p = q \<equiv> eventually (\<lambda>n. P n = Q n) \<U>"
+  by (simp only: FreeUltrafilterNat.eventually_iff_iff)
 
 lemma transfer_if_bool [transfer_intro]:
-  "\<lbrakk>p \<equiv> eventually P \<U>; x \<equiv> eventually X \<U>; y \<equiv> eventually Y \<U>\<rbrakk>
-    \<Longrightarrow> (if p then x else y) \<equiv> eventually (\<lambda>n. if P n then X n else Y n) \<U>"
-by (simp only: if_bool_eq_conj transfer_conj transfer_imp transfer_not)
+  "p \<equiv> eventually P \<U> \<Longrightarrow> x \<equiv> eventually X \<U> \<Longrightarrow> y \<equiv> eventually Y \<U> \<Longrightarrow>
+    (if p then x else y) \<equiv> eventually (\<lambda>n. if P n then X n else Y n) \<U>"
+  by (simp only: if_bool_eq_conj transfer_conj transfer_imp transfer_not)
 
 lemma transfer_eq [transfer_intro]:
-  "\<lbrakk>x \<equiv> star_n X; y \<equiv> star_n Y\<rbrakk> \<Longrightarrow> x = y \<equiv> eventually (\<lambda>n. X n = Y n) \<U>"
-by (simp only: star_n_eq_iff)
+  "x \<equiv> star_n X \<Longrightarrow> y \<equiv> star_n Y \<Longrightarrow> x = y \<equiv> eventually (\<lambda>n. X n = Y n) \<U>"
+  by (simp only: star_n_eq_iff)
 
 lemma transfer_if [transfer_intro]:
-  "\<lbrakk>p \<equiv> eventually (\<lambda>n. P n) \<U>; x \<equiv> star_n X; y \<equiv> star_n Y\<rbrakk>
-    \<Longrightarrow> (if p then x else y) \<equiv> star_n (\<lambda>n. if P n then X n else Y n)"
-apply (rule eq_reflection)
-apply (auto simp add: star_n_eq_iff transfer_not elim!: eventually_mono)
-done
+  "p \<equiv> eventually (\<lambda>n. P n) \<U> \<Longrightarrow> x \<equiv> star_n X \<Longrightarrow> y \<equiv> star_n Y \<Longrightarrow>
+    (if p then x else y) \<equiv> star_n (\<lambda>n. if P n then X n else Y n)"
+  by (rule eq_reflection) (auto simp: star_n_eq_iff transfer_not elim!: eventually_mono)
 
 lemma transfer_fun_eq [transfer_intro]:
-  "\<lbrakk>\<And>X. f (star_n X) = g (star_n X)
-    \<equiv> eventually (\<lambda>n. F n (X n) = G n (X n)) \<U>\<rbrakk>
-      \<Longrightarrow> f = g \<equiv> eventually (\<lambda>n. F n = G n) \<U>"
-by (simp only: fun_eq_iff transfer_all)
+  "(\<And>X. f (star_n X) = g (star_n X) \<equiv> eventually (\<lambda>n. F n (X n) = G n (X n)) \<U>) \<Longrightarrow>
+    f = g \<equiv> eventually (\<lambda>n. F n = G n) \<U>"
+  by (simp only: fun_eq_iff transfer_all)
 
 lemma transfer_star_n [transfer_intro]: "star_n X \<equiv> star_n (\<lambda>n. X n)"
-by (rule reflexive)
+  by (rule reflexive)
 
 lemma transfer_bool [transfer_intro]: "p \<equiv> eventually (\<lambda>n. p) \<U>"
-by (simp add: FreeUltrafilterNat.proper)
+  by (simp add: FreeUltrafilterNat.proper)
 
 
 subsection \<open>Standard elements\<close>
 
-definition
-  Standard :: "'a star set" where
-  "Standard = range star_of"
+definition Standard :: "'a star set"
+  where "Standard = range star_of"
 
-text \<open>Transfer tactic should remove occurrences of @{term star_of}\<close>
+text \<open>Transfer tactic should remove occurrences of @{term star_of}.\<close>
 setup \<open>Transfer_Principle.add_const @{const_name star_of}\<close>
 
-lemma star_of_inject: "(star_of x = star_of y) = (x = y)"
-by (transfer, rule refl)
+lemma star_of_inject: "star_of x = star_of y \<longleftrightarrow> x = y"
+  by transfer (rule refl)
 
 lemma Standard_star_of [simp]: "star_of x \<in> Standard"
-by (simp add: Standard_def)
+  by (simp add: Standard_def)
+
 
 subsection \<open>Internal functions\<close>
 
-text \<open>Transfer tactic should remove occurrences of @{term Ifun}\<close>
+text \<open>Transfer tactic should remove occurrences of @{term Ifun}.\<close>
 setup \<open>Transfer_Principle.add_const @{const_name Ifun}\<close>
 
 lemma Ifun_star_of [simp]: "star_of f \<star> star_of x = star_of (f x)"
-by (transfer, rule refl)
+  by transfer (rule refl)
 
-lemma Standard_Ifun [simp]:
-  "\<lbrakk>f \<in> Standard; x \<in> Standard\<rbrakk> \<Longrightarrow> f \<star> x \<in> Standard"
-by (auto simp add: Standard_def)
+lemma Standard_Ifun [simp]: "f \<in> Standard \<Longrightarrow> x \<in> Standard \<Longrightarrow> f \<star> x \<in> Standard"
+  by (auto simp add: Standard_def)
 
-text \<open>Nonstandard extensions of functions\<close>
 
-definition
-  starfun :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a star \<Rightarrow> 'b star)"  ("*f* _" [80] 80) where
-  "starfun f == \<lambda>x. star_of f \<star> x"
+text \<open>Nonstandard extensions of functions.\<close>
 
-definition
-  starfun2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a star \<Rightarrow> 'b star \<Rightarrow> 'c star)"
-    ("*f2* _" [80] 80) where
-  "starfun2 f == \<lambda>x y. star_of f \<star> x \<star> y"
+definition starfun :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a star \<Rightarrow> 'b star"  ("*f* _" [80] 80)
+  where "starfun f \<equiv> \<lambda>x. star_of f \<star> x"
+
+definition starfun2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a star \<Rightarrow> 'b star \<Rightarrow> 'c star"  ("*f2* _" [80] 80)
+  where "starfun2 f \<equiv> \<lambda>x y. star_of f \<star> x \<star> y"
 
 declare starfun_def [transfer_unfold]
 declare starfun2_def [transfer_unfold]
 
 lemma starfun_star_n: "( *f* f) (star_n X) = star_n (\<lambda>n. f (X n))"
-by (simp only: starfun_def star_of_def Ifun_star_n)
+  by (simp only: starfun_def star_of_def Ifun_star_n)
 
-lemma starfun2_star_n:
-  "( *f2* f) (star_n X) (star_n Y) = star_n (\<lambda>n. f (X n) (Y n))"
-by (simp only: starfun2_def star_of_def Ifun_star_n)
+lemma starfun2_star_n: "( *f2* f) (star_n X) (star_n Y) = star_n (\<lambda>n. f (X n) (Y n))"
+  by (simp only: starfun2_def star_of_def Ifun_star_n)
 
 lemma starfun_star_of [simp]: "( *f* f) (star_of x) = star_of (f x)"
-by (transfer, rule refl)
+  by transfer (rule refl)
 
 lemma starfun2_star_of [simp]: "( *f2* f) (star_of x) = *f* f x"
-by (transfer, rule refl)
+  by transfer (rule refl)
 
 lemma Standard_starfun [simp]: "x \<in> Standard \<Longrightarrow> starfun f x \<in> Standard"
-by (simp add: starfun_def)
+  by (simp add: starfun_def)
 
-lemma Standard_starfun2 [simp]:
-  "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> starfun2 f x y \<in> Standard"
-by (simp add: starfun2_def)
+lemma Standard_starfun2 [simp]: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> starfun2 f x y \<in> Standard"
+  by (simp add: starfun2_def)
 
 lemma Standard_starfun_iff:
   assumes inj: "\<And>x y. f x = f y \<Longrightarrow> x = y"
-  shows "(starfun f x \<in> Standard) = (x \<in> Standard)"
+  shows "starfun f x \<in> Standard \<longleftrightarrow> x \<in> Standard"
 proof
   assume "x \<in> Standard"
-  thus "starfun f x \<in> Standard" by simp
+  then show "starfun f x \<in> Standard" by simp
 next
-  have inj': "\<And>x y. starfun f x = starfun f y \<Longrightarrow> x = y"
-    using inj by transfer
+  from inj have inj': "\<And>x y. starfun f x = starfun f y \<Longrightarrow> x = y"
+    by transfer
   assume "starfun f x \<in> Standard"
   then obtain b where b: "starfun f x = star_of b"
     unfolding Standard_def ..
-  hence "\<exists>x. starfun f x = star_of b" ..
-  hence "\<exists>a. f a = b" by transfer
+  then have "\<exists>x. starfun f x = star_of b" ..
+  then have "\<exists>a. f a = b" by transfer
   then obtain a where "f a = b" ..
-  hence "starfun f (star_of a) = star_of b" by transfer
+  then have "starfun f (star_of a) = star_of b" by transfer
   with b have "starfun f x = starfun f (star_of a)" by simp
-  hence "x = star_of a" by (rule inj')
-  thus "x \<in> Standard"
-    unfolding Standard_def by auto
+  then have "x = star_of a" by (rule inj')
+  then show "x \<in> Standard" by (simp add: Standard_def)
 qed
 
 lemma Standard_starfun2_iff:
   assumes inj: "\<And>a b a' b'. f a b = f a' b' \<Longrightarrow> a = a' \<and> b = b'"
-  shows "(starfun2 f x y \<in> Standard) = (x \<in> Standard \<and> y \<in> Standard)"
+  shows "starfun2 f x y \<in> Standard \<longleftrightarrow> x \<in> Standard \<and> y \<in> Standard"
 proof
   assume "x \<in> Standard \<and> y \<in> Standard"
-  thus "starfun2 f x y \<in> Standard" by simp
+  then show "starfun2 f x y \<in> Standard" by simp
 next
   have inj': "\<And>x y z w. starfun2 f x y = starfun2 f z w \<Longrightarrow> x = z \<and> y = w"
     using inj by transfer
   assume "starfun2 f x y \<in> Standard"
   then obtain c where c: "starfun2 f x y = star_of c"
     unfolding Standard_def ..
-  hence "\<exists>x y. starfun2 f x y = star_of c" by auto
-  hence "\<exists>a b. f a b = c" by transfer
+  then have "\<exists>x y. starfun2 f x y = star_of c" by auto
+  then have "\<exists>a b. f a b = c" by transfer
   then obtain a b where "f a b = c" by auto
-  hence "starfun2 f (star_of a) (star_of b) = star_of c"
-    by transfer
-  with c have "starfun2 f x y = starfun2 f (star_of a) (star_of b)"
-    by simp
-  hence "x = star_of a \<and> y = star_of b"
-    by (rule inj')
-  thus "x \<in> Standard \<and> y \<in> Standard"
-    unfolding Standard_def by auto
+  then have "starfun2 f (star_of a) (star_of b) = star_of c" by transfer
+  with c have "starfun2 f x y = starfun2 f (star_of a) (star_of b)" by simp
+  then have "x = star_of a \<and> y = star_of b" by (rule inj')
+  then show "x \<in> Standard \<and> y \<in> Standard" by (simp add: Standard_def)
 qed
 
 
 subsection \<open>Internal predicates\<close>
 
-definition unstar :: "bool star \<Rightarrow> bool" where
-  "unstar b \<longleftrightarrow> b = star_of True"
+definition unstar :: "bool star \<Rightarrow> bool"
+  where "unstar b \<longleftrightarrow> b = star_of True"
 
-lemma unstar_star_n: "unstar (star_n P) = (eventually P \<U>)"
-by (simp add: unstar_def star_of_def star_n_eq_iff)
+lemma unstar_star_n: "unstar (star_n P) \<longleftrightarrow> eventually P \<U>"
+  by (simp add: unstar_def star_of_def star_n_eq_iff)
 
 lemma unstar_star_of [simp]: "unstar (star_of p) = p"
-by (simp add: unstar_def star_of_inject)
+  by (simp add: unstar_def star_of_inject)
 
-text \<open>Transfer tactic should remove occurrences of @{term unstar}\<close>
+text \<open>Transfer tactic should remove occurrences of @{term unstar}.\<close>
 setup \<open>Transfer_Principle.add_const @{const_name unstar}\<close>
 
-lemma transfer_unstar [transfer_intro]:
-  "p \<equiv> star_n P \<Longrightarrow> unstar p \<equiv> eventually P \<U>"
-by (simp only: unstar_star_n)
+lemma transfer_unstar [transfer_intro]: "p \<equiv> star_n P \<Longrightarrow> unstar p \<equiv> eventually P \<U>"
+  by (simp only: unstar_star_n)
 
-definition
-  starP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> bool"  ("*p* _" [80] 80) where
-  "*p* P = (\<lambda>x. unstar (star_of P \<star> x))"
+definition starP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> bool"  ("*p* _" [80] 80)
+  where "*p* P = (\<lambda>x. unstar (star_of P \<star> x))"
 
-definition
-  starP2 :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> 'b star \<Rightarrow> bool"  ("*p2* _" [80] 80) where
-  "*p2* P = (\<lambda>x y. unstar (star_of P \<star> x \<star> y))"
+definition starP2 :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> 'b star \<Rightarrow> bool"  ("*p2* _" [80] 80)
+  where "*p2* P = (\<lambda>x y. unstar (star_of P \<star> x \<star> y))"
 
 declare starP_def [transfer_unfold]
 declare starP2_def [transfer_unfold]
 
-lemma starP_star_n: "( *p* P) (star_n X) = (eventually (\<lambda>n. P (X n)) \<U>)"
-by (simp only: starP_def star_of_def Ifun_star_n unstar_star_n)
+lemma starP_star_n: "( *p* P) (star_n X) = eventually (\<lambda>n. P (X n)) \<U>"
+  by (simp only: starP_def star_of_def Ifun_star_n unstar_star_n)
 
-lemma starP2_star_n:
-  "( *p2* P) (star_n X) (star_n Y) = (eventually (\<lambda>n. P (X n) (Y n)) \<U>)"
-by (simp only: starP2_def star_of_def Ifun_star_n unstar_star_n)
+lemma starP2_star_n: "( *p2* P) (star_n X) (star_n Y) = (eventually (\<lambda>n. P (X n) (Y n)) \<U>)"
+  by (simp only: starP2_def star_of_def Ifun_star_n unstar_star_n)
 
 lemma starP_star_of [simp]: "( *p* P) (star_of x) = P x"
-by (transfer, rule refl)
+  by transfer (rule refl)
 
 lemma starP2_star_of [simp]: "( *p2* P) (star_of x) = *p* P x"
-by (transfer, rule refl)
+  by transfer (rule refl)
 
 
 subsection \<open>Internal sets\<close>
 
-definition
-  Iset :: "'a set star \<Rightarrow> 'a star set" where
-  "Iset A = {x. ( *p2* op \<in>) x A}"
+definition Iset :: "'a set star \<Rightarrow> 'a star set"
+  where "Iset A = {x. ( *p2* op \<in>) x A}"
 
-lemma Iset_star_n:
-  "(star_n X \<in> Iset (star_n A)) = (eventually (\<lambda>n. X n \<in> A n) \<U>)"
-by (simp add: Iset_def starP2_star_n)
+lemma Iset_star_n: "(star_n X \<in> Iset (star_n A)) = (eventually (\<lambda>n. X n \<in> A n) \<U>)"
+  by (simp add: Iset_def starP2_star_n)
 
-text \<open>Transfer tactic should remove occurrences of @{term Iset}\<close>
+text \<open>Transfer tactic should remove occurrences of @{term Iset}.\<close>
 setup \<open>Transfer_Principle.add_const @{const_name Iset}\<close>
 
 lemma transfer_mem [transfer_intro]:
-  "\<lbrakk>x \<equiv> star_n X; a \<equiv> Iset (star_n A)\<rbrakk>
-    \<Longrightarrow> x \<in> a \<equiv> eventually (\<lambda>n. X n \<in> A n) \<U>"
-by (simp only: Iset_star_n)
+  "x \<equiv> star_n X \<Longrightarrow> a \<equiv> Iset (star_n A) \<Longrightarrow> x \<in> a \<equiv> eventually (\<lambda>n. X n \<in> A n) \<U>"
+  by (simp only: Iset_star_n)
 
 lemma transfer_Collect [transfer_intro]:
-  "\<lbrakk>\<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>\<rbrakk>
-    \<Longrightarrow> Collect p \<equiv> Iset (star_n (\<lambda>n. Collect (P n)))"
-by (simp add: atomize_eq set_eq_iff all_star_eq Iset_star_n)
+  "(\<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>) \<Longrightarrow>
+    Collect p \<equiv> Iset (star_n (\<lambda>n. Collect (P n)))"
+  by (simp add: atomize_eq set_eq_iff all_star_eq Iset_star_n)
 
 lemma transfer_set_eq [transfer_intro]:
-  "\<lbrakk>a \<equiv> Iset (star_n A); b \<equiv> Iset (star_n B)\<rbrakk>
-    \<Longrightarrow> a = b \<equiv> eventually (\<lambda>n. A n = B n) \<U>"
-by (simp only: set_eq_iff transfer_all transfer_iff transfer_mem)
+  "a \<equiv> Iset (star_n A) \<Longrightarrow> b \<equiv> Iset (star_n B) \<Longrightarrow> a = b \<equiv> eventually (\<lambda>n. A n = B n) \<U>"
+  by (simp only: set_eq_iff transfer_all transfer_iff transfer_mem)
 
 lemma transfer_ball [transfer_intro]:
-  "\<lbrakk>a \<equiv> Iset (star_n A); \<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>\<rbrakk>
-    \<Longrightarrow> \<forall>x\<in>a. p x \<equiv> eventually (\<lambda>n. \<forall>x\<in>A n. P n x) \<U>"
-by (simp only: Ball_def transfer_all transfer_imp transfer_mem)
+  "a \<equiv> Iset (star_n A) \<Longrightarrow> (\<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>) \<Longrightarrow>
+    \<forall>x\<in>a. p x \<equiv> eventually (\<lambda>n. \<forall>x\<in>A n. P n x) \<U>"
+  by (simp only: Ball_def transfer_all transfer_imp transfer_mem)
 
 lemma transfer_bex [transfer_intro]:
-  "\<lbrakk>a \<equiv> Iset (star_n A); \<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>\<rbrakk>
-    \<Longrightarrow> \<exists>x\<in>a. p x \<equiv> eventually (\<lambda>n. \<exists>x\<in>A n. P n x) \<U>"
-by (simp only: Bex_def transfer_ex transfer_conj transfer_mem)
+  "a \<equiv> Iset (star_n A) \<Longrightarrow> (\<And>X. p (star_n X) \<equiv> eventually (\<lambda>n. P n (X n)) \<U>) \<Longrightarrow>
+    \<exists>x\<in>a. p x \<equiv> eventually (\<lambda>n. \<exists>x\<in>A n. P n x) \<U>"
+  by (simp only: Bex_def transfer_ex transfer_conj transfer_mem)
 
-lemma transfer_Iset [transfer_intro]:
-  "\<lbrakk>a \<equiv> star_n A\<rbrakk> \<Longrightarrow> Iset a \<equiv> Iset (star_n (\<lambda>n. A n))"
-by simp
+lemma transfer_Iset [transfer_intro]: "a \<equiv> star_n A \<Longrightarrow> Iset a \<equiv> Iset (star_n (\<lambda>n. A n))"
+  by simp
+
 
 text \<open>Nonstandard extensions of sets.\<close>
 
-definition
-  starset :: "'a set \<Rightarrow> 'a star set" ("*s* _" [80] 80) where
-  "starset A = Iset (star_of A)"
+definition starset :: "'a set \<Rightarrow> 'a star set" ("*s* _" [80] 80)
+  where "starset A = Iset (star_of A)"
 
 declare starset_def [transfer_unfold]
 
-lemma starset_mem: "(star_of x \<in> *s* A) = (x \<in> A)"
-by (transfer, rule refl)
+lemma starset_mem: "star_of x \<in> *s* A \<longleftrightarrow> x \<in> A"
+  by transfer (rule refl)
 
 lemma starset_UNIV: "*s* (UNIV::'a set) = (UNIV::'a star set)"
-by (transfer UNIV_def, rule refl)
+  by (transfer UNIV_def) (rule refl)
 
 lemma starset_empty: "*s* {} = {}"
-by (transfer empty_def, rule refl)
+  by (transfer empty_def) (rule refl)
 
 lemma starset_insert: "*s* (insert x A) = insert (star_of x) ( *s* A)"
-by (transfer insert_def Un_def, rule refl)
+  by (transfer insert_def Un_def) (rule refl)
 
 lemma starset_Un: "*s* (A \<union> B) = *s* A \<union> *s* B"
-by (transfer Un_def, rule refl)
+  by (transfer Un_def) (rule refl)
 
 lemma starset_Int: "*s* (A \<inter> B) = *s* A \<inter> *s* B"
-by (transfer Int_def, rule refl)
+  by (transfer Int_def) (rule refl)
 
 lemma starset_Compl: "*s* -A = -( *s* A)"
-by (transfer Compl_eq, rule refl)
+  by (transfer Compl_eq) (rule refl)
 
 lemma starset_diff: "*s* (A - B) = *s* A - *s* B"
-by (transfer set_diff_eq, rule refl)
+  by (transfer set_diff_eq) (rule refl)
 
 lemma starset_image: "*s* (f ` A) = ( *f* f) ` ( *s* A)"
-by (transfer image_def, rule refl)
+  by (transfer image_def) (rule refl)
 
 lemma starset_vimage: "*s* (f -` A) = ( *f* f) -` ( *s* A)"
-by (transfer vimage_def, rule refl)
+  by (transfer vimage_def) (rule refl)
 
-lemma starset_subset: "( *s* A \<subseteq> *s* B) = (A \<subseteq> B)"
-by (transfer subset_eq, rule refl)
+lemma starset_subset: "( *s* A \<subseteq> *s* B) \<longleftrightarrow> A \<subseteq> B"
+  by (transfer subset_eq) (rule refl)
 
-lemma starset_eq: "( *s* A = *s* B) = (A = B)"
-by (transfer, rule refl)
+lemma starset_eq: "( *s* A = *s* B) \<longleftrightarrow> A = B"
+  by transfer (rule refl)
 
 lemmas starset_simps [simp] =
   starset_mem     starset_UNIV
@@ -425,127 +395,77 @@
 
 instantiation star :: (zero) zero
 begin
-
-definition
-  star_zero_def:    "0 \<equiv> star_of 0"
-
-instance ..
-
+  definition star_zero_def: "0 \<equiv> star_of 0"
+  instance ..
 end
 
 instantiation star :: (one) one
 begin
-
-definition
-  star_one_def:     "1 \<equiv> star_of 1"
-
-instance ..
-
+  definition star_one_def: "1 \<equiv> star_of 1"
+  instance ..
 end
 
 instantiation star :: (plus) plus
 begin
-
-definition
-  star_add_def:     "(op +) \<equiv> *f2* (op +)"
-
-instance ..
-
+  definition star_add_def: "(op +) \<equiv> *f2* (op +)"
+  instance ..
 end
 
 instantiation star :: (times) times
 begin
-
-definition
-  star_mult_def:    "(op *) \<equiv> *f2* (op *)"
-
-instance ..
-
+  definition star_mult_def: "(op *) \<equiv> *f2* (op *)"
+  instance ..
 end
 
 instantiation star :: (uminus) uminus
 begin
-
-definition
-  star_minus_def:   "uminus \<equiv> *f* uminus"
-
-instance ..
-
+  definition star_minus_def: "uminus \<equiv> *f* uminus"
+  instance ..
 end
 
 instantiation star :: (minus) minus
 begin
-
-definition
-  star_diff_def:    "(op -) \<equiv> *f2* (op -)"
-
-instance ..
-
+  definition star_diff_def: "(op -) \<equiv> *f2* (op -)"
+  instance ..
 end
 
 instantiation star :: (abs) abs
 begin
-
-definition
-  star_abs_def:     "abs \<equiv> *f* abs"
-
-instance ..
-
+  definition star_abs_def: "abs \<equiv> *f* abs"
+  instance ..
 end
 
 instantiation star :: (sgn) sgn
 begin
-
-definition
-  star_sgn_def:     "sgn \<equiv> *f* sgn"
-
-instance ..
-
+  definition star_sgn_def: "sgn \<equiv> *f* sgn"
+  instance ..
 end
 
 instantiation star :: (divide) divide
 begin
-
-definition
-  star_divide_def:  "divide \<equiv> *f2* divide"
-
-instance ..
-
+  definition star_divide_def:  "divide \<equiv> *f2* divide"
+  instance ..
 end
 
 instantiation star :: (inverse) inverse
 begin
-
-definition
-  star_inverse_def: "inverse \<equiv> *f* inverse"
-
-instance ..
-
+  definition star_inverse_def: "inverse \<equiv> *f* inverse"
+  instance ..
 end
 
 instance star :: (Rings.dvd) Rings.dvd ..
 
 instantiation star :: (modulo) modulo
 begin
-
-definition
-  star_mod_def:     "(op mod) \<equiv> *f2* (op mod)"
-
-instance ..
-
+  definition star_mod_def: "(op mod) \<equiv> *f2* (op mod)"
+  instance ..
 end
 
 instantiation star :: (ord) ord
 begin
-
-definition
-  star_le_def:      "(op \<le>) \<equiv> *p2* (op \<le>)"
-
-definition
-  star_less_def:    "(op <) \<equiv> *p2* (op <)"
-
-instance ..
-
+  definition star_le_def: "(op \<le>) \<equiv> *p2* (op \<le>)"
+  definition star_less_def: "(op <) \<equiv> *p2* (op <)"
+  instance ..
 end
 
 lemmas star_class_defs [transfer_unfold] =
@@ -555,37 +475,38 @@
   star_le_def       star_less_def     star_abs_def       star_sgn_def
   star_mod_def
 
-text \<open>Class operations preserve standard elements\<close>
+
+text \<open>Class operations preserve standard elements.\<close>
 
 lemma Standard_zero: "0 \<in> Standard"
-by (simp add: star_zero_def)
+  by (simp add: star_zero_def)
 
 lemma Standard_one: "1 \<in> Standard"
-by (simp add: star_one_def)
+  by (simp add: star_one_def)
 
-lemma Standard_add: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x + y \<in> Standard"
-by (simp add: star_add_def)
+lemma Standard_add: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> x + y \<in> Standard"
+  by (simp add: star_add_def)
 
-lemma Standard_diff: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x - y \<in> Standard"
-by (simp add: star_diff_def)
+lemma Standard_diff: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> x - y \<in> Standard"
+  by (simp add: star_diff_def)
 
 lemma Standard_minus: "x \<in> Standard \<Longrightarrow> - x \<in> Standard"
-by (simp add: star_minus_def)
+  by (simp add: star_minus_def)
 
-lemma Standard_mult: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x * y \<in> Standard"
-by (simp add: star_mult_def)
+lemma Standard_mult: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> x * y \<in> Standard"
+  by (simp add: star_mult_def)
 
-lemma Standard_divide: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x / y \<in> Standard"
-by (simp add: star_divide_def)
+lemma Standard_divide: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> x / y \<in> Standard"
+  by (simp add: star_divide_def)
 
 lemma Standard_inverse: "x \<in> Standard \<Longrightarrow> inverse x \<in> Standard"
-by (simp add: star_inverse_def)
+  by (simp add: star_inverse_def)
 
 lemma Standard_abs: "x \<in> Standard \<Longrightarrow> \<bar>x\<bar> \<in> Standard"
-by (simp add: star_abs_def)
+  by (simp add: star_abs_def)
 
-lemma Standard_mod: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x mod y \<in> Standard"
-by (simp add: star_mod_def)
+lemma Standard_mod: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> x mod y \<in> Standard"
+  by (simp add: star_mod_def)
 
 lemmas Standard_simps [simp] =
   Standard_zero  Standard_one
@@ -593,41 +514,44 @@
   Standard_mult  Standard_divide  Standard_inverse
   Standard_abs   Standard_mod
 
-text \<open>@{term star_of} preserves class operations\<close>
+
+text \<open>@{term star_of} preserves class operations.\<close>
 
 lemma star_of_add: "star_of (x + y) = star_of x + star_of y"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma star_of_diff: "star_of (x - y) = star_of x - star_of y"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma star_of_minus: "star_of (-x) = - star_of x"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma star_of_mult: "star_of (x * y) = star_of x * star_of y"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma star_of_divide: "star_of (x / y) = star_of x / star_of y"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma star_of_inverse: "star_of (inverse x) = inverse (star_of x)"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma star_of_mod: "star_of (x mod y) = star_of x mod star_of y"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma star_of_abs: "star_of \<bar>x\<bar> = \<bar>star_of x\<bar>"
-by transfer (rule refl)
+  by transfer (rule refl)
 
-text \<open>@{term star_of} preserves numerals\<close>
+
+text \<open>@{term star_of} preserves numerals.\<close>
 
 lemma star_of_zero: "star_of 0 = 0"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma star_of_one: "star_of 1 = 1"
-by transfer (rule refl)
+  by transfer (rule refl)
 
-text \<open>@{term star_of} preserves orderings\<close>
+
+text \<open>@{term star_of} preserves orderings.\<close>
 
 lemma star_of_less: "(star_of x < star_of y) = (x < y)"
 by transfer (rule refl)
@@ -638,7 +562,8 @@
 lemma star_of_eq: "(star_of x = star_of y) = (x = y)"
 by transfer (rule refl)
 
-text\<open>As above, for 0\<close>
+
+text \<open>As above, for \<open>0\<close>.\<close>
 
 lemmas star_of_0_less = star_of_less [of 0, simplified star_of_zero]
 lemmas star_of_0_le   = star_of_le   [of 0, simplified star_of_zero]
@@ -648,7 +573,8 @@
 lemmas star_of_le_0   = star_of_le   [of _ 0, simplified star_of_zero]
 lemmas star_of_eq_0   = star_of_eq   [of _ 0, simplified star_of_zero]
 
-text\<open>As above, for 1\<close>
+
+text \<open>As above, for \<open>1\<close>.\<close>
 
 lemmas star_of_1_less = star_of_less [of 1, simplified star_of_one]
 lemmas star_of_1_le   = star_of_le   [of 1, simplified star_of_one]
@@ -669,36 +595,27 @@
   star_of_1_less  star_of_1_le    star_of_1_eq
   star_of_less_1  star_of_le_1    star_of_eq_1
 
+
 subsection \<open>Ordering and lattice classes\<close>
 
 instance star :: (order) order
-apply (intro_classes)
-apply (transfer, rule less_le_not_le)
-apply (transfer, rule order_refl)
-apply (transfer, erule (1) order_trans)
-apply (transfer, erule (1) order_antisym)
-done
+  apply intro_classes
+     apply (transfer, rule less_le_not_le)
+    apply (transfer, rule order_refl)
+   apply (transfer, erule (1) order_trans)
+  apply (transfer, erule (1) order_antisym)
+  done
 
 instantiation star :: (semilattice_inf) semilattice_inf
 begin
-
-definition
-  star_inf_def [transfer_unfold]: "inf \<equiv> *f2* inf"
-
-instance
-  by (standard; transfer) auto
-
+  definition star_inf_def [transfer_unfold]: "inf \<equiv> *f2* inf"
+  instance by (standard; transfer) auto
 end
 
 instantiation star :: (semilattice_sup) semilattice_sup
 begin
-
-definition
-  star_sup_def [transfer_unfold]: "sup \<equiv> *f2* sup"
-
-instance
-  by (standard; transfer) auto
-
+  definition star_sup_def [transfer_unfold]: "sup \<equiv> *f2* sup"
+  instance by (standard; transfer) auto
 end
 
 instance star :: (lattice) lattice ..
@@ -706,102 +623,98 @@
 instance star :: (distrib_lattice) distrib_lattice
   by (standard; transfer) (auto simp add: sup_inf_distrib1)
 
-lemma Standard_inf [simp]:
-  "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> inf x y \<in> Standard"
-by (simp add: star_inf_def)
+lemma Standard_inf [simp]: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> inf x y \<in> Standard"
+  by (simp add: star_inf_def)
 
-lemma Standard_sup [simp]:
-  "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> sup x y \<in> Standard"
-by (simp add: star_sup_def)
+lemma Standard_sup [simp]: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> sup x y \<in> Standard"
+  by (simp add: star_sup_def)
 
 lemma star_of_inf [simp]: "star_of (inf x y) = inf (star_of x) (star_of y)"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma star_of_sup [simp]: "star_of (sup x y) = sup (star_of x) (star_of y)"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 instance star :: (linorder) linorder
-by (intro_classes, transfer, rule linorder_linear)
+  by (intro_classes, transfer, rule linorder_linear)
 
 lemma star_max_def [transfer_unfold]: "max = *f2* max"
-apply (rule ext, rule ext)
-apply (unfold max_def, transfer, fold max_def)
-apply (rule refl)
-done
+  apply (rule ext, rule ext)
+  apply (unfold max_def, transfer, fold max_def)
+  apply (rule refl)
+  done
 
 lemma star_min_def [transfer_unfold]: "min = *f2* min"
-apply (rule ext, rule ext)
-apply (unfold min_def, transfer, fold min_def)
-apply (rule refl)
-done
+  apply (rule ext, rule ext)
+  apply (unfold min_def, transfer, fold min_def)
+  apply (rule refl)
+  done
 
-lemma Standard_max [simp]:
-  "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> max x y \<in> Standard"
-by (simp add: star_max_def)
+lemma Standard_max [simp]: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> max x y \<in> Standard"
+  by (simp add: star_max_def)
 
-lemma Standard_min [simp]:
-  "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> min x y \<in> Standard"
-by (simp add: star_min_def)
+lemma Standard_min [simp]: "x \<in> Standard \<Longrightarrow> y \<in> Standard \<Longrightarrow> min x y \<in> Standard"
+  by (simp add: star_min_def)
 
 lemma star_of_max [simp]: "star_of (max x y) = max (star_of x) (star_of y)"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma star_of_min [simp]: "star_of (min x y) = min (star_of x) (star_of y)"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 
 subsection \<open>Ordered group classes\<close>
 
 instance star :: (semigroup_add) semigroup_add
-by (intro_classes, transfer, rule add.assoc)
+  by (intro_classes, transfer, rule add.assoc)
 
 instance star :: (ab_semigroup_add) ab_semigroup_add
-by (intro_classes, transfer, rule add.commute)
+  by (intro_classes, transfer, rule add.commute)
 
 instance star :: (semigroup_mult) semigroup_mult
-by (intro_classes, transfer, rule mult.assoc)
+  by (intro_classes, transfer, rule mult.assoc)
 
 instance star :: (ab_semigroup_mult) ab_semigroup_mult
-by (intro_classes, transfer, rule mult.commute)
+  by (intro_classes, transfer, rule mult.commute)
 
 instance star :: (comm_monoid_add) comm_monoid_add
-by (intro_classes, transfer, rule comm_monoid_add_class.add_0)
+  by (intro_classes, transfer, rule comm_monoid_add_class.add_0)
 
 instance star :: (monoid_mult) monoid_mult
-apply (intro_classes)
-apply (transfer, rule mult_1_left)
-apply (transfer, rule mult_1_right)
-done
+  apply intro_classes
+   apply (transfer, rule mult_1_left)
+  apply (transfer, rule mult_1_right)
+  done
 
 instance star :: (power) power ..
 
 instance star :: (comm_monoid_mult) comm_monoid_mult
-by (intro_classes, transfer, rule mult_1)
+  by (intro_classes, transfer, rule mult_1)
 
 instance star :: (cancel_semigroup_add) cancel_semigroup_add
-apply (intro_classes)
-apply (transfer, erule add_left_imp_eq)
-apply (transfer, erule add_right_imp_eq)
-done
+  apply intro_classes
+   apply (transfer, erule add_left_imp_eq)
+  apply (transfer, erule add_right_imp_eq)
+  done
 
 instance star :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add
-by intro_classes (transfer, simp add: diff_diff_eq)+
+  by intro_classes (transfer, simp add: diff_diff_eq)+
 
 instance star :: (cancel_comm_monoid_add) cancel_comm_monoid_add ..
 
 instance star :: (ab_group_add) ab_group_add
-apply (intro_classes)
-apply (transfer, rule left_minus)
-apply (transfer, rule diff_conv_add_uminus)
-done
+  apply intro_classes
+   apply (transfer, rule left_minus)
+  apply (transfer, rule diff_conv_add_uminus)
+  done
 
 instance star :: (ordered_ab_semigroup_add) ordered_ab_semigroup_add
-by (intro_classes, transfer, rule add_left_mono)
+  by (intro_classes, transfer, rule add_left_mono)
 
 instance star :: (ordered_cancel_ab_semigroup_add) ordered_cancel_ab_semigroup_add ..
 
 instance star :: (ordered_ab_semigroup_add_imp_le) ordered_ab_semigroup_add_imp_le
-by (intro_classes, transfer, rule add_le_imp_le_left)
+  by (intro_classes, transfer, rule add_le_imp_le_left)
 
 instance star :: (ordered_comm_monoid_add) ordered_comm_monoid_add ..
 instance star :: (ordered_ab_semigroup_monoid_add_imp_le) ordered_ab_semigroup_monoid_add_imp_le ..
@@ -809,8 +722,7 @@
 instance star :: (ordered_ab_group_add) ordered_ab_group_add ..
 
 instance star :: (ordered_ab_group_add_abs) ordered_ab_group_add_abs
-  by intro_classes (transfer,
-    simp add: abs_ge_self abs_leI abs_triangle_ineq)+
+  by intro_classes (transfer, simp add: abs_ge_self abs_leI abs_triangle_ineq)+
 
 instance star :: (linordered_cancel_ab_semigroup_add) linordered_cancel_ab_semigroup_add ..
 
@@ -909,14 +821,13 @@
 
 instance star :: (linordered_field) linordered_field ..
 
+
 subsection \<open>Power\<close>
 
-lemma star_power_def [transfer_unfold]:
-  "(op ^) \<equiv> \<lambda>x n. ( *f* (\<lambda>x. x ^ n)) x"
+lemma star_power_def [transfer_unfold]: "(op ^) \<equiv> \<lambda>x n. ( *f* (\<lambda>x. x ^ n)) x"
 proof (rule eq_reflection, rule ext, rule ext)
-  fix n :: nat
-  show "\<And>x::'a star. x ^ n = ( *f* (\<lambda>x. x ^ n)) x"
-  proof (induct n)
+  show "x ^ n = ( *f* (\<lambda>x. x ^ n)) x" for n :: nat and x :: "'a star"
+  proof (induct n arbitrary: x)
     case 0
     have "\<And>x::'a star. ( *f* (\<lambda>x. 1)) x = 1"
       by transfer simp
@@ -940,18 +851,17 @@
 
 instance star :: (numeral) numeral ..
 
-lemma star_numeral_def [transfer_unfold]:
-  "numeral k = star_of (numeral k)"
-by (induct k, simp_all only: numeral.simps star_of_one star_of_add)
+lemma star_numeral_def [transfer_unfold]: "numeral k = star_of (numeral k)"
+  by (induct k) (simp_all only: numeral.simps star_of_one star_of_add)
 
 lemma Standard_numeral [simp]: "numeral k \<in> Standard"
-by (simp add: star_numeral_def)
+  by (simp add: star_numeral_def)
 
 lemma star_of_numeral [simp]: "star_of (numeral k) = numeral k"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma star_of_nat_def [transfer_unfold]: "of_nat n = star_of (of_nat n)"
-by (induct n, simp_all)
+  by (induct n) simp_all
 
 lemmas star_of_compare_numeral [simp] =
   star_of_less [of "numeral k", simplified star_of_numeral]
@@ -968,63 +878,66 @@
   star_of_eq   [of _ "- numeral k", simplified star_of_numeral] for k
 
 lemma Standard_of_nat [simp]: "of_nat n \<in> Standard"
-by (simp add: star_of_nat_def)
+  by (simp add: star_of_nat_def)
 
 lemma star_of_of_nat [simp]: "star_of (of_nat n) = of_nat n"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 lemma star_of_int_def [transfer_unfold]: "of_int z = star_of (of_int z)"
-by (rule_tac z=z in int_diff_cases, simp)
+  by (rule int_diff_cases [of z]) simp
 
 lemma Standard_of_int [simp]: "of_int z \<in> Standard"
-by (simp add: star_of_int_def)
+  by (simp add: star_of_int_def)
 
 lemma star_of_of_int [simp]: "star_of (of_int z) = of_int z"
-by transfer (rule refl)
+  by transfer (rule refl)
 
 instance star :: (semiring_char_0) semiring_char_0
 proof
-  have "inj (star_of :: 'a \<Rightarrow> 'a star)" by (rule injI) simp
-  then have "inj (star_of \<circ> of_nat :: nat \<Rightarrow> 'a star)" using inj_of_nat by (rule inj_comp)
-  then show "inj (of_nat :: nat \<Rightarrow> 'a star)" by (simp add: comp_def)
+  have "inj (star_of :: 'a \<Rightarrow> 'a star)"
+    by (rule injI) simp
+  then have "inj (star_of \<circ> of_nat :: nat \<Rightarrow> 'a star)"
+    using inj_of_nat by (rule inj_comp)
+  then show "inj (of_nat :: nat \<Rightarrow> 'a star)"
+    by (simp add: comp_def)
 qed
 
 instance star :: (ring_char_0) ring_char_0 ..
 
 instance star :: (semiring_parity) semiring_parity
-apply intro_classes
-apply(transfer, rule odd_one)
-apply(transfer, erule (1) odd_even_add)
-apply(transfer, erule even_multD)
-apply(transfer, erule odd_ex_decrement)
-done
+  apply intro_classes
+     apply (transfer, rule odd_one)
+    apply (transfer, erule (1) odd_even_add)
+   apply (transfer, erule even_multD)
+  apply (transfer, erule odd_ex_decrement)
+  done
 
 instance star :: (semiring_div_parity) semiring_div_parity
-apply intro_classes
-apply(transfer, rule parity)
-apply(transfer, rule one_mod_two_eq_one)
-apply(transfer, rule zero_not_eq_two)
-done
+  apply intro_classes
+    apply (transfer, rule parity)
+   apply (transfer, rule one_mod_two_eq_one)
+  apply (transfer, rule zero_not_eq_two)
+  done
 
 instantiation star :: (semiring_numeral_div) semiring_numeral_div
 begin
 
 definition divmod_star :: "num \<Rightarrow> num \<Rightarrow> 'a star \<times> 'a star"
-where
-  divmod_star_def: "divmod_star m n = (numeral m div numeral n, numeral m mod numeral n)"
+  where divmod_star_def: "divmod_star m n = (numeral m div numeral n, numeral m mod numeral n)"
 
 definition divmod_step_star :: "num \<Rightarrow> 'a star \<times> 'a star \<Rightarrow> 'a star \<times> 'a star"
-where
-  "divmod_step_star l qr = (let (q, r) = qr
-    in if r \<ge> numeral l then (2 * q + 1, r - numeral l)
-    else (2 * q, r))"
+  where "divmod_step_star l qr =
+    (let (q, r) = qr
+     in if r \<ge> numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))"
 
-instance proof
-  show "divmod m n = (numeral m div numeral n :: 'a star, numeral m mod numeral n)"
-    for m n by (fact divmod_star_def)
-  show "divmod_step l qr = (let (q, r) = qr
-    in if r \<ge> numeral l then (2 * q + 1, r - numeral l)
-    else (2 * q, r))" for l and qr :: "'a star \<times> 'a star"
+instance
+proof
+  show "divmod m n = (numeral m div numeral n :: 'a star, numeral m mod numeral n)" for m n
+    by (fact divmod_star_def)
+  show "divmod_step l qr =
+    (let (q, r) = qr
+     in if r \<ge> numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))"
+    for l and qr :: "'a star \<times> 'a star"
     by (fact divmod_step_star_def)
 qed (transfer,
   fact
@@ -1046,13 +959,13 @@
 subsection \<open>Finite class\<close>
 
 lemma starset_finite: "finite A \<Longrightarrow> *s* A = star_of ` A"
-by (erule finite_induct, simp_all)
+  by (erule finite_induct) simp_all
 
 instance star :: (finite) finite
-apply (intro_classes)
-apply (subst starset_UNIV [symmetric])
-apply (subst starset_finite [OF finite])
-apply (rule finite_imageI [OF finite])
-done
+  apply intro_classes
+  apply (subst starset_UNIV [symmetric])
+  apply (subst starset_finite [OF finite])
+  apply (rule finite_imageI [OF finite])
+  done
 
 end
--- a/src/HOL/Nonstandard_Analysis/document/root.tex	Mon Oct 31 16:26:36 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/document/root.tex	Tue Nov 01 00:44:24 2016 +0100
@@ -1,6 +1,6 @@
 \documentclass[11pt,a4paper]{article}
 \usepackage{graphicx,isabelle,isabellesym,latexsym,textcomp}
-\usepackage{amsmath}
+\usepackage{amsmath,amssymb}
 \usepackage{stmaryrd}
 \usepackage{pdfsetup}
 
--- a/src/HOL/Nonstandard_Analysis/transfer.ML	Mon Oct 31 16:26:36 2016 +0100
+++ b/src/HOL/Nonstandard_Analysis/transfer.ML	Tue Nov 01 00:44:24 2016 +0100
@@ -13,7 +13,7 @@
 structure Transfer_Principle: TRANSFER_PRINCIPLE =
 struct
 
-structure TransferData = Generic_Data
+structure Data = Generic_Data
 (
   type T = {
     intros: thm list,
@@ -55,13 +55,13 @@
 fun transfer_star_tac ctxt =
   let
     fun thm_of (Const (@{const_name Ifun}, _) $ t $ u) = @{thm transfer_Ifun} OF [thm_of t, thm_of u]
-    | thm_of (Const (@{const_name star_of}, _) $ _) = @{thm star_of_def}
-    | thm_of (Const (@{const_name star_n}, _) $ _) = @{thm Pure.reflexive}
-    | thm_of _ = raise MATCH;
+      | thm_of (Const (@{const_name star_of}, _) $ _) = @{thm star_of_def}
+      | thm_of (Const (@{const_name star_n}, _) $ _) = @{thm Pure.reflexive}
+      | thm_of _ = raise MATCH;
 
     fun thm_of_goal (Const (@{const_name Pure.eq}, _) $ t $ (Const (@{const_name star_n}, _) $ _)) =
-      thm_of t
-    | thm_of_goal _ = raise MATCH;
+          thm_of t
+      | thm_of_goal _ = raise MATCH;
   in
     SUBGOAL (fn (t, i) =>
       resolve_tac ctxt [thm_of_goal (strip_all_body t |> Logic.strip_imp_concl)] i
@@ -71,44 +71,43 @@
 
 fun transfer_thm_of ctxt ths t =
   let
-    val {intros,unfolds,refolds,consts} = TransferData.get (Context.Proof ctxt);
+    val {intros,unfolds,refolds,consts} = Data.get (Context.Proof ctxt);
     val meta = Local_Defs.meta_rewrite_rule ctxt;
     val ths' = map meta ths;
     val unfolds' = map meta unfolds and refolds' = map meta refolds;
     val (_$_$t') = Thm.concl_of (Raw_Simplifier.rewrite ctxt true unfolds' (Thm.cterm_of ctxt t))
     val u = unstar_term consts t'
-   val tac =
+    val tac =
       rewrite_goals_tac ctxt (ths' @ refolds' @ unfolds') THEN
       ALLGOALS (Object_Logic.full_atomize_tac ctxt) THEN
       match_tac ctxt [transitive_thm] 1 THEN
       resolve_tac ctxt [@{thm transfer_start}] 1 THEN
       REPEAT_ALL_NEW (resolve_tac ctxt intros ORELSE' transfer_star_tac ctxt) 1 THEN
       match_tac ctxt [reflexive_thm] 1
-  in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end
+  in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end;
 
-fun transfer_tac ctxt ths =
-    SUBGOAL (fn (t, _) =>
-        (fn th =>
-            let
-              val tr = transfer_thm_of ctxt ths t
-              val (_$l$r) = Thm.concl_of tr;
-              val trs = if l aconv r then [] else [tr];
-            in rewrite_goals_tac ctxt trs th end))
+fun transfer_tac ctxt ths = SUBGOAL (fn (t, _) => (fn th =>
+  let
+    val tr = transfer_thm_of ctxt ths t
+    val (_$ l $ r) = Thm.concl_of tr;
+    val trs = if l aconv r then [] else [tr];
+  in rewrite_goals_tac ctxt trs th end));
 
 local
 
-fun map_intros f = TransferData.map
+fun map_intros f = Data.map
   (fn {intros,unfolds,refolds,consts} =>
     {intros=f intros, unfolds=unfolds, refolds=refolds, consts=consts})
 
-fun map_unfolds f = TransferData.map
+fun map_unfolds f = Data.map
   (fn {intros,unfolds,refolds,consts} =>
     {intros=intros, unfolds=f unfolds, refolds=refolds, consts=consts})
 
-fun map_refolds f = TransferData.map
+fun map_refolds f = Data.map
   (fn {intros,unfolds,refolds,consts} =>
     {intros=intros, unfolds=unfolds, refolds=f refolds, consts=consts})
 in
+
 val intro_add = Thm.declaration_attribute (map_intros o Thm.add_thm);
 val intro_del = Thm.declaration_attribute (map_intros o Thm.del_thm);
 
@@ -117,9 +116,10 @@
 
 val refold_add = Thm.declaration_attribute (map_refolds o Thm.add_thm);
 val refold_del = Thm.declaration_attribute (map_refolds o Thm.del_thm);
+
 end
 
-fun add_const c = Context.theory_map (TransferData.map
+fun add_const c = Context.theory_map (Data.map
   (fn {intros,unfolds,refolds,consts} =>
     {intros=intros, unfolds=unfolds, refolds=refolds, consts=c::consts}))