src/HOL/Nonstandard_Analysis/NSA.thy
changeset 62479 716336f19aa9
parent 61982 3af5a06577c7
child 63901 4ce989e962e0
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Nonstandard_Analysis/NSA.thy	Mon Feb 29 22:34:36 2016 +0100
@@ -0,0 +1,2219 @@
+(*  Title:      HOL/Nonstandard_Analysis/NSA.thy
+    Author:     Jacques D. Fleuriot, University of Cambridge
+    Author:     Lawrence C Paulson, University of Cambridge
+*)
+
+section\<open>Infinite Numbers, Infinitesimals, Infinitely Close Relation\<close>
+
+theory NSA
+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
+  Infinitesimal  :: "('a::real_normed_vector) star set" where
+  "Infinitesimal = {x. \<forall>r \<in> Reals. 0 < r --> 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
+  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
+  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
+  monad     :: "'a::real_normed_vector star => '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}"
+
+lemma SReal_def: "\<real> == {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"
+
+lemma Standard_hnorm [simp]: "x \<in> Standard \<Longrightarrow> hnorm x \<in> Standard"
+by (simp add: hnorm_def)
+
+lemma star_of_norm [simp]: "star_of (norm x) = hnorm (star_of x)"
+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_eq_zero [simp]:
+  "\<And>x::'a::real_normed_vector star. (hnorm x = 0) = (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_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_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_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_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) = (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_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_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 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)
+
+lemma hnorm_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 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)
+
+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)
+
+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
+
+subsection\<open>Closure Laws for the Standard Reals\<close>
+
+lemma Reals_minus_iff [simp]: "(-x \<in> \<real>) = (x \<in> \<real>)"
+apply auto
+apply (drule Reals_minus, auto)
+done
+
+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_hypreal_of_real [simp]: "hypreal_of_real x \<in> \<real>"
+by (simp add: Reals_eq_Standard)
+
+lemma SReal_divide_numeral: "r \<in> \<real> ==> 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
+
+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
+
+lemma SReal_UNIV_real: "{x. hypreal_of_real x \<in> \<real>} = (UNIV::real set)"
+by simp
+
+lemma SReal_iff: "(x \<in> \<real>) = (\<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)
+
+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
+
+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_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
+
+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])
+
+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
+
+
+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_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_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_minus_iff: "(-x \<in> HFinite) = (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
+
+lemma SReal_subset_HFinite: "(\<real>::hypreal set) \<subseteq> HFinite"
+by (auto simp add: SReal_def)
+
+lemma HFiniteD: "x \<in> HFinite ==> \<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_hnorm_iff [iff]:
+  "(hnorm (x::hypreal) \<in> HFinite) = (x \<in> HFinite)"
+by (simp add: HFinite_def)
+
+lemma HFinite_numeral [simp]: "numeral w \<in> HFinite"
+unfolding star_numeral_def by (rule HFinite_star_of)
+
+(** As always with numerals, 0 and 1 are special cases **)
+
+lemma HFinite_0 [simp]: "0 \<in> HFinite"
+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)
+
+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 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
+
+
+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 InfinitesimalD:
+      "x \<in> Infinitesimal ==> \<forall>r \<in> Reals. 0 < r --> 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 InfinitesimalD2:
+  "\<lbrakk>x \<in> Infinitesimal; 0 < r\<rbrakk> \<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)
+
+lemma hypreal_sum_of_halves: "x/(2::hypreal) + x/(2::hypreal) = x"
+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_minus_iff [simp]: "(-x:Infinitesimal) = (x:Infinitesimal)"
+by (simp add: Infinitesimal_def)
+
+lemma Infinitesimal_hnorm_iff:
+  "(hnorm x \<in> Infinitesimal) = (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_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)
+
+lemma Infinitesimal_diff:
+     "[| x \<in> Infinitesimal;  y \<in> Infinitesimal |] ==> 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_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_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
+
+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
+
+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 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
+
+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 HInfiniteI: "(\<And>r. r \<in> \<real> \<Longrightarrow> r < hnorm x) \<Longrightarrow> x \<in> HInfinite"
+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 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 hypreal_add_zero_less_le_mono: "[|r < x; (0::hypreal) \<le> y|] ==> r < x+y"
+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_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_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_minus_iff: "(-x \<in> HInfinite) = (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_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 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)
+
+lemma not_Infinitesimal_not_zero: "x \<notin> Infinitesimal ==> x \<noteq> 0"
+by auto
+
+lemma not_Infinitesimal_not_zero2: "x \<in> HFinite - Infinitesimal ==> x \<noteq> 0"
+by auto
+
+lemma HFinite_diff_Infinitesimal_hrabs:
+  "(x::hypreal) \<in> HFinite - Infinitesimal ==> \<bar>x\<bar> \<in> HFinite - Infinitesimal"
+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_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 hrabs_le_Infinitesimal:
+     "[| e \<in> Infinitesimal; \<bar>x::hypreal\<bar> \<le> e |] ==> x \<in> Infinitesimal"
+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 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)
+
+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)
+
+
+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 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 hrealpow_hyperpow_Infinitesimal_iff:
+     "(x ^ n \<in> Infinitesimal) = (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 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
+
+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 HFinite_Infinitesimal_not_zero: "x \<in> HFinite-Infinitesimal ==> 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
+
+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_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_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])
+
+
+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 approx_minus_iff: " (x \<approx> y) = (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_refl [iff]: "x \<approx> x"
+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 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_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_trans2: "[| r \<approx> x; s \<approx> x |] ==> 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_reorient: "(x \<approx> y) = (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.*)
+simproc_setup approx_reorient_simproc
+  ("0 \<approx> x" | "1 \<approx> y" | "numeral w \<approx> z" | "- 1 \<approx> y" | "- numeral w \<approx> r") =
+\<open>
+  let val rule = @{thm approx_reorient} RS eq_reflection
+      fun proc phi ss ct =
+        case Thm.term_of ct of
+          _ $ t $ u => if can HOLogic.dest_number u then NONE
+            else if can HOLogic.dest_number t then SOME rule else NONE
+        | _ => NONE
+  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 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 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 approx_add: "[| a \<approx> b; c \<approx> d |] ==> 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)
+  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_minus2: "-a \<approx> -b ==> 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_add_minus: "[| a \<approx> b; c \<approx> d |] ==> 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"
+  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_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:
+  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_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_eq_imp: "a = b ==> 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 bex_Infinitesimal_iff: "(\<exists>y \<in> Infinitesimal. x - z = y) = (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 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_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_self2: "y \<in> Infinitesimal ==> 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_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_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 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_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_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_mono2: "b \<approx> c ==> 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_right_iff [simp]: "(b + a \<approx> c + a) = (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_star_of_HFinite: "x \<approx> star_of D ==> 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 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_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_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:
+  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_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_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::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::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::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::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"
+proof (unfold approx_def)
+  assume "x - y \<in> Infinitesimal"
+  hence 1: "hnorm (x - y) \<in> Infinitesimal"
+    by (simp only: Infinitesimal_hnorm_iff)
+  moreover have 2: "(0::real star) \<in> Infinitesimal"
+    by (rule Infinitesimal_zero)
+  moreover have 3: "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)"
+    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"
+    by (simp only: Infinitesimal_hrabs_iff)
+qed
+
+
+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_SReal2:
+     "(y::hypreal) \<in> Infinitesimal ==> \<forall>r \<in> Reals. 0 < r --> y < r"
+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_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_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
+
+lemma SReal_Infinitesimal_zero:
+  "[| (x::hypreal) \<in> \<real>; x \<in> Infinitesimal|] ==> x = 0"
+by (cut_tac SReal_Int_Infinitesimal_zero, 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 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)
+
+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_HFinite_diff_Infinitesimal:
+     "x \<noteq> 0 ==> 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])
+
+(*again: 1 is a special case, but not 0 this time*)
+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
+
+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 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
+
+(*The premise y\<noteq>0 is essential; otherwise x/y =0 and we lose the
+  HFinite premise.*)
+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
+
+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
+
+(*------------------------------------------------------------------
+       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>
+
+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 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 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
+
+(*And also for 0 \<approx> #nn and 1 \<approx> #nn, #nn \<approx> 0 and #nn \<approx> 1.*)
+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)
+
+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_zero_iff [simp]: "(star_of k \<approx> 0) = (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 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)
+
+
+subsection\<open>Existence of Unique Real Infinitely Close\<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_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_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_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)
+
+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_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) ==> \<exists>Y. isLub \<real> P Y"
+by (auto simp add: isLub_def leastP_def isUb_def)
+
+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
+
+(* lemma about lubs *)
+
+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::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_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
+
+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_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 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_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_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
+
+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
+
+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
+
+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_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_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
+
+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
+
+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
+
+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)
+
+
+
+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
+
+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
+
+text\<open>There is a unique real infinitely close\<close>
+lemma st_part_Ex1: "x \<in> HFinite ==> EX! 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
+
+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
+
+lemma HFinite_not_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
+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 HInfinite_HFinite_disj: "x \<in> HInfinite | 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 HFinite_HInfinite_iff: "(x \<in> HFinite) = (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)
+
+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_inverse2:
+  fixes x :: "'a::real_normed_div_algebra star"
+  shows "x \<in> HFinite - Infinitesimal ==> inverse x \<in> HFinite"
+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
+
+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
+
+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
+
+(*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
+
+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
+
+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
+
+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
+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 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 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_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 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:
+     "[| 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_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 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 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
+
+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
+
+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_zero_gt_one:
+  "[| (x::hypreal) \<in> HInfinite; 0 < x |] ==> 1 < x"
+by (auto intro: HInfinite_gt_SReal)
+
+
+lemma not_HInfinite_one [simp]: "1 \<notin> HInfinite"
+apply (simp (no_asm) add: HInfinite_HFinite_iff)
+done
+
+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)
+
+
+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)
+
+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 mem_monad_iff: "(u \<in> monad x) = (-u \<in> monad (-x))"
+by (simp add: monad_def)
+
+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 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_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 mem_monad_self [simp]: "x \<in> monad x"
+by (simp add: monad_def)
+
+
+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_monad2: "x \<approx> y ==> {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 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_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 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 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 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_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 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_less_zero:
+     "[|(x::hypreal) < 0; x \<notin> Infinitesimal; x \<approx> y|] ==> y < 0"
+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_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_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_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 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
+
+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
+
+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
+*)
+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
+
+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
+
+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
+
+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
+
+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)
+
+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
+
+(*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_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 HFinite_square_cancel [simp]:
+  "(x::hypreal)*x + y*y \<in> HFinite ==> x*x \<in> HFinite"
+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 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 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
+
+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 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
+
+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 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
+
+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 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 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
+
+
+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_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_HFinite: "x \<in> HFinite ==> 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_SReal_eq: "x \<in> \<real> ==> 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])
+
+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 approx_st_eq:
+  assumes x: "x \<in> HFinite" and y: "y \<in> HFinite" and xy: "x \<approx> y"
+  shows "st x = st y"
+proof -
+  have "st x \<approx> x" "st y \<approx> y" "st x \<in> \<real>" "st y \<in> \<real>"
+    by (simp_all add: st_approx_self st_SReal x y)
+  with xy show ?thesis
+    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_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_SReal2:
+     "[| x \<in> \<real>; e \<in> Infinitesimal |] ==> 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 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_numeral [simp]: "st (numeral w) = numeral w"
+by (rule Reals_numeral [THEN st_SReal_eq])
+
+lemma st_neg_numeral [simp]: "st (- numeral w) = - numeral w"
+proof -
+  from Reals_numeral have "numeral w \<in> \<real>" .
+  then have "- numeral w \<in> \<real>" by simp
+  with st_SReal_eq show ?thesis .
+qed
+
+lemma st_0 [simp]: "st 0 = 0"
+by (simp add: st_SReal_eq)
+
+lemma st_1 [simp]: "st 1 = 1"
+by (simp add: st_SReal_eq)
+
+lemma st_neg_1 [simp]: "st (- 1) = - 1"
+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)
+
+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)
+
+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)
+
+lemma st_Infinitesimal: "x \<in> Infinitesimal ==> st x = 0"
+by (simp add: st_unique mem_infmal_iff)
+
+lemma st_not_Infinitesimal: "st(x) \<noteq> 0 ==> 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_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_idempotent [simp]: "x \<in> HFinite ==> 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
+
+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
+
+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_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_ge: "[| x \<le> 0;  x \<in> HFinite |] ==> 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
+
+
+
+subsection \<open>Alternative Definitions using Free Ultrafilter\<close>
+
+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
+
+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
+
+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)
+
+subsubsection \<open>@{term HInfinite}\<close>
+
+lemma lemma_Compl_eq: "- {n. u < norm (f n)} = {n. norm (f n) \<le> u}"
+by auto
+
+lemma lemma_Compl_eq2: "- {n. norm (f n) < u} = {n. u \<le> norm (f n)}"
+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
+
+(*-------------------------------------
+  Exclude this type of sets from free
+  ultrafilter for Infinite numbers!
+ -------------------------------------*)
+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
+
+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
+
+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_HIa: "{n. u < norm (X n)} Int {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)
+proof -
+  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
+    by (simp add: eventually_False FreeUltrafilterNat.proper)
+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)
+
+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 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
+
+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
+
+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)
+
+(*------------------------------------------------------------------------
+         Infinitesimals as smaller than 1/n for all n::nat (> 0)
+ ------------------------------------------------------------------------*)
+
+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_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
+  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
+
+
+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
+
+
+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>
+
+lemma Suc_Un_eq: "{n. n < Suc m} = {n. n < m} Un {n. n = m}"
+by (auto simp add: less_Suc_eq)
+
+(*-------------------------------------------
+  Prove that any segment is finite and hence cannot belong to FreeUltrafilterNat
+ -------------------------------------------*)
+
+lemma finite_real_of_nat_segment: "finite {n::nat. real n < real (m::nat)}"
+  by (auto intro: finite_Collect_less_nat)
+
+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
+
+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 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)
+
+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
+
+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)
+
+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
+
+(*--------------------------------------------------------------
+ The complement of {n. \<bar>real n\<bar> \<le> u} =
+ {n. u < \<bar>real n\<bar>} is in FreeUltrafilterNat
+ by property of (free) ultrafilters
+ --------------------------------------------------------------*)
+
+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)
+
+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
+
+(*-----------------------------------------------
+       Epsilon is a member of Infinitesimal
+ -----------------------------------------------*)
+
+lemma Infinitesimal_epsilon [simp]: "\<epsilon> \<in> Infinitesimal"
+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])
+
+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.
+ -----------------------------------------------------------------------*)
+
+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
+
+lemma finite_inverse_real_of_posnat_gt_real:
+     "0 < u ==> 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
+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)
+
+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 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)
+
+(*--------------------------------------------------------------
+    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)
+
+
+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])
+
+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)
+
+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
+ -----------------------------------------------------*)
+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)
+
+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)
+
+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)
+
+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)
+
+end