move nonstandard analysis theories to NSA directory
authorhuffman
Thu, 03 Jul 2008 17:47:22 +0200
changeset 27468 0783dd1dc13d
parent 27467 c0c4a6bd2cbc
child 27469 00ee6d56de8b
move nonstandard analysis theories to NSA directory
src/HOL/NSA/CLim.thy
src/HOL/NSA/CStar.thy
src/HOL/NSA/Examples/NSPrimes.thy
src/HOL/NSA/Examples/ROOT.ML
src/HOL/NSA/Filter.thy
src/HOL/NSA/HDeriv.thy
src/HOL/NSA/HLim.thy
src/HOL/NSA/HLog.thy
src/HOL/NSA/HSEQ.thy
src/HOL/NSA/HSeries.thy
src/HOL/NSA/HTranscendental.thy
src/HOL/NSA/HyperDef.thy
src/HOL/NSA/HyperNat.thy
src/HOL/NSA/Hypercomplex.thy
src/HOL/NSA/Hyperreal.thy
src/HOL/NSA/NSA.thy
src/HOL/NSA/NSCA.thy
src/HOL/NSA/NSComplex.thy
src/HOL/NSA/NatStar.thy
src/HOL/NSA/ROOT.ML
src/HOL/NSA/Star.thy
src/HOL/NSA/StarDef.thy
src/HOL/NSA/document/root.tex
src/HOL/NSA/hypreal_arith.ML
src/HOL/NSA/transfer.ML
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/CLim.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,187 @@
+(*  Title       : CLim.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 2001 University of Edinburgh
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
+*)
+
+header{*Limits, Continuity and Differentiation for Complex Functions*}
+
+theory CLim
+imports CStar
+begin
+
+(*not in simpset?*)
+declare hypreal_epsilon_not_zero [simp]
+
+(*??generalize*)
+lemma lemma_complex_mult_inverse_squared [simp]:
+     "x \<noteq> (0::complex) \<Longrightarrow> (x * inverse(x) ^ 2) = inverse x"
+by (simp add: numeral_2_eq_2)
+
+text{*Changing the quantified variable. Install earlier?*}
+lemma all_shift: "(\<forall>x::'a::comm_ring_1. P x) = (\<forall>x. P (x-a))";
+apply auto 
+apply (drule_tac x="x+a" in spec) 
+apply (simp add: diff_minus add_assoc) 
+done
+
+lemma complex_add_minus_iff [simp]: "(x + - a = (0::complex)) = (x=a)"
+by (simp add: diff_eq_eq diff_minus [symmetric])
+
+lemma complex_add_eq_0_iff [iff]: "(x+y = (0::complex)) = (y = -x)"
+apply auto
+apply (drule sym [THEN diff_eq_eq [THEN iffD2]], auto)
+done
+
+
+subsection{*Limit of Complex to Complex Function*}
+
+lemma NSLIM_Re: "f -- a --NS> L ==> (%x. Re(f x)) -- a --NS> Re(L)"
+by (simp add: NSLIM_def starfunC_approx_Re_Im_iff 
+              hRe_hcomplex_of_complex)
+
+lemma NSLIM_Im: "f -- a --NS> L ==> (%x. Im(f x)) -- a --NS> Im(L)"
+by (simp add: NSLIM_def starfunC_approx_Re_Im_iff 
+              hIm_hcomplex_of_complex)
+
+(** get this result easily now **)
+lemma LIM_Re: "f -- a --> L ==> (%x. Re(f x)) -- a --> Re(L)"
+by (simp add: LIM_NSLIM_iff NSLIM_Re)
+
+lemma LIM_Im: "f -- a --> L ==> (%x. Im(f x)) -- a --> Im(L)"
+by (simp add: LIM_NSLIM_iff NSLIM_Im)
+
+lemma LIM_cnj: "f -- a --> L ==> (%x. cnj (f x)) -- a --> cnj L"
+by (simp add: LIM_def complex_cnj_diff [symmetric])
+
+lemma LIM_cnj_iff: "((%x. cnj (f x)) -- a --> cnj L) = (f -- a --> L)"
+by (simp add: LIM_def complex_cnj_diff [symmetric])
+
+lemma starfun_norm: "( *f* (\<lambda>x. norm (f x))) = (\<lambda>x. hnorm (( *f* f) x))"
+by transfer (rule refl)
+
+lemma star_of_Re [simp]: "star_of (Re x) = hRe (star_of x)"
+by transfer (rule refl)
+
+lemma star_of_Im [simp]: "star_of (Im x) = hIm (star_of x)"
+by transfer (rule refl)
+
+text""
+(** another equivalence result **)
+lemma NSCLIM_NSCRLIM_iff:
+   "(f -- x --NS> L) = ((%y. cmod(f y - L)) -- x --NS> 0)"
+by (simp add: NSLIM_def starfun_norm
+    approx_approx_zero_iff [symmetric] approx_minus_iff [symmetric])
+
+(** much, much easier standard proof **)
+lemma CLIM_CRLIM_iff: "(f -- x --> L) = ((%y. cmod(f y - L)) -- x --> 0)"
+by (simp add: LIM_def)
+
+(* so this is nicer nonstandard proof *)
+lemma NSCLIM_NSCRLIM_iff2:
+     "(f -- x --NS> L) = ((%y. cmod(f y - L)) -- x --NS> 0)"
+by (simp add: LIM_NSLIM_iff [symmetric] CLIM_CRLIM_iff)
+
+lemma NSLIM_NSCRLIM_Re_Im_iff:
+     "(f -- a --NS> L) = ((%x. Re(f x)) -- a --NS> Re(L) &
+                            (%x. Im(f x)) -- a --NS> Im(L))"
+apply (auto intro: NSLIM_Re NSLIM_Im)
+apply (auto simp add: NSLIM_def starfun_Re starfun_Im)
+apply (auto dest!: spec)
+apply (simp add: hcomplex_approx_iff)
+done
+
+lemma LIM_CRLIM_Re_Im_iff:
+     "(f -- a --> L) = ((%x. Re(f x)) -- a --> Re(L) &
+                         (%x. Im(f x)) -- a --> Im(L))"
+by (simp add: LIM_NSLIM_iff NSLIM_NSCRLIM_Re_Im_iff)
+
+
+subsection{*Continuity*}
+
+lemma NSLIM_isContc_iff:
+     "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)"
+by (rule NSLIM_h_iff)
+
+subsection{*Functions from Complex to Reals*}
+
+lemma isNSContCR_cmod [simp]: "isNSCont cmod (a)"
+by (auto intro: approx_hnorm
+         simp add: starfunCR_cmod hcmod_hcomplex_of_complex [symmetric] 
+                    isNSCont_def)
+
+lemma isContCR_cmod [simp]: "isCont cmod (a)"
+by (simp add: isNSCont_isCont_iff [symmetric])
+
+lemma isCont_Re: "isCont f a ==> isCont (%x. Re (f x)) a"
+by (simp add: isCont_def LIM_Re)
+
+lemma isCont_Im: "isCont f a ==> isCont (%x. Im (f x)) a"
+by (simp add: isCont_def LIM_Im)
+
+subsection{* Differentiation of Natural Number Powers*}
+
+lemma CDERIV_pow [simp]:
+     "DERIV (%x. x ^ n) x :> (complex_of_real (real n)) * (x ^ (n - Suc 0))"
+apply (induct n)
+apply (drule_tac [2] DERIV_ident [THEN DERIV_mult])
+apply (auto simp add: left_distrib real_of_nat_Suc)
+apply (case_tac "n")
+apply (auto simp add: mult_ac add_commute)
+done
+
+text{*Nonstandard version*}
+lemma NSCDERIV_pow:
+     "NSDERIV (%x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - 1))"
+by (simp add: NSDERIV_DERIV_iff)
+
+text{*Can't relax the premise @{term "x \<noteq> 0"}: it isn't continuous at zero*}
+lemma NSCDERIV_inverse:
+     "(x::complex) \<noteq> 0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ 2))"
+unfolding numeral_2_eq_2
+by (rule NSDERIV_inverse)
+
+lemma CDERIV_inverse:
+     "(x::complex) \<noteq> 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ 2))"
+unfolding numeral_2_eq_2
+by (rule DERIV_inverse)
+
+
+subsection{*Derivative of Reciprocals (Function @{term inverse})*}
+
+lemma CDERIV_inverse_fun:
+     "[| DERIV f x :> d; f(x) \<noteq> (0::complex) |]
+      ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))"
+unfolding numeral_2_eq_2
+by (rule DERIV_inverse_fun)
+
+lemma NSCDERIV_inverse_fun:
+     "[| NSDERIV f x :> d; f(x) \<noteq> (0::complex) |]
+      ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))"
+unfolding numeral_2_eq_2
+by (rule NSDERIV_inverse_fun)
+
+
+subsection{* Derivative of Quotient*}
+
+lemma CDERIV_quotient:
+     "[| DERIV f x :> d; DERIV g x :> e; g(x) \<noteq> (0::complex) |]
+       ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ 2)"
+unfolding numeral_2_eq_2
+by (rule DERIV_quotient)
+
+lemma NSCDERIV_quotient:
+     "[| NSDERIV f x :> d; NSDERIV g x :> e; g(x) \<noteq> (0::complex) |]
+       ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ 2)"
+unfolding numeral_2_eq_2
+by (rule NSDERIV_quotient)
+
+
+subsection{*Caratheodory Formulation of Derivative at a Point: Standard Proof*}
+
+lemma CARAT_CDERIVD:
+     "(\<forall>z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l
+      ==> NSDERIV f x :> l"
+by clarify (rule CARAT_DERIVD)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/CStar.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,59 @@
+(*  Title       : CStar.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 2001 University of Edinburgh
+*)
+
+header{*Star-transforms in NSA, Extending Sets of Complex Numbers
+      and Complex Functions*}
+
+theory CStar
+imports NSCA
+begin
+
+subsection{*Properties of the *-Transform Applied to Sets of Reals*}
+
+lemma STARC_hcomplex_of_complex_Int:
+     "*s* X Int SComplex = hcomplex_of_complex ` X"
+by (auto simp add: Standard_def)
+
+lemma lemma_not_hcomplexA:
+     "x \<notin> hcomplex_of_complex ` A ==> \<forall>y \<in> A. x \<noteq> hcomplex_of_complex y"
+by auto
+
+subsection{*Theorems about Nonstandard Extensions of Functions*}
+
+lemma starfunC_hcpow: "!!Z. ( *f* (%z. z ^ n)) Z = Z pow hypnat_of_nat n"
+by transfer (rule refl)
+
+lemma starfunCR_cmod: "*f* cmod = hcmod"
+by transfer (rule refl)
+
+subsection{*Internal Functions - Some Redundancy With *f* Now*}
+
+(** subtraction: ( *fn) - ( *gn) = *(fn - gn) **)
+(*
+lemma starfun_n_diff:
+   "( *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_diff)
+done
+*)
+(** composition: ( *fn) o ( *gn) = *(fn o gn) **)
+
+lemma starfun_Re: "( *f* (\<lambda>x. Re (f x))) = (\<lambda>x. hRe (( *f* f) x))"
+by transfer (rule refl)
+
+lemma starfun_Im: "( *f* (\<lambda>x. Im (f x))) = (\<lambda>x. hIm (( *f* f) x))"
+by transfer (rule refl)
+
+lemma starfunC_eq_Re_Im_iff:
+    "(( *f* f) x = z) = ((( *f* (%x. Re(f x))) x = hRe (z)) &
+                          (( *f* (%x. Im(f x))) x = hIm (z)))"
+by (simp add: hcomplex_hRe_hIm_cancel_iff starfun_Re starfun_Im)
+
+lemma starfunC_approx_Re_Im_iff:
+    "(( *f* f) x @= z) = ((( *f* (%x. Re(f x))) x @= hRe (z)) &
+                            (( *f* (%x. Im(f x))) x @= hIm (z)))"
+by (simp add: hcomplex_approx_iff starfun_Re starfun_Im)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/Examples/NSPrimes.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,343 @@
+(*  Title       : NSPrimes.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 2002 University of Edinburgh
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
+*)
+
+header{*The Nonstandard Primes as an Extension of the Prime Numbers*}
+
+theory NSPrimes
+imports "~~/src/HOL/NumberTheory/Factorization" Hyperreal
+begin
+
+text{*These can be used to derive an alternative proof of the infinitude of
+primes by considering a property of nonstandard sets.*}
+
+definition
+  hdvd  :: "[hypnat, hypnat] => bool"       (infixl "hdvd" 50) where
+  [transfer_unfold]: "(M::hypnat) hdvd N = ( *p2* (op dvd)) M N"
+
+definition
+  starprime :: "hypnat set" where
+  [transfer_unfold]: "starprime = ( *s* {p. prime p})"
+
+definition
+  choicefun :: "'a set => 'a" where
+  "choicefun E = (@x. \<exists>X \<in> Pow(E) -{{}}. x : X)"
+
+consts injf_max :: "nat => ('a::{order} set) => 'a"
+primrec
+  injf_max_zero: "injf_max 0 E = choicefun E"
+  injf_max_Suc:  "injf_max (Suc n) E = choicefun({e. e:E & injf_max n E < e})"
+
+
+lemma dvd_by_all: "\<forall>M. \<exists>N. 0 < N & (\<forall>m. 0 < m & (m::nat) <= M --> m dvd N)"
+apply (rule allI)
+apply (induct_tac "M", auto)
+apply (rule_tac x = "N * (Suc n) " in exI)
+apply (safe, force)
+apply (drule le_imp_less_or_eq, erule disjE)
+apply (force intro!: dvd_mult2)
+apply (force intro!: dvd_mult)
+done
+
+lemmas dvd_by_all2 = dvd_by_all [THEN spec, standard]
+
+lemma hypnat_of_nat_le_zero_iff: "(hypnat_of_nat n <= 0) = (n = 0)"
+by (transfer, simp)
+declare hypnat_of_nat_le_zero_iff [simp]
+
+
+(* Goldblatt: Exercise 5.11(2) - p. 57 *)
+lemma hdvd_by_all: "\<forall>M. \<exists>N. 0 < N & (\<forall>m. 0 < m & (m::hypnat) <= M --> m hdvd N)"
+by (transfer, rule dvd_by_all)
+
+lemmas hdvd_by_all2 = hdvd_by_all [THEN spec, standard]
+
+(* Goldblatt: Exercise 5.11(2) - p. 57 *)
+lemma hypnat_dvd_all_hypnat_of_nat:
+     "\<exists>(N::hypnat). 0 < N & (\<forall>n \<in> -{0::nat}. hypnat_of_nat(n) hdvd N)"
+apply (cut_tac hdvd_by_all)
+apply (drule_tac x = whn in spec, auto)
+apply (rule exI, auto)
+apply (drule_tac x = "hypnat_of_nat n" in spec)
+apply (auto simp add: linorder_not_less star_of_eq_0)
+done
+
+
+text{*The nonstandard extension of the set prime numbers consists of precisely
+those hypernaturals exceeding 1 that have no nontrivial factors*}
+
+(* Goldblatt: Exercise 5.11(3a) - p 57  *)
+lemma starprime:
+  "starprime = {p. 1 < p & (\<forall>m. m hdvd p --> m = 1 | m = p)}"
+by (transfer, auto simp add: prime_def)
+
+lemma prime_two:  "prime 2"
+apply (unfold prime_def, auto)
+apply (frule dvd_imp_le)
+apply (auto dest: dvd_0_left)
+apply (case_tac m, simp, arith)
+done
+declare prime_two [simp]
+
+(* proof uses course-of-value induction *)
+lemma prime_factor_exists [rule_format]: "Suc 0 < n --> (\<exists>k. prime k & k dvd n)"
+apply (rule_tac n = n in nat_less_induct, auto)
+apply (case_tac "prime n")
+apply (rule_tac x = n in exI, auto)
+apply (drule conjI [THEN not_prime_ex_mk], auto)
+apply (drule_tac x = m in spec, auto)
+apply (rule_tac x = ka in exI)
+apply (auto intro: dvd_mult2)
+done
+
+(* Goldblatt Exercise 5.11(3b) - p 57  *)
+lemma hyperprime_factor_exists [rule_format]:
+  "!!n. 1 < n ==> (\<exists>k \<in> starprime. k hdvd n)"
+by (transfer, simp add: prime_factor_exists)
+
+(* Goldblatt Exercise 3.10(1) - p. 29 *)
+lemma NatStar_hypnat_of_nat: "finite A ==> *s* A = hypnat_of_nat ` A"
+by (rule starset_finite)
+
+
+subsection{*Another characterization of infinite set of natural numbers*}
+
+lemma finite_nat_set_bounded: "finite N ==> \<exists>n. (\<forall>i \<in> N. i<(n::nat))"
+apply (erule_tac F = N in finite_induct, auto)
+apply (rule_tac x = "Suc n + x" in exI, auto)
+done
+
+lemma finite_nat_set_bounded_iff: "finite N = (\<exists>n. (\<forall>i \<in> N. i<(n::nat)))"
+by (blast intro: finite_nat_set_bounded bounded_nat_set_is_finite)
+
+lemma not_finite_nat_set_iff: "(~ finite N) = (\<forall>n. \<exists>i \<in> N. n <= (i::nat))"
+by (auto simp add: finite_nat_set_bounded_iff not_less)
+
+lemma bounded_nat_set_is_finite2: "(\<forall>i \<in> N. i<=(n::nat)) ==> finite N"
+apply (rule finite_subset)
+ apply (rule_tac [2] finite_atMost, auto)
+done
+
+lemma finite_nat_set_bounded2: "finite N ==> \<exists>n. (\<forall>i \<in> N. i<=(n::nat))"
+apply (erule_tac F = N in finite_induct, auto)
+apply (rule_tac x = "n + x" in exI, auto)
+done
+
+lemma finite_nat_set_bounded_iff2: "finite N = (\<exists>n. (\<forall>i \<in> N. i<=(n::nat)))"
+by (blast intro: finite_nat_set_bounded2 bounded_nat_set_is_finite2)
+
+lemma not_finite_nat_set_iff2: "(~ finite N) = (\<forall>n. \<exists>i \<in> N. n < (i::nat))"
+by (auto simp add: finite_nat_set_bounded_iff2 not_le)
+
+
+subsection{*An injective function cannot define an embedded natural number*}
+
+lemma lemma_infinite_set_singleton: "\<forall>m n. m \<noteq> n --> f n \<noteq> f m
+      ==>  {n. f n = N} = {} |  (\<exists>m. {n. f n = N} = {m})"
+apply auto
+apply (drule_tac x = x in spec, auto)
+apply (subgoal_tac "\<forall>n. (f n = f x) = (x = n) ")
+apply auto
+done
+
+lemma inj_fun_not_hypnat_in_SHNat:
+  assumes inj_f: "inj (f::nat=>nat)"
+  shows "starfun f whn \<notin> Nats"
+proof
+  from inj_f have inj_f': "inj (starfun f)"
+    by (transfer inj_on_def Ball_def UNIV_def)
+  assume "starfun f whn \<in> Nats"
+  then obtain N where N: "starfun f whn = hypnat_of_nat N"
+    by (auto simp add: Nats_def)
+  hence "\<exists>n. starfun f n = hypnat_of_nat N" ..
+  hence "\<exists>n. f n = N" by transfer
+  then obtain n where n: "f n = N" ..
+  hence "starfun f (hypnat_of_nat n) = hypnat_of_nat N"
+    by transfer
+  with N have "starfun f whn = starfun f (hypnat_of_nat n)"
+    by simp
+  with inj_f' have "whn = hypnat_of_nat n"
+    by (rule injD)
+  thus "False"
+    by (simp add: whn_neq_hypnat_of_nat)
+qed
+
+lemma range_subset_mem_starsetNat:
+   "range f <= A ==> starfun f whn \<in> *s* A"
+apply (rule_tac x="whn" in spec)
+apply (transfer, auto)
+done
+
+(*--------------------------------------------------------------------------------*)
+(* Gleason Proposition 11-5.5. pg 149, pg 155 (ex. 3) and pg. 360                 *)
+(* Let E be a nonvoid ordered set with no maximal elements (note: effectively an  *)
+(* infinite set if we take E = N (Nats)). Then there exists an order-preserving   *)
+(* injection from N to E. Of course, (as some doofus will undoubtedly point out!  *)
+(* :-)) can use notion of least element in proof (i.e. no need for choice) if     *)
+(* dealing with nats as we have well-ordering property                            *)
+(*--------------------------------------------------------------------------------*)
+
+lemma lemmaPow3: "E \<noteq> {} ==> \<exists>x. \<exists>X \<in> (Pow E - {{}}). x: X"
+by auto
+
+lemma choicefun_mem_set: "E \<noteq> {} ==> choicefun E \<in> E"
+apply (unfold choicefun_def)
+apply (rule lemmaPow3 [THEN someI2_ex], auto)
+done
+declare choicefun_mem_set [simp]
+
+lemma injf_max_mem_set: "[| E \<noteq>{}; \<forall>x. \<exists>y \<in> E. x < y |] ==> injf_max n E \<in> E"
+apply (induct_tac "n", force)
+apply (simp (no_asm) add: choicefun_def)
+apply (rule lemmaPow3 [THEN someI2_ex], auto)
+done
+
+lemma injf_max_order_preserving: "\<forall>x. \<exists>y \<in> E. x < y ==> injf_max n E < injf_max (Suc n) E"
+apply (simp (no_asm) add: choicefun_def)
+apply (rule lemmaPow3 [THEN someI2_ex], auto)
+done
+
+lemma injf_max_order_preserving2: "\<forall>x. \<exists>y \<in> E. x < y
+      ==> \<forall>n m. m < n --> injf_max m E < injf_max n E"
+apply (rule allI)
+apply (induct_tac "n", auto)
+apply (simp (no_asm) add: choicefun_def)
+apply (rule lemmaPow3 [THEN someI2_ex])
+apply (auto simp add: less_Suc_eq)
+apply (drule_tac x = m in spec)
+apply (drule subsetD, auto)
+apply (drule_tac x = "injf_max m E" in order_less_trans, auto)
+done
+
+lemma inj_injf_max: "\<forall>x. \<exists>y \<in> E. x < y ==> inj (%n. injf_max n E)"
+apply (rule inj_onI)
+apply (rule ccontr, auto)
+apply (drule injf_max_order_preserving2)
+apply (metis linorder_antisym_conv3 order_less_le)
+done
+
+lemma infinite_set_has_order_preserving_inj:
+     "[| (E::('a::{order} set)) \<noteq> {}; \<forall>x. \<exists>y \<in> E. x < y |]
+      ==> \<exists>f. range f <= E & inj (f::nat => 'a) & (\<forall>m. f m < f(Suc m))"
+apply (rule_tac x = "%n. injf_max n E" in exI, safe)
+apply (rule injf_max_mem_set)
+apply (rule_tac [3] inj_injf_max)
+apply (rule_tac [4] injf_max_order_preserving, auto)
+done
+
+text{*Only need the existence of an injective function from N to A for proof*}
+
+lemma hypnat_infinite_has_nonstandard:
+     "~ finite A ==> hypnat_of_nat ` A < ( *s* A)"
+apply auto
+apply (subgoal_tac "A \<noteq> {}")
+prefer 2 apply force
+apply (drule infinite_set_has_order_preserving_inj)
+apply (erule not_finite_nat_set_iff2 [THEN iffD1], auto)
+apply (drule inj_fun_not_hypnat_in_SHNat)
+apply (drule range_subset_mem_starsetNat)
+apply (auto simp add: SHNat_eq)
+done
+
+lemma starsetNat_eq_hypnat_of_nat_image_finite: "*s* A =  hypnat_of_nat ` A ==> finite A"
+apply (rule ccontr)
+apply (auto dest: hypnat_infinite_has_nonstandard)
+done
+
+lemma finite_starsetNat_iff: "( *s* A = hypnat_of_nat ` A) = (finite A)"
+by (blast intro!: starsetNat_eq_hypnat_of_nat_image_finite NatStar_hypnat_of_nat)
+
+lemma hypnat_infinite_has_nonstandard_iff: "(~ finite A) = (hypnat_of_nat ` A < *s* A)"
+apply (rule iffI)
+apply (blast intro!: hypnat_infinite_has_nonstandard)
+apply (auto simp add: finite_starsetNat_iff [symmetric])
+done
+
+subsection{*Existence of Infinitely Many Primes: a Nonstandard Proof*}
+
+lemma lemma_not_dvd_hypnat_one: "~ (\<forall>n \<in> - {0}. hypnat_of_nat n hdvd 1)"
+apply auto
+apply (rule_tac x = 2 in bexI)
+apply (transfer, auto)
+done
+declare lemma_not_dvd_hypnat_one [simp]
+
+lemma lemma_not_dvd_hypnat_one2: "\<exists>n \<in> - {0}. ~ hypnat_of_nat n hdvd 1"
+apply (cut_tac lemma_not_dvd_hypnat_one)
+apply (auto simp del: lemma_not_dvd_hypnat_one)
+done
+declare lemma_not_dvd_hypnat_one2 [simp]
+
+(* not needed here *)
+lemma hypnat_gt_zero_gt_one:
+  "!!N. [| 0 < (N::hypnat); N \<noteq> 1 |] ==> 1 < N"
+by (transfer, simp)
+
+lemma hypnat_add_one_gt_one:
+    "!!N. 0 < N ==> 1 < (N::hypnat) + 1"
+by (transfer, simp)
+
+lemma zero_not_prime: "\<not> prime 0"
+apply safe
+apply (drule prime_g_zero, auto)
+done
+declare zero_not_prime [simp]
+
+lemma hypnat_of_nat_zero_not_prime: "hypnat_of_nat 0 \<notin> starprime"
+by (transfer, simp)
+declare hypnat_of_nat_zero_not_prime [simp]
+
+lemma hypnat_zero_not_prime:
+   "0 \<notin> starprime"
+by (cut_tac hypnat_of_nat_zero_not_prime, simp)
+declare hypnat_zero_not_prime [simp]
+
+lemma one_not_prime: "\<not> prime 1"
+apply safe
+apply (drule prime_g_one, auto)
+done
+declare one_not_prime [simp]
+
+lemma one_not_prime2: "\<not> prime(Suc 0)"
+apply safe
+apply (drule prime_g_one, auto)
+done
+declare one_not_prime2 [simp]
+
+lemma hypnat_of_nat_one_not_prime: "hypnat_of_nat 1 \<notin> starprime"
+by (transfer, simp)
+declare hypnat_of_nat_one_not_prime [simp]
+
+lemma hypnat_one_not_prime: "1 \<notin> starprime"
+by (cut_tac hypnat_of_nat_one_not_prime, simp)
+declare hypnat_one_not_prime [simp]
+
+lemma hdvd_diff: "!!k m n. [| k hdvd m; k hdvd n |] ==> k hdvd (m - n)"
+by (transfer, rule dvd_diff)
+
+lemma dvd_one_eq_one: "x dvd (1::nat) ==> x = 1"
+by (unfold dvd_def, auto)
+
+lemma hdvd_one_eq_one: "!!x. x hdvd 1 ==> x = 1"
+by (transfer, rule dvd_one_eq_one)
+
+theorem not_finite_prime: "~ finite {p. prime p}"
+apply (rule hypnat_infinite_has_nonstandard_iff [THEN iffD2])
+apply (cut_tac hypnat_dvd_all_hypnat_of_nat)
+apply (erule exE)
+apply (erule conjE)
+apply (subgoal_tac "1 < N + 1")
+prefer 2 apply (blast intro: hypnat_add_one_gt_one)
+apply (drule hyperprime_factor_exists)
+apply auto
+apply (subgoal_tac "k \<notin> hypnat_of_nat ` {p. prime p}")
+apply (force simp add: starprime_def, safe)
+apply (drule_tac x = x in bspec)
+apply (rule ccontr, simp)
+apply (drule hdvd_diff, assumption)
+apply (auto dest: hdvd_one_eq_one)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/Examples/ROOT.ML	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,1 @@
+use_thy "NSPrimes";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/Filter.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,409 @@
+(*  Title       : Filter.thy
+    ID          : $Id$
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
+    Conversion to locales by Brian Huffman, 2005
+*) 
+
+header {* Filters and Ultrafilters *}
+
+theory Filter
+imports "~~/src/HOL/Library/Zorn" "~~/src/HOL/Library/Infinite_Set"
+begin
+
+subsection {* Definitions and basic properties *}
+
+subsubsection {* Filters *}
+
+locale filter =
+  fixes F :: "'a set set"
+  assumes UNIV [iff]:  "UNIV \<in> F"
+  assumes empty [iff]: "{} \<notin> F"
+  assumes Int:         "\<lbrakk>u \<in> F; v \<in> F\<rbrakk> \<Longrightarrow> u \<inter> v \<in> F"
+  assumes subset:      "\<lbrakk>u \<in> F; u \<subseteq> v\<rbrakk> \<Longrightarrow> v \<in> F"
+
+lemma (in filter) memD: "A \<in> F \<Longrightarrow> - A \<notin> F"
+proof
+  assume "A \<in> F" and "- A \<in> F"
+  hence "A \<inter> (- A) \<in> F" by (rule Int)
+  thus "False" by simp
+qed
+
+lemma (in filter) not_memI: "- A \<in> F \<Longrightarrow> A \<notin> F"
+by (drule memD, simp)
+
+lemma (in filter) Int_iff: "(x \<inter> y \<in> F) = (x \<in> F \<and> y \<in> F)"
+by (auto elim: subset intro: Int)
+
+subsubsection {* Ultrafilters *}
+
+locale ultrafilter = filter +
+  assumes ultra: "A \<in> F \<or> - A \<in> F"
+
+lemma (in ultrafilter) memI: "- A \<notin> F \<Longrightarrow> A \<in> F"
+by (cut_tac ultra [of A], simp)
+
+lemma (in ultrafilter) not_memD: "A \<notin> F \<Longrightarrow> - A \<in> F"
+by (rule memI, simp)
+
+lemma (in ultrafilter) not_mem_iff: "(A \<notin> F) = (- A \<in> F)"
+by (rule iffI [OF not_memD not_memI])
+
+lemma (in ultrafilter) Compl_iff: "(- A \<in> F) = (A \<notin> F)"
+by (rule iffI [OF not_memI not_memD])
+
+lemma (in ultrafilter) Un_iff: "(x \<union> y \<in> F) = (x \<in> F \<or> y \<in> F)"
+ apply (rule iffI)
+  apply (erule contrapos_pp)
+  apply (simp add: Int_iff not_mem_iff)
+ apply (auto elim: subset)
+done
+
+subsubsection {* Free Ultrafilters *}
+
+locale freeultrafilter = ultrafilter +
+  assumes infinite: "A \<in> F \<Longrightarrow> infinite A"
+
+lemma (in freeultrafilter) finite: "finite A \<Longrightarrow> A \<notin> F"
+by (erule contrapos_pn, erule infinite)
+
+lemma (in freeultrafilter) singleton: "{x} \<notin> F"
+by (rule finite, simp)
+
+lemma (in freeultrafilter) insert_iff [simp]: "(insert x A \<in> F) = (A \<in> F)"
+apply (subst insert_is_Un)
+apply (subst Un_iff)
+apply (simp add: singleton)
+done
+
+lemma (in freeultrafilter) filter: "filter F" by unfold_locales
+
+lemma (in freeultrafilter) ultrafilter: "ultrafilter F"
+  by unfold_locales
+
+
+subsection {* Collect properties *}
+
+lemma (in filter) Collect_ex:
+  "({n. \<exists>x. P n x} \<in> F) = (\<exists>X. {n. P n (X n)} \<in> F)"
+proof
+  assume "{n. \<exists>x. P n x} \<in> F"
+  hence "{n. P n (SOME x. P n x)} \<in> F"
+    by (auto elim: someI subset)
+  thus "\<exists>X. {n. P n (X n)} \<in> F" by fast
+next
+  show "\<exists>X. {n. P n (X n)} \<in> F \<Longrightarrow> {n. \<exists>x. P n x} \<in> F"
+    by (auto elim: subset)
+qed
+
+lemma (in filter) Collect_conj:
+  "({n. P n \<and> Q n} \<in> F) = ({n. P n} \<in> F \<and> {n. Q n} \<in> F)"
+by (subst Collect_conj_eq, rule Int_iff)
+
+lemma (in ultrafilter) Collect_not:
+  "({n. \<not> P n} \<in> F) = ({n. P n} \<notin> F)"
+by (subst Collect_neg_eq, rule Compl_iff)
+
+lemma (in ultrafilter) Collect_disj:
+  "({n. P n \<or> Q n} \<in> F) = ({n. P n} \<in> F \<or> {n. Q n} \<in> F)"
+by (subst Collect_disj_eq, rule Un_iff)
+
+lemma (in ultrafilter) Collect_all:
+  "({n. \<forall>x. P n x} \<in> F) = (\<forall>X. {n. P n (X n)} \<in> F)"
+apply (rule Not_eq_iff [THEN iffD1])
+apply (simp add: Collect_not [symmetric])
+apply (rule Collect_ex)
+done
+
+subsection {* Maximal filter = Ultrafilter *}
+
+text {*
+   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"}
+*}
+text {*
+  Lemmas 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.
+*}
+
+lemma extend_lemma1: "UNIV \<in> F \<Longrightarrow> A \<in> {X. \<exists>f\<in>F. A \<inter> f \<subseteq> X}"
+by blast
+
+lemma extend_lemma2: "F \<subseteq> {X. \<exists>f\<in>F. A \<inter> f \<subseteq> X}"
+by blast
+
+lemma (in filter) extend_filter:
+assumes A: "- A \<notin> F"
+shows "filter {X. \<exists>f\<in>F. A \<inter> f \<subseteq> X}" (is "filter ?X")
+proof (rule filter.intro)
+  show "UNIV \<in> ?X" by blast
+next
+  show "{} \<notin> ?X"
+  proof (clarify)
+    fix f assume f: "f \<in> F" and Af: "A \<inter> f \<subseteq> {}"
+    from Af have fA: "f \<subseteq> - A" by blast
+    from f fA have "- A \<in> F" by (rule subset)
+    with A show "False" by simp
+  qed
+next
+  fix u and v
+  assume u: "u \<in> ?X" and v: "v \<in> ?X"
+  from u obtain f where f: "f \<in> F" and Af: "A \<inter> f \<subseteq> u" by blast
+  from v obtain g where g: "g \<in> F" and Ag: "A \<inter> g \<subseteq> v" by blast
+  from f g have fg: "f \<inter> g \<in> F" by (rule Int)
+  from Af Ag have Afg: "A \<inter> (f \<inter> g) \<subseteq> u \<inter> v" by blast
+  from fg Afg show "u \<inter> v \<in> ?X" by blast
+next
+  fix u and v
+  assume uv: "u \<subseteq> v" and u: "u \<in> ?X"
+  from u obtain f where f: "f \<in> F" and Afu: "A \<inter> f \<subseteq> u" by blast
+  from Afu uv have Afv: "A \<inter> f \<subseteq> v" by blast
+  from f Afv have "\<exists>f\<in>F. A \<inter> f \<subseteq> v" by blast
+  thus "v \<in> ?X" by simp
+qed
+
+lemma (in filter) max_filter_ultrafilter:
+assumes max: "\<And>G. \<lbrakk>filter G; F \<subseteq> G\<rbrakk> \<Longrightarrow> F = G"
+shows "ultrafilter_axioms F"
+proof (rule ultrafilter_axioms.intro)
+  fix A show "A \<in> F \<or> - A \<in> F"
+  proof (rule disjCI)
+    let ?X = "{X. \<exists>f\<in>F. A \<inter> f \<subseteq> X}"
+    assume AF: "- A \<notin> F"
+    from AF have X: "filter ?X" by (rule extend_filter)
+    from UNIV have AX: "A \<in> ?X" by (rule extend_lemma1)
+    have FX: "F \<subseteq> ?X" by (rule extend_lemma2)
+    from X FX have "F = ?X" by (rule max)
+    with AX show "A \<in> F" by simp
+  qed
+qed
+
+lemma (in ultrafilter) max_filter:
+assumes G: "filter G" and sub: "F \<subseteq> G" shows "F = G"
+proof
+  show "F \<subseteq> G" using sub .
+  show "G \<subseteq> F"
+  proof
+    fix A assume A: "A \<in> G"
+    from G A have "- A \<notin> G" by (rule filter.memD)
+    with sub have B: "- A \<notin> F" by blast
+    thus "A \<in> F" by (rule memI)
+  qed
+qed
+
+subsection {* Ultrafilter Theorem *}
+
+text "A locale makes proof of ultrafilter Theorem more modular"
+locale (open) UFT =
+  fixes   frechet :: "'a set set"
+  and     superfrechet :: "'a set set set"
+
+  assumes infinite_UNIV: "infinite (UNIV :: 'a set)"
+
+  defines frechet_def: "frechet \<equiv> {A. finite (- A)}"
+  and     superfrechet_def: "superfrechet \<equiv> {G. filter G \<and> frechet \<subseteq> G}"
+
+lemma (in UFT) superfrechetI:
+  "\<lbrakk>filter G; frechet \<subseteq> G\<rbrakk> \<Longrightarrow> G \<in> superfrechet"
+by (simp add: superfrechet_def)
+
+lemma (in UFT) superfrechetD1:
+  "G \<in> superfrechet \<Longrightarrow> filter G"
+by (simp add: superfrechet_def)
+
+lemma (in UFT) superfrechetD2:
+  "G \<in> superfrechet \<Longrightarrow> frechet \<subseteq> G"
+by (simp add: superfrechet_def)
+
+text {* A few properties of free filters *}
+
+lemma filter_cofinite:
+assumes inf: "infinite (UNIV :: 'a set)"
+shows "filter {A:: 'a set. finite (- A)}" (is "filter ?F")
+proof (rule filter.intro)
+  show "UNIV \<in> ?F" by simp
+next
+  show "{} \<notin> ?F" using inf by simp
+next
+  fix u v assume "u \<in> ?F" and "v \<in> ?F"
+  thus "u \<inter> v \<in> ?F" by simp
+next
+  fix u v assume uv: "u \<subseteq> v" and u: "u \<in> ?F"
+  from uv have vu: "- v \<subseteq> - u" by simp
+  from u show "v \<in> ?F"
+    by (simp add: finite_subset [OF vu])
+qed
+
+text {*
+   We prove: 1. Existence of maximal filter i.e. ultrafilter;
+             2. Freeness property i.e ultrafilter is free.
+             Use a locale to prove various lemmas and then 
+             export main result: The ultrafilter Theorem
+*}
+
+lemma (in UFT) filter_frechet: "filter frechet"
+by (unfold frechet_def, rule filter_cofinite [OF infinite_UNIV])
+
+lemma (in UFT) frechet_in_superfrechet: "frechet \<in> superfrechet"
+by (rule superfrechetI [OF filter_frechet subset_refl])
+
+lemma (in UFT) lemma_mem_chain_filter:
+  "\<lbrakk>c \<in> chain superfrechet; x \<in> c\<rbrakk> \<Longrightarrow> filter x"
+by (unfold chain_def superfrechet_def, blast)
+
+
+subsubsection {* Unions of chains of superfrechets *}
+
+text "In this section we prove that superfrechet is closed
+with respect to unions of non-empty chains. We must show
+  1) Union of a chain is a filter,
+  2) Union of a chain contains frechet.
+
+Number 2 is trivial, but 1 requires us to prove all the filter rules."
+
+lemma (in UFT) Union_chain_UNIV:
+"\<lbrakk>c \<in> chain superfrechet; c \<noteq> {}\<rbrakk> \<Longrightarrow> UNIV \<in> \<Union>c"
+proof -
+  assume 1: "c \<in> chain superfrechet" and 2: "c \<noteq> {}"
+  from 2 obtain x where 3: "x \<in> c" by blast
+  from 1 3 have "filter x" by (rule lemma_mem_chain_filter)
+  hence "UNIV \<in> x" by (rule filter.UNIV)
+  with 3 show "UNIV \<in> \<Union>c" by blast
+qed
+
+lemma (in UFT) Union_chain_empty:
+"c \<in> chain superfrechet \<Longrightarrow> {} \<notin> \<Union>c"
+proof
+  assume 1: "c \<in> chain superfrechet" and 2: "{} \<in> \<Union>c"
+  from 2 obtain x where 3: "x \<in> c" and 4: "{} \<in> x" ..
+  from 1 3 have "filter x" by (rule lemma_mem_chain_filter)
+  hence "{} \<notin> x" by (rule filter.empty)
+  with 4 show "False" by simp
+qed
+
+lemma (in UFT) Union_chain_Int:
+"\<lbrakk>c \<in> chain superfrechet; u \<in> \<Union>c; v \<in> \<Union>c\<rbrakk> \<Longrightarrow> u \<inter> v \<in> \<Union>c"
+proof -
+  assume c: "c \<in> chain superfrechet"
+  assume "u \<in> \<Union>c"
+    then obtain x where ux: "u \<in> x" and xc: "x \<in> c" ..
+  assume "v \<in> \<Union>c"
+    then obtain y where vy: "v \<in> y" and yc: "y \<in> c" ..
+  from c xc yc have "x \<subseteq> y \<or> y \<subseteq> x" by (rule chainD)
+  with xc yc have xyc: "x \<union> y \<in> c"
+    by (auto simp add: Un_absorb1 Un_absorb2)
+  with c have fxy: "filter (x \<union> y)" by (rule lemma_mem_chain_filter)
+  from ux have uxy: "u \<in> x \<union> y" by simp
+  from vy have vxy: "v \<in> x \<union> y" by simp
+  from fxy uxy vxy have "u \<inter> v \<in> x \<union> y" by (rule filter.Int)
+  with xyc show "u \<inter> v \<in> \<Union>c" ..
+qed
+
+lemma (in UFT) Union_chain_subset:
+"\<lbrakk>c \<in> chain superfrechet; u \<in> \<Union>c; u \<subseteq> v\<rbrakk> \<Longrightarrow> v \<in> \<Union>c"
+proof -
+  assume c: "c \<in> chain superfrechet"
+     and u: "u \<in> \<Union>c" and uv: "u \<subseteq> v"
+  from u obtain x where ux: "u \<in> x" and xc: "x \<in> c" ..
+  from c xc have fx: "filter x" by (rule lemma_mem_chain_filter)
+  from fx ux uv have vx: "v \<in> x" by (rule filter.subset)
+  with xc show "v \<in> \<Union>c" ..
+qed
+
+lemma (in UFT) Union_chain_filter:
+assumes chain: "c \<in> chain superfrechet" and nonempty: "c \<noteq> {}"
+shows "filter (\<Union>c)"
+proof (rule filter.intro)
+  show "UNIV \<in> \<Union>c" using chain nonempty by (rule Union_chain_UNIV)
+next
+  show "{} \<notin> \<Union>c" using chain by (rule Union_chain_empty)
+next
+  fix u v assume "u \<in> \<Union>c" and "v \<in> \<Union>c"
+  with chain show "u \<inter> v \<in> \<Union>c" by (rule Union_chain_Int)
+next
+  fix u v assume "u \<in> \<Union>c" and "u \<subseteq> v"
+  with chain show "v \<in> \<Union>c" by (rule Union_chain_subset)
+qed
+
+lemma (in UFT) lemma_mem_chain_frechet_subset:
+  "\<lbrakk>c \<in> chain superfrechet; x \<in> c\<rbrakk> \<Longrightarrow> frechet \<subseteq> x"
+by (unfold superfrechet_def chain_def, blast)
+
+lemma (in UFT) Union_chain_superfrechet:
+  "\<lbrakk>c \<noteq> {}; c \<in> chain superfrechet\<rbrakk> \<Longrightarrow> \<Union>c \<in> superfrechet"
+proof (rule superfrechetI)
+  assume 1: "c \<in> chain superfrechet" and 2: "c \<noteq> {}"
+  thus "filter (\<Union>c)" by (rule Union_chain_filter)
+  from 2 obtain x where 3: "x \<in> c" by blast
+  from 1 3 have "frechet \<subseteq> x" by (rule lemma_mem_chain_frechet_subset)
+  also from 3 have "x \<subseteq> \<Union>c" by blast
+  finally show "frechet \<subseteq> \<Union>c" .
+qed
+
+subsubsection {* Existence of free ultrafilter *}
+
+lemma (in UFT) max_cofinite_filter_Ex:
+  "\<exists>U\<in>superfrechet. \<forall>G\<in>superfrechet. U \<subseteq> G \<longrightarrow> U = G"
+proof (rule Zorn_Lemma2 [rule_format])
+  fix c assume c: "c \<in> chain superfrechet"
+  show "\<exists>U\<in>superfrechet. \<forall>G\<in>c. G \<subseteq> U" (is "?U")
+  proof (cases)
+    assume "c = {}"
+    with frechet_in_superfrechet show "?U" by blast
+  next
+    assume A: "c \<noteq> {}"
+    from A c have "\<Union>c \<in> superfrechet"
+      by (rule Union_chain_superfrechet)
+    thus "?U" by blast
+  qed
+qed
+
+lemma (in UFT) mem_superfrechet_all_infinite:
+  "\<lbrakk>U \<in> superfrechet; A \<in> U\<rbrakk> \<Longrightarrow> infinite A"
+proof
+  assume U: "U \<in> superfrechet" and A: "A \<in> U" and fin: "finite A"
+  from U have fil: "filter U" and fre: "frechet \<subseteq> U"
+    by (simp_all add: superfrechet_def)
+  from fin have "- A \<in> frechet" by (simp add: frechet_def)
+  with fre have cA: "- A \<in> U" by (rule subsetD)
+  from fil A cA have "A \<inter> - A \<in> U" by (rule filter.Int)
+  with fil show "False" by (simp add: filter.empty)
+qed
+
+text {* There exists a free ultrafilter on any infinite set *}
+
+lemma (in UFT) freeultrafilter_ex:
+  "\<exists>U::'a set set. freeultrafilter U"
+proof -
+  from max_cofinite_filter_Ex obtain U
+    where U: "U \<in> superfrechet"
+      and max [rule_format]: "\<forall>G\<in>superfrechet. U \<subseteq> G \<longrightarrow> U = G" ..
+  from U have fil: "filter U" by (rule superfrechetD1)
+  from U have fre: "frechet \<subseteq> U" by (rule superfrechetD2)
+  have ultra: "ultrafilter_axioms U"
+  proof (rule filter.max_filter_ultrafilter [OF fil])
+    fix G assume G: "filter G" and UG: "U \<subseteq> G"
+    from fre UG have "frechet \<subseteq> G" by simp
+    with G have "G \<in> superfrechet" by (rule superfrechetI)
+    from this UG show "U = G" by (rule max)
+  qed
+  have free: "freeultrafilter_axioms U"
+  proof (rule freeultrafilter_axioms.intro)
+    fix A assume "A \<in> U"
+    with U show "infinite A" by (rule mem_superfrechet_all_infinite)
+  qed
+  show ?thesis
+  proof
+    from fil ultra free show "freeultrafilter U"
+      by (rule freeultrafilter.intro [OF ultrafilter.intro])
+      (* FIXME: unfold_locales should use chained facts *)
+  qed
+qed
+
+lemmas freeultrafilter_Ex = UFT.freeultrafilter_ex
+
+hide (open) const filter
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/HDeriv.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,472 @@
+(*  Title       : Deriv.thy
+    ID          : $Id$
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
+*)
+
+header{* Differentiation (Nonstandard) *}
+
+theory HDeriv
+imports Deriv HLim
+begin
+
+text{*Nonstandard Definitions*}
+
+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 @= 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
+  increment :: "[real=>real,real,hypreal] => hypreal" where
+  [code func del]: "increment f x h = (@inc. f NSdifferentiable x &
+           inc = ( *f* f)(hypreal_of_real x + h) - hypreal_of_real (f x))"
+
+
+subsection {* Derivatives *}
+
+lemma DERIV_NS_iff:
+      "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --NS> D)"
+by (simp add: deriv_def LIM_NSLIM_iff)
+
+lemma NS_DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) -- 0 --NS> 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 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
+
+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
+
+text {*First NSDERIV in terms of NSLIM*}
+
+text{*first equivalence *}
+lemma NSDERIV_NSLIM_iff:
+      "(NSDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --NS> 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{*second equivalence *}
+lemma NSDERIV_NSLIM_iff2:
+     "(NSDERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --NS> D)"
+by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff  diff_minus [symmetric]
+              LIM_NSLIM_iff [symmetric])
+
+(* while we're at it! *)
+
+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)
+
+(*FIXME DELETE*)
+lemma hypreal_not_eq_minus_iff:
+  "(x \<noteq> a) = (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
+
+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 simp add: diff_minus)
+apply (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: diff_minus)
+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 diff_minus)
+done
+
+text{*Differentiability implies continuity
+         nice and simple "algebraic" proof*}
+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_divide_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{*Differentiation rules for combinations of functions
+      follow from clear, straightforard, algebraic
+      manipulations*}
+text{*Constant function*}
+
+(* use simple constant nslimit theorem *)
+lemma NSDERIV_const [simp]: "(NSDERIV (%x. k) x :> 0)"
+by (simp add: NSDERIV_NSLIM_iff)
+
+text{*Sum of functions- proved easily*}
+
+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: diff_def add_ac)
+done
+
+text{*Product of functions - Proof is trivial but tedious
+  and long due to rearrangement of terms*}
+
+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 mult_ac)
+
+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
+
+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)
+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: left_distrib right_distrib 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{*Multiplying by a constant*}
+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
+
+text{*Negation of function*}
+lemma NSDERIV_minus: "NSDERIV f x :> D ==> NSDERIV (%x. -(f x)) x :> -D"
+proof (simp add: NSDERIV_NSLIM_iff)
+  assume "(\<lambda>h. (f (x + h) - f x) / h) -- 0 --NS> D"
+  hence deriv: "(\<lambda>h. - ((f(x+h) - f x) / h)) -- 0 --NS> - 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 diff_def)
+  with deriv
+  show "(\<lambda>h. (- f (x + h) + f x) / h) -- 0 --NS> - D" by simp
+qed
+
+text{*Subtraction*}
+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)
+
+lemma NSDERIV_diff:
+     "[| NSDERIV f x :> Da; NSDERIV g x :> Db |]
+      ==> NSDERIV (%x. f x - g x) x :> Da-Db"
+apply (simp add: diff_minus)
+apply (blast intro: NSDERIV_add_minus)
+done
+
+text{*  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.*}
+
+(* lemmas *)
+
+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
+
+(* can be proved differently using NSLIM_isCont_iff *)
+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
+
+(*---------------------------------------------------------------
+   from one version of differentiability
+
+                f(x) - f(a)
+              --------------- \<approx> Db
+                  x - a
+ ---------------------------------------------------------------*)
+
+lemma NSDERIVD1: "[| NSDERIV f (g x) :> Da;
+         ( *f* g) (star_of(x) + xa) \<noteq> star_of (g x);
+         ( *f* g) (star_of(x) + xa) \<approx> star_of (g x)
+      |] ==> (( *f* f) (( *f* g) (star_of(x) + xa))
+                   - 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 diff_minus [symmetric])
+
+(*--------------------------------------------------------------
+   from other version of differentiability
+
+                f(x + h) - f(x)
+               ----------------- \<approx> Db
+                       h
+ --------------------------------------------------------------*)
+
+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)
+
+lemma lemma_chain: "(z::'a::real_normed_field star) \<noteq> 0 ==> x*y = (x*inverse(z))*(z*y)"
+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)
+qed
+
+text{*This proof uses both definitions of differentiability.*}
+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{*Differentiation of natural number powers*}
+lemma NSDERIV_Id [simp]: "NSDERIV (%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)
+
+(*Can't get rid of x \<noteq> 0 because it isn't continuous at zero*)
+lemma NSDERIV_inverse:
+  fixes x :: "'a::{real_normed_field,recpower}"
+  shows "x \<noteq> 0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ Suc (Suc 0)))"
+apply (simp add: nsderiv_def)
+apply (rule ballI, simp, clarify)
+apply (frule (1) Infinitesimal_add_not_zero)
+apply (simp add: add_commute)
+(*apply (auto simp add: starfun_inverse_inverse realpow_two
+        simp del: minus_mult_left [symmetric] minus_mult_right [symmetric])*)
+apply (simp add: inverse_add nonzero_inverse_mult_distrib [symmetric] power_Suc
+              nonzero_inverse_minus_eq [symmetric] add_ac mult_ac diff_def
+            del: inverse_mult_distrib inverse_minus_eq
+                 minus_mult_left [symmetric] minus_mult_right [symmetric])
+apply (subst mult_commute, simp add: nonzero_mult_divide_cancel_right)
+apply (simp (no_asm_simp) add: mult_assoc [symmetric] left_distrib
+            del: minus_mult_left [symmetric] minus_mult_right [symmetric])
+apply (rule_tac y = "inverse (- (star_of x * star_of x))" in approx_trans)
+apply (rule inverse_add_Infinitesimal_approx2)
+apply (auto dest!: hypreal_of_real_HFinite_diff_Infinitesimal
+            simp add: inverse_minus_eq [symmetric] HFinite_minus_iff)
+apply (rule Infinitesimal_HFinite_mult, auto)
+done
+
+subsubsection {* Equivalence of NS and Standard definitions *}
+
+lemma divideR_eq_divide: "x /\<^sub>R y = x / y"
+by (simp add: real_scaleR_def divide_inverse mult_commute)
+
+text{*Now equivalence between NSDERIV and DERIV*}
+lemma NSDERIV_DERIV_iff: "(NSDERIV f x :> D) = (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{*Derivative of inverse*}
+
+lemma NSDERIV_inverse_fun:
+  fixes x :: "'a::{real_normed_field,recpower}"
+  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: realpow_Suc)
+
+text{*Derivative of quotient*}
+
+lemma NSDERIV_quotient:
+  fixes x :: "'a::{real_normed_field,recpower}"
+  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: realpow_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 hypreal_eq_minus_iff3: "(x = y + z) = (x + -z = (y::hypreal))"
+by auto
+
+lemma CARAT_DERIVD:
+  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_divide_cancel_right)
+  thus ?thesis using all
+    by (simp add: NSDERIV_iff2 starfun_if_eq cong: if_cong)
+qed
+
+subsubsection {* Differentiability predicate *}
+
+lemma NSdifferentiableD: "f NSdifferentiable x ==> \<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)
+
+
+subsection {*(NS) Increment*}
+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
+
+(* 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 hypreal_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: left_distrib)
+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)
+
+
+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: left_distrib [symmetric] mem_infmal_iff [symmetric])
+apply (erule Infinitesimal_subset_HFinite [THEN subsetD])
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/HLim.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,334 @@
+(*  Title       : HLim.thy
+    ID          : $Id$
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
+*)
+
+header{* Limits and Continuity (Nonstandard) *}
+
+theory HLim
+imports Star Lim
+begin
+
+text{*Nonstandard Definitions*}
+
+definition
+  NSLIM :: "['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool"
+            ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60) where
+  [code func del]: "f -- a --NS> L =
+    (\<forall>x. (x \<noteq> star_of a & x @= star_of a --> ( *f* f) x @= star_of L))"
+
+definition
+  isNSCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool" where
+    --{*NS definition dispenses with limit notions*}
+  [code func del]: "isNSCont f a = (\<forall>y. y @= star_of a -->
+         ( *f* f) y @= star_of (f a))"
+
+definition
+  isNSUCont :: "['a::real_normed_vector => 'b::real_normed_vector] => bool" where
+  [code func del]: "isNSUCont f = (\<forall>x y. x @= y --> ( *f* f) x @= ( *f* f) y)"
+
+
+subsection {* Limits of Functions *}
+
+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 -- a --NS> L"
+by (simp add: NSLIM_def)
+
+lemma NSLIM_D:
+  "\<lbrakk>f -- a --NS> 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{*Proving properties of limits using nonstandard definition.
+      The properties hold for standard limits as well!*}
+
+lemma NSLIM_mult:
+  fixes l m :: "'a::real_normed_algebra"
+  shows "[| f -- x --NS> l; g -- x --NS> m |]
+      ==> (%x. f(x) * g(x)) -- x --NS> (l * m)"
+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 NSLIM_scaleR:
+  "[| f -- x --NS> l; g -- x --NS> m |]
+      ==> (%x. f(x) *\<^sub>R g(x)) -- x --NS> (l *\<^sub>R m)"
+by (auto simp add: NSLIM_def intro!: approx_scaleR_HFinite)
+
+lemma NSLIM_add:
+     "[| f -- x --NS> l; g -- x --NS> m |]
+      ==> (%x. f(x) + g(x)) -- x --NS> (l + m)"
+by (auto simp add: NSLIM_def intro!: approx_add)
+
+lemma NSLIM_const [simp]: "(%x. k) -- x --NS> k"
+by (simp add: NSLIM_def)
+
+lemma NSLIM_minus: "f -- a --NS> L ==> (%x. -f(x)) -- a --NS> -L"
+by (simp add: NSLIM_def)
+
+lemma NSLIM_diff:
+  "\<lbrakk>f -- x --NS> l; g -- x --NS> m\<rbrakk> \<Longrightarrow> (\<lambda>x. f x - g x) -- x --NS> (l - m)"
+by (simp only: diff_def NSLIM_add NSLIM_minus)
+
+lemma NSLIM_add_minus: "[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f(x) + -g(x)) -- x --NS> (l + -m)"
+by (simp only: NSLIM_add NSLIM_minus)
+
+lemma NSLIM_inverse:
+  fixes L :: "'a::real_normed_div_algebra"
+  shows "[| f -- a --NS> L;  L \<noteq> 0 |]
+      ==> (%x. inverse(f(x))) -- a --NS> (inverse L)"
+apply (simp add: NSLIM_def, clarify)
+apply (drule spec)
+apply (auto simp add: star_of_approx_inverse)
+done
+
+lemma NSLIM_zero:
+  assumes f: "f -- a --NS> l" shows "(%x. f(x) - l) -- a --NS> 0"
+proof -
+  have "(\<lambda>x. f x - l) -- a --NS> l - l"
+    by (rule NSLIM_diff [OF f NSLIM_const])
+  thus ?thesis by simp
+qed
+
+lemma NSLIM_zero_cancel: "(%x. f(x) - l) -- x --NS> 0 ==> f -- x --NS> l"
+apply (drule_tac g = "%x. l" and m = l in NSLIM_add)
+apply (auto simp add: diff_minus add_assoc)
+done
+
+lemma NSLIM_const_not_eq:
+  fixes a :: "'a::real_normed_algebra_1"
+  shows "k \<noteq> L \<Longrightarrow> \<not> (\<lambda>x. k) -- a --NS> 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_not_zero:
+  fixes a :: "'a::real_normed_algebra_1"
+  shows "k \<noteq> 0 \<Longrightarrow> \<not> (\<lambda>x. k) -- a --NS> 0"
+by (rule NSLIM_const_not_eq)
+
+lemma NSLIM_const_eq:
+  fixes a :: "'a::real_normed_algebra_1"
+  shows "(\<lambda>x. k) -- a --NS> L \<Longrightarrow> k = L"
+apply (rule ccontr)
+apply (blast dest: NSLIM_const_not_eq)
+done
+
+lemma NSLIM_unique:
+  fixes a :: "'a::real_normed_algebra_1"
+  shows "\<lbrakk>f -- a --NS> L; f -- a --NS> M\<rbrakk> \<Longrightarrow> L = M"
+apply (drule (1) NSLIM_diff)
+apply (auto dest!: NSLIM_const_eq)
+done
+
+lemma NSLIM_mult_zero:
+  fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra"
+  shows "[| f -- x --NS> 0; g -- x --NS> 0 |] ==> (%x. f(x)*g(x)) -- x --NS> 0"
+by (drule NSLIM_mult, auto)
+
+lemma NSLIM_self: "(%x. x) -- a --NS> a"
+by (simp add: NSLIM_def)
+
+subsubsection {* Equivalence of @{term LIM} and @{term NSLIM} *}
+
+lemma LIM_NSLIM:
+  assumes f: "f -- a --> L" shows "f -- a --NS> 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"
+      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"
+      by transfer
+    from approx have "x - star_of a \<in> Infinitesimal"
+      by (unfold approx_def)
+    hence "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"
+    by (unfold approx_def)
+qed
+
+lemma NSLIM_LIM:
+  assumes f: "f -- a --NS> L" shows "f -- a --> 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"
+  proof (rule exI, safe)
+    show "0 < epsilon" by (rule hypreal_epsilon_gt_zero)
+  next
+    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"
+      by (rule hnorm_less_Infinitesimal)
+    hence "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"
+      by (unfold approx_def)
+    thus "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"
+    by transfer
+qed
+
+theorem LIM_NSLIM_iff: "(f -- x --> L) = (f -- x --NS> L)"
+by (blast intro: LIM_NSLIM NSLIM_LIM)
+
+
+subsection {* Continuity *}
+
+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 isNSCont_NSLIM: "isNSCont f a ==> f -- a --NS> (f a) "
+by (simp add: isNSCont_def NSLIM_def)
+
+lemma NSLIM_isNSCont: "f -- a --NS> (f a) ==> isNSCont f a"
+apply (simp add: isNSCont_def NSLIM_def, auto)
+apply (case_tac "y = star_of a", auto)
+done
+
+text{*NS continuity can be defined using NS Limit in
+    similar fashion to standard def of continuity*}
+lemma isNSCont_NSLIM_iff: "(isNSCont f a) = (f -- a --NS> (f a))"
+by (blast intro: isNSCont_NSLIM NSLIM_isNSCont)
+
+text{*Hence, NS continuity can be given
+  in terms of standard limit*}
+lemma isNSCont_LIM_iff: "(isNSCont f a) = (f -- a --> (f a))"
+by (simp add: LIM_NSLIM_iff isNSCont_NSLIM_iff)
+
+text{*Moreover, it's trivial now that NS continuity
+  is equivalent to standard continuity*}
+lemma isNSCont_isCont_iff: "(isNSCont f a) = (isCont f a)"
+apply (simp add: isCont_def)
+apply (rule isNSCont_LIM_iff)
+done
+
+text{*Standard continuity ==> NS continuity*}
+lemma isCont_isNSCont: "isCont f a ==> isNSCont f a"
+by (erule isNSCont_isCont_iff [THEN iffD2])
+
+text{*NS continuity ==> Standard continuity*}
+lemma isNSCont_isCont: "isNSCont f a ==> isCont f a"
+by (erule isNSCont_isCont_iff [THEN iffD1])
+
+text{*Alternative definition of continuity*}
+
+(* Prove equivalence between NS limits - *)
+(* seems easier than using standard def  *)
+lemma NSLIM_h_iff: "(f -- a --NS> L) = ((%h. f(a + h)) -- 0 --NS> 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 diff_def [symmetric])
+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 approx_refl star_n_zero_num)
+done
+
+lemma NSLIM_isCont_iff: "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)"
+by (rule NSLIM_h_iff)
+
+lemma isNSCont_minus: "isNSCont f a ==> isNSCont (%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_const [simp]: "isNSCont (%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
+
+
+subsection {* Uniform Continuity *}
+
+lemma isNSUContD: "[| isNSUCont f; x \<approx> y|] ==> ( *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)
+  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"
+      by (auto simp add: isUCont_def)
+    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"
+      by transfer
+    from approx have "x - y \<in> Infinitesimal"
+      by (unfold approx_def)
+    hence "hnorm (x - y) < star_of s"
+      using s by (rule InfinitesimalD2)
+    thus "hnorm (starfun f x - starfun f y) < star_of r"
+      by (rule less_r')
+  qed
+  thus "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, 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)
+  next
+    fix x y :: "'a star"
+    assume "hnorm (x - y) < epsilon"
+    with Infinitesimal_epsilon
+    have "x - y \<in> Infinitesimal"
+      by (rule hnorm_less_Infinitesimal)
+    hence "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"
+      by (unfold approx_def)
+    thus "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"
+    by transfer
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/HLog.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,156 @@
+(*  Title       : HLog.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 2000,2001 University of Edinburgh
+*)
+
+header{*Logarithms: Non-Standard Version*}
+
+theory HLog
+imports Log HTranscendental
+begin
+
+
+(* should be in NSA.ML *)
+lemma epsilon_ge_zero [simp]: "0 \<le> epsilon"
+by (simp add: epsilon_def star_n_zero_num star_n_le)
+
+lemma hpfinite_witness: "epsilon : {x. 0 \<le> x & x : HFinite}"
+by auto
+
+
+definition
+  powhr  :: "[hypreal,hypreal] => hypreal"     (infixr "powhr" 80) where
+  [transfer_unfold, code func del]: "x powhr a = starfun2 (op powr) x a"
+  
+definition
+  hlog :: "[hypreal,hypreal] => hypreal" where
+  [transfer_unfold, code func del]: "hlog a x = starfun2 log a x"
+
+lemma powhr: "(star_n X) powhr (star_n Y) = star_n (%n. (X n) powr (Y n))"
+by (simp add: powhr_def starfun2_star_n)
+
+lemma powhr_one_eq_one [simp]: "!!a. 1 powhr a = 1"
+by (transfer, simp)
+
+lemma powhr_mult:
+  "!!a x y. [| 0 < x; 0 < y |] ==> (x * y) powhr a = (x powhr a) * (y powhr a)"
+by (transfer, rule powr_mult)
+
+lemma powhr_gt_zero [simp]: "!!a x. 0 < x powhr a"
+by (transfer, simp)
+
+lemma powhr_not_zero [simp]: "x powhr a \<noteq> 0"
+by (rule powhr_gt_zero [THEN hypreal_not_refl2, THEN not_sym])
+
+lemma powhr_divide:
+  "!!a x y. [| 0 < x; 0 < y |] ==> (x / y) powhr a = (x powhr a)/(y powhr a)"
+by (transfer, rule powr_divide)
+
+lemma powhr_add: "!!a b x. x powhr (a + b) = (x powhr a) * (x powhr b)"
+by (transfer, rule powr_add)
+
+lemma powhr_powhr: "!!a b x. (x powhr a) powhr b = x powhr (a * b)"
+by (transfer, rule powr_powr)
+
+lemma powhr_powhr_swap: "!!a b x. (x powhr a) powhr b = (x powhr b) powhr a"
+by (transfer, rule powr_powr_swap)
+
+lemma powhr_minus: "!!a x. x powhr (-a) = inverse (x powhr a)"
+by (transfer, rule powr_minus)
+
+lemma powhr_minus_divide: "x powhr (-a) = 1/(x powhr a)"
+by (simp add: divide_inverse powhr_minus)
+
+lemma powhr_less_mono: "!!a b x. [| a < b; 1 < x |] ==> x powhr a < x powhr b"
+by (transfer, simp)
+
+lemma powhr_less_cancel: "!!a b x. [| x powhr a < x powhr b; 1 < x |] ==> a < b"
+by (transfer, simp)
+
+lemma powhr_less_cancel_iff [simp]:
+     "1 < x ==> (x powhr a < x powhr b) = (a < b)"
+by (blast intro: powhr_less_cancel powhr_less_mono)
+
+lemma powhr_le_cancel_iff [simp]:
+     "1 < x ==> (x powhr a \<le> x powhr b) = (a \<le> b)"
+by (simp add: linorder_not_less [symmetric])
+
+lemma hlog:
+     "hlog (star_n X) (star_n Y) =  
+      star_n (%n. log (X n) (Y n))"
+by (simp add: hlog_def starfun2_star_n)
+
+lemma hlog_starfun_ln: "!!x. ( *f* ln) x = hlog (( *f* exp) 1) x"
+by (transfer, rule log_ln)
+
+lemma powhr_hlog_cancel [simp]:
+     "!!a x. [| 0 < a; a \<noteq> 1; 0 < x |] ==> a powhr (hlog a x) = x"
+by (transfer, simp)
+
+lemma hlog_powhr_cancel [simp]:
+     "!!a y. [| 0 < a; a \<noteq> 1 |] ==> hlog a (a powhr y) = y"
+by (transfer, simp)
+
+lemma hlog_mult:
+     "!!a x y. [| 0 < a; a \<noteq> 1; 0 < x; 0 < y  |]  
+      ==> hlog a (x * y) = hlog a x + hlog a y"
+by (transfer, rule log_mult)
+
+lemma hlog_as_starfun:
+     "!!a x. [| 0 < a; a \<noteq> 1 |] ==> hlog a x = ( *f* ln) x / ( *f* ln) a"
+by (transfer, simp add: log_def)
+
+lemma hlog_eq_div_starfun_ln_mult_hlog:
+     "!!a b x. [| 0 < a; a \<noteq> 1; 0 < b; b \<noteq> 1; 0 < x |]  
+      ==> hlog a x = (( *f* ln) b/( *f*ln) a) * hlog b x"
+by (transfer, rule log_eq_div_ln_mult_log)
+
+lemma powhr_as_starfun: "!!a x. x powhr a = ( *f* exp) (a * ( *f* ln) x)"
+by (transfer, simp add: powr_def)
+
+lemma HInfinite_powhr:
+     "[| x : HInfinite; 0 < x; a : HFinite - Infinitesimal;  
+         0 < a |] ==> x powhr a : HInfinite"
+apply (auto intro!: starfun_ln_ge_zero starfun_ln_HInfinite HInfinite_HFinite_not_Infinitesimal_mult2 starfun_exp_HInfinite 
+       simp add: order_less_imp_le HInfinite_gt_zero_gt_one powhr_as_starfun zero_le_mult_iff)
+done
+
+lemma hlog_hrabs_HInfinite_Infinitesimal:
+     "[| x : HFinite - Infinitesimal; a : HInfinite; 0 < a |]  
+      ==> hlog a (abs x) : Infinitesimal"
+apply (frule HInfinite_gt_zero_gt_one)
+apply (auto intro!: starfun_ln_HFinite_not_Infinitesimal
+            HInfinite_inverse_Infinitesimal Infinitesimal_HFinite_mult2 
+        simp add: starfun_ln_HInfinite not_Infinitesimal_not_zero 
+          hlog_as_starfun hypreal_not_refl2 [THEN not_sym] divide_inverse)
+done
+
+lemma hlog_HInfinite_as_starfun:
+     "[| a : HInfinite; 0 < a |] ==> hlog a x = ( *f* ln) x / ( *f* ln) a"
+by (rule hlog_as_starfun, auto)
+
+lemma hlog_one [simp]: "!!a. hlog a 1 = 0"
+by (transfer, simp)
+
+lemma hlog_eq_one [simp]: "!!a. [| 0 < a; a \<noteq> 1 |] ==> hlog a a = 1"
+by (transfer, rule log_eq_one)
+
+lemma hlog_inverse:
+     "[| 0 < a; a \<noteq> 1; 0 < x |] ==> hlog a (inverse x) = - hlog a x"
+apply (rule add_left_cancel [of "hlog a x", THEN iffD1])
+apply (simp add: hlog_mult [symmetric])
+done
+
+lemma hlog_divide:
+     "[| 0 < a; a \<noteq> 1; 0 < x; 0 < y|] ==> hlog a (x/y) = hlog a x - hlog a y"
+by (simp add: hlog_mult hlog_inverse divide_inverse)
+
+lemma hlog_less_cancel_iff [simp]:
+     "!!a x y. [| 1 < a; 0 < x; 0 < y |] ==> (hlog a x < hlog a y) = (x < y)"
+by (transfer, simp)
+
+lemma hlog_le_cancel_iff [simp]:
+     "[| 1 < a; 0 < x; 0 < y |] ==> (hlog a x \<le> hlog a y) = (x \<le> y)"
+by (simp add: linorder_not_less [symmetric])
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/HSEQ.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,531 @@
+(*  Title       : HSEQ.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Convergence of sequences and series
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
+    Additional contributions by Jeremy Avigad and Brian Huffman
+*)
+
+header {* Sequences and Convergence (Nonstandard) *}
+
+theory HSEQ
+imports SEQ NatStar
+begin
+
+definition
+  NSLIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool"
+    ("((_)/ ----NS> (_))" [60, 60] 60) where
+    --{*Nonstandard definition of convergence of sequence*}
+  [code func del]: "X ----NS> L = (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> star_of L)"
+
+definition
+  nslim :: "(nat => 'a::real_normed_vector) => 'a" where
+    --{*Nonstandard definition of limit using choice operator*}
+  "nslim X = (THE L. X ----NS> L)"
+
+definition
+  NSconvergent :: "(nat => 'a::real_normed_vector) => bool" where
+    --{*Nonstandard definition of convergence*}
+  "NSconvergent X = (\<exists>L. X ----NS> L)"
+
+definition
+  NSBseq :: "(nat => 'a::real_normed_vector) => bool" where
+    --{*Nonstandard definition for bounded sequence*}
+  [code func del]: "NSBseq X = (\<forall>N \<in> HNatInfinite. ( *f* X) N : HFinite)"
+
+definition
+  NSCauchy :: "(nat => 'a::real_normed_vector) => bool" where
+    --{*Nonstandard definition*}
+  [code func del]: "NSCauchy X = (\<forall>M \<in> HNatInfinite. \<forall>N \<in> HNatInfinite. ( *f* X) M \<approx> ( *f* X) N)"
+
+subsection {* Limits of Sequences *}
+
+lemma NSLIMSEQ_iff:
+    "(X ----NS> L) = (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> star_of L)"
+by (simp add: NSLIMSEQ_def)
+
+lemma NSLIMSEQ_I:
+  "(\<And>N. N \<in> HNatInfinite \<Longrightarrow> starfun X N \<approx> star_of L) \<Longrightarrow> X ----NS> L"
+by (simp add: NSLIMSEQ_def)
+
+lemma NSLIMSEQ_D:
+  "\<lbrakk>X ----NS> L; N \<in> HNatInfinite\<rbrakk> \<Longrightarrow> starfun X N \<approx> star_of L"
+by (simp add: NSLIMSEQ_def)
+
+lemma NSLIMSEQ_const: "(%n. k) ----NS> k"
+by (simp add: NSLIMSEQ_def)
+
+lemma NSLIMSEQ_add:
+      "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n + Y n) ----NS> a + b"
+by (auto intro: approx_add simp add: NSLIMSEQ_def starfun_add [symmetric])
+
+lemma NSLIMSEQ_add_const: "f ----NS> a ==> (%n.(f n + b)) ----NS> a + b"
+by (simp only: NSLIMSEQ_add NSLIMSEQ_const)
+
+lemma NSLIMSEQ_mult:
+  fixes a b :: "'a::real_normed_algebra"
+  shows "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n * Y n) ----NS> a * b"
+by (auto intro!: approx_mult_HFinite simp add: NSLIMSEQ_def)
+
+lemma NSLIMSEQ_minus: "X ----NS> a ==> (%n. -(X n)) ----NS> -a"
+by (auto simp add: NSLIMSEQ_def)
+
+lemma NSLIMSEQ_minus_cancel: "(%n. -(X n)) ----NS> -a ==> X ----NS> a"
+by (drule NSLIMSEQ_minus, simp)
+
+(* FIXME: delete *)
+lemma NSLIMSEQ_add_minus:
+     "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n + -Y n) ----NS> a + -b"
+by (simp add: NSLIMSEQ_add NSLIMSEQ_minus)
+
+lemma NSLIMSEQ_diff:
+     "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n - Y n) ----NS> a - b"
+by (simp add: diff_minus NSLIMSEQ_add NSLIMSEQ_minus)
+
+lemma NSLIMSEQ_diff_const: "f ----NS> a ==> (%n.(f n - b)) ----NS> a - b"
+by (simp add: NSLIMSEQ_diff NSLIMSEQ_const)
+
+lemma NSLIMSEQ_inverse:
+  fixes a :: "'a::real_normed_div_algebra"
+  shows "[| X ----NS> a;  a ~= 0 |] ==> (%n. inverse(X n)) ----NS> inverse(a)"
+by (simp add: NSLIMSEQ_def star_of_approx_inverse)
+
+lemma NSLIMSEQ_mult_inverse:
+  fixes a b :: "'a::real_normed_field"
+  shows
+     "[| X ----NS> a;  Y ----NS> b;  b ~= 0 |] ==> (%n. X n / Y n) ----NS> a/b"
+by (simp add: NSLIMSEQ_mult NSLIMSEQ_inverse divide_inverse)
+
+lemma starfun_hnorm: "\<And>x. hnorm (( *f* f) x) = ( *f* (\<lambda>x. norm (f x))) x"
+by transfer simp
+
+lemma NSLIMSEQ_norm: "X ----NS> a \<Longrightarrow> (\<lambda>n. norm (X n)) ----NS> norm a"
+by (simp add: NSLIMSEQ_def starfun_hnorm [symmetric] approx_hnorm)
+
+text{*Uniqueness of limit*}
+lemma NSLIMSEQ_unique: "[| X ----NS> a; X ----NS> b |] ==> a = b"
+apply (simp add: NSLIMSEQ_def)
+apply (drule HNatInfinite_whn [THEN [2] bspec])+
+apply (auto dest: approx_trans3)
+done
+
+lemma NSLIMSEQ_pow [rule_format]:
+  fixes a :: "'a::{real_normed_algebra,recpower}"
+  shows "(X ----NS> a) --> ((%n. (X n) ^ m) ----NS> a ^ m)"
+apply (induct "m")
+apply (auto simp add: power_Suc intro: NSLIMSEQ_mult NSLIMSEQ_const)
+done
+
+text{*We can now try and derive a few properties of sequences,
+     starting with the limit comparison property for sequences.*}
+
+lemma NSLIMSEQ_le:
+       "[| f ----NS> l; g ----NS> m;
+           \<exists>N. \<forall>n \<ge> N. f(n) \<le> g(n)
+        |] ==> l \<le> (m::real)"
+apply (simp add: NSLIMSEQ_def, safe)
+apply (drule starfun_le_mono)
+apply (drule HNatInfinite_whn [THEN [2] bspec])+
+apply (drule_tac x = whn in spec)
+apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+
+apply clarify
+apply (auto intro: hypreal_of_real_le_add_Infininitesimal_cancel2)
+done
+
+lemma NSLIMSEQ_le_const: "[| X ----NS> (r::real); \<forall>n. a \<le> X n |] ==> a \<le> r"
+by (erule NSLIMSEQ_le [OF NSLIMSEQ_const], auto)
+
+lemma NSLIMSEQ_le_const2: "[| X ----NS> (r::real); \<forall>n. X n \<le> a |] ==> r \<le> a"
+by (erule NSLIMSEQ_le [OF _ NSLIMSEQ_const], auto)
+
+text{*Shift a convergent series by 1:
+  By the equivalence between Cauchiness and convergence and because
+  the successor of an infinite hypernatural is also infinite.*}
+
+lemma NSLIMSEQ_Suc: "f ----NS> l ==> (%n. f(Suc n)) ----NS> l"
+apply (unfold NSLIMSEQ_def, safe)
+apply (drule_tac x="N + 1" in bspec)
+apply (erule HNatInfinite_add)
+apply (simp add: starfun_shift_one)
+done
+
+lemma NSLIMSEQ_imp_Suc: "(%n. f(Suc n)) ----NS> l ==> f ----NS> l"
+apply (unfold NSLIMSEQ_def, safe)
+apply (drule_tac x="N - 1" in bspec) 
+apply (erule Nats_1 [THEN [2] HNatInfinite_diff])
+apply (simp add: starfun_shift_one one_le_HNatInfinite)
+done
+
+lemma NSLIMSEQ_Suc_iff: "((%n. f(Suc n)) ----NS> l) = (f ----NS> l)"
+by (blast intro: NSLIMSEQ_imp_Suc NSLIMSEQ_Suc)
+
+subsubsection {* Equivalence of @{term LIMSEQ} and @{term NSLIMSEQ} *}
+
+lemma LIMSEQ_NSLIMSEQ:
+  assumes X: "X ----> L" shows "X ----NS> L"
+proof (rule NSLIMSEQ_I)
+  fix N assume N: "N \<in> HNatInfinite"
+  have "starfun X N - star_of L \<in> Infinitesimal"
+  proof (rule InfinitesimalI2)
+    fix r::real assume r: "0 < r"
+    from LIMSEQ_D [OF X r]
+    obtain no where "\<forall>n\<ge>no. norm (X n - L) < r" ..
+    hence "\<forall>n\<ge>star_of no. hnorm (starfun X n - star_of L) < star_of r"
+      by transfer
+    thus "hnorm (starfun X N - star_of L) < star_of r"
+      using N by (simp add: star_of_le_HNatInfinite)
+  qed
+  thus "starfun X N \<approx> star_of L"
+    by (unfold approx_def)
+qed
+
+lemma NSLIMSEQ_LIMSEQ:
+  assumes X: "X ----NS> L" shows "X ----> L"
+proof (rule LIMSEQ_I)
+  fix r::real assume r: "0 < r"
+  have "\<exists>no. \<forall>n\<ge>no. hnorm (starfun X n - star_of L) < star_of r"
+  proof (intro exI allI impI)
+    fix n assume "whn \<le> n"
+    with HNatInfinite_whn have "n \<in> HNatInfinite"
+      by (rule HNatInfinite_upward_closed)
+    with X have "starfun X n \<approx> star_of L"
+      by (rule NSLIMSEQ_D)
+    hence "starfun X n - star_of L \<in> Infinitesimal"
+      by (unfold approx_def)
+    thus "hnorm (starfun X n - star_of L) < star_of r"
+      using r by (rule InfinitesimalD2)
+  qed
+  thus "\<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
+    by transfer
+qed
+
+theorem LIMSEQ_NSLIMSEQ_iff: "(f ----> L) = (f ----NS> L)"
+by (blast intro: LIMSEQ_NSLIMSEQ NSLIMSEQ_LIMSEQ)
+
+(* Used once by Integration/Rats.thy in AFP *)
+lemma NSLIMSEQ_finite_set:
+     "!!(f::nat=>nat). \<forall>n. n \<le> f n ==> finite {n. f n \<le> u}"
+by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans)
+
+subsubsection {* Derived theorems about @{term NSLIMSEQ} *}
+
+text{*We prove the NS version from the standard one, since the NS proof
+   seems more complicated than the standard one above!*}
+lemma NSLIMSEQ_norm_zero: "((\<lambda>n. norm (X n)) ----NS> 0) = (X ----NS> 0)"
+by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_norm_zero)
+
+lemma NSLIMSEQ_rabs_zero: "((%n. \<bar>f n\<bar>) ----NS> 0) = (f ----NS> (0::real))"
+by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_rabs_zero)
+
+text{*Generalization to other limits*}
+lemma NSLIMSEQ_imp_rabs: "f ----NS> (l::real) ==> (%n. \<bar>f n\<bar>) ----NS> \<bar>l\<bar>"
+apply (simp add: NSLIMSEQ_def)
+apply (auto intro: approx_hrabs 
+            simp add: starfun_abs)
+done
+
+lemma NSLIMSEQ_inverse_zero:
+     "\<forall>y::real. \<exists>N. \<forall>n \<ge> N. y < f(n)
+      ==> (%n. inverse(f n)) ----NS> 0"
+by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_zero)
+
+lemma NSLIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----NS> 0"
+by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat)
+
+lemma NSLIMSEQ_inverse_real_of_nat_add:
+     "(%n. r + inverse(real(Suc n))) ----NS> r"
+by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add)
+
+lemma NSLIMSEQ_inverse_real_of_nat_add_minus:
+     "(%n. r + -inverse(real(Suc n))) ----NS> r"
+by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add_minus)
+
+lemma NSLIMSEQ_inverse_real_of_nat_add_minus_mult:
+     "(%n. r*( 1 + -inverse(real(Suc n)))) ----NS> r"
+by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add_minus_mult)
+
+
+subsection {* Convergence *}
+
+lemma nslimI: "X ----NS> L ==> nslim X = L"
+apply (simp add: nslim_def)
+apply (blast intro: NSLIMSEQ_unique)
+done
+
+lemma lim_nslim_iff: "lim X = nslim X"
+by (simp add: lim_def nslim_def LIMSEQ_NSLIMSEQ_iff)
+
+lemma NSconvergentD: "NSconvergent X ==> \<exists>L. (X ----NS> L)"
+by (simp add: NSconvergent_def)
+
+lemma NSconvergentI: "(X ----NS> L) ==> NSconvergent X"
+by (auto simp add: NSconvergent_def)
+
+lemma convergent_NSconvergent_iff: "convergent X = NSconvergent X"
+by (simp add: convergent_def NSconvergent_def LIMSEQ_NSLIMSEQ_iff)
+
+lemma NSconvergent_NSLIMSEQ_iff: "NSconvergent X = (X ----NS> nslim X)"
+by (auto intro: theI NSLIMSEQ_unique simp add: NSconvergent_def nslim_def)
+
+
+subsection {* Bounded Monotonic Sequences *}
+
+lemma NSBseqD: "[| NSBseq X;  N: HNatInfinite |] ==> ( *f* X) N : HFinite"
+by (simp add: NSBseq_def)
+
+lemma Standard_subset_HFinite: "Standard \<subseteq> HFinite"
+unfolding Standard_def by auto
+
+lemma NSBseqD2: "NSBseq X \<Longrightarrow> ( *f* X) N \<in> HFinite"
+apply (cases "N \<in> HNatInfinite")
+apply (erule (1) NSBseqD)
+apply (rule subsetD [OF Standard_subset_HFinite])
+apply (simp add: HNatInfinite_def Nats_eq_Standard)
+done
+
+lemma NSBseqI: "\<forall>N \<in> HNatInfinite. ( *f* X) N : HFinite ==> NSBseq X"
+by (simp add: NSBseq_def)
+
+text{*The standard definition implies the nonstandard definition*}
+
+lemma Bseq_NSBseq: "Bseq X ==> NSBseq X"
+proof (unfold NSBseq_def, safe)
+  assume X: "Bseq X"
+  fix N assume N: "N \<in> HNatInfinite"
+  from BseqD [OF X] obtain K where "\<forall>n. norm (X n) \<le> K" by fast
+  hence "\<forall>N. hnorm (starfun X N) \<le> star_of K" by transfer
+  hence "hnorm (starfun X N) \<le> star_of K" by simp
+  also have "star_of K < star_of (K + 1)" by simp
+  finally have "\<exists>x\<in>Reals. hnorm (starfun X N) < x" by (rule bexI, simp)
+  thus "starfun X N \<in> HFinite" by (simp add: HFinite_def)
+qed
+
+text{*The nonstandard definition implies the standard definition*}
+
+lemma SReal_less_omega: "r \<in> \<real> \<Longrightarrow> r < \<omega>"
+apply (insert HInfinite_omega)
+apply (simp add: HInfinite_def)
+apply (simp add: order_less_imp_le)
+done
+
+lemma NSBseq_Bseq: "NSBseq X \<Longrightarrow> Bseq X"
+proof (rule ccontr)
+  let ?n = "\<lambda>K. LEAST n. K < norm (X n)"
+  assume "NSBseq X"
+  hence finite: "( *f* X) (( *f* ?n) \<omega>) \<in> HFinite"
+    by (rule NSBseqD2)
+  assume "\<not> Bseq X"
+  hence "\<forall>K>0. \<exists>n. K < norm (X n)"
+    by (simp add: Bseq_def linorder_not_le)
+  hence "\<forall>K>0. K < norm (X (?n K))"
+    by (auto intro: LeastI_ex)
+  hence "\<forall>K>0. K < hnorm (( *f* X) (( *f* ?n) K))"
+    by transfer
+  hence "\<omega> < hnorm (( *f* X) (( *f* ?n) \<omega>))"
+    by simp
+  hence "\<forall>r\<in>\<real>. r < hnorm (( *f* X) (( *f* ?n) \<omega>))"
+    by (simp add: order_less_trans [OF SReal_less_omega])
+  hence "( *f* X) (( *f* ?n) \<omega>) \<in> HInfinite"
+    by (simp add: HInfinite_def)
+  with finite show "False"
+    by (simp add: HFinite_HInfinite_iff)
+qed
+
+text{* Equivalence of nonstandard and standard definitions
+  for a bounded sequence*}
+lemma Bseq_NSBseq_iff: "(Bseq X) = (NSBseq X)"
+by (blast intro!: NSBseq_Bseq Bseq_NSBseq)
+
+text{*A convergent sequence is bounded: 
+ Boundedness as a necessary condition for convergence. 
+ The nonstandard version has no existential, as usual *}
+
+lemma NSconvergent_NSBseq: "NSconvergent X ==> NSBseq X"
+apply (simp add: NSconvergent_def NSBseq_def NSLIMSEQ_def)
+apply (blast intro: HFinite_star_of approx_sym approx_HFinite)
+done
+
+text{*Standard Version: easily now proved using equivalence of NS and
+ standard definitions *}
+
+lemma convergent_Bseq: "convergent X ==> Bseq X"
+by (simp add: NSconvergent_NSBseq convergent_NSconvergent_iff Bseq_NSBseq_iff)
+
+subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
+
+lemma NSBseq_isUb: "NSBseq X ==> \<exists>U::real. isUb UNIV {x. \<exists>n. X n = x} U"
+by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isUb)
+
+lemma NSBseq_isLub: "NSBseq X ==> \<exists>U::real. isLub UNIV {x. \<exists>n. X n = x} U"
+by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isLub)
+
+subsubsection{*A Bounded and Monotonic Sequence Converges*}
+
+text{* The best of both worlds: Easier to prove this result as a standard
+   theorem and then use equivalence to "transfer" it into the
+   equivalent nonstandard form if needed!*}
+
+lemma Bmonoseq_NSLIMSEQ: "\<forall>n \<ge> m. X n = X m ==> \<exists>L. (X ----NS> L)"
+by (auto dest!: Bmonoseq_LIMSEQ simp add: LIMSEQ_NSLIMSEQ_iff)
+
+lemma NSBseq_mono_NSconvergent:
+     "[| NSBseq X; \<forall>m. \<forall>n \<ge> m. X m \<le> X n |] ==> NSconvergent (X::nat=>real)"
+by (auto intro: Bseq_mono_convergent 
+         simp add: convergent_NSconvergent_iff [symmetric] 
+                   Bseq_NSBseq_iff [symmetric])
+
+
+subsection {* Cauchy Sequences *}
+
+lemma NSCauchyI:
+  "(\<And>M N. \<lbrakk>M \<in> HNatInfinite; N \<in> HNatInfinite\<rbrakk> \<Longrightarrow> starfun X M \<approx> starfun X N)
+   \<Longrightarrow> NSCauchy X"
+by (simp add: NSCauchy_def)
+
+lemma NSCauchyD:
+  "\<lbrakk>NSCauchy X; M \<in> HNatInfinite; N \<in> HNatInfinite\<rbrakk>
+   \<Longrightarrow> starfun X M \<approx> starfun X N"
+by (simp add: NSCauchy_def)
+
+subsubsection{*Equivalence Between NS and Standard*}
+
+lemma Cauchy_NSCauchy:
+  assumes X: "Cauchy X" shows "NSCauchy X"
+proof (rule NSCauchyI)
+  fix M assume M: "M \<in> HNatInfinite"
+  fix N assume N: "N \<in> HNatInfinite"
+  have "starfun X M - starfun X N \<in> Infinitesimal"
+  proof (rule InfinitesimalI2)
+    fix r :: real assume r: "0 < r"
+    from CauchyD [OF X r]
+    obtain k where "\<forall>m\<ge>k. \<forall>n\<ge>k. norm (X m - X n) < r" ..
+    hence "\<forall>m\<ge>star_of k. \<forall>n\<ge>star_of k.
+           hnorm (starfun X m - starfun X n) < star_of r"
+      by transfer
+    thus "hnorm (starfun X M - starfun X N) < star_of r"
+      using M N by (simp add: star_of_le_HNatInfinite)
+  qed
+  thus "starfun X M \<approx> starfun X N"
+    by (unfold approx_def)
+qed
+
+lemma NSCauchy_Cauchy:
+  assumes X: "NSCauchy X" shows "Cauchy X"
+proof (rule CauchyI)
+  fix r::real assume r: "0 < r"
+  have "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. hnorm (starfun X m - starfun X n) < star_of r"
+  proof (intro exI allI impI)
+    fix M assume "whn \<le> M"
+    with HNatInfinite_whn have M: "M \<in> HNatInfinite"
+      by (rule HNatInfinite_upward_closed)
+    fix N assume "whn \<le> N"
+    with HNatInfinite_whn have N: "N \<in> HNatInfinite"
+      by (rule HNatInfinite_upward_closed)
+    from X M N have "starfun X M \<approx> starfun X N"
+      by (rule NSCauchyD)
+    hence "starfun X M - starfun X N \<in> Infinitesimal"
+      by (unfold approx_def)
+    thus "hnorm (starfun X M - starfun X N) < star_of r"
+      using r by (rule InfinitesimalD2)
+  qed
+  thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. norm (X m - X n) < r"
+    by transfer
+qed
+
+theorem NSCauchy_Cauchy_iff: "NSCauchy X = Cauchy X"
+by (blast intro!: NSCauchy_Cauchy Cauchy_NSCauchy)
+
+subsubsection {* Cauchy Sequences are Bounded *}
+
+text{*A Cauchy sequence is bounded -- nonstandard version*}
+
+lemma NSCauchy_NSBseq: "NSCauchy X ==> NSBseq X"
+by (simp add: Cauchy_Bseq Bseq_NSBseq_iff [symmetric] NSCauchy_Cauchy_iff)
+
+subsubsection {* Cauchy Sequences are Convergent *}
+
+text{*Equivalence of Cauchy criterion and convergence:
+  We will prove this using our NS formulation which provides a
+  much easier proof than using the standard definition. We do not
+  need to use properties of subsequences such as boundedness,
+  monotonicity etc... Compare with Harrison's corresponding proof
+  in HOL which is much longer and more complicated. Of course, we do
+  not have problems which he encountered with guessing the right
+  instantiations for his 'espsilon-delta' proof(s) in this case
+  since the NS formulations do not involve existential quantifiers.*}
+
+lemma NSconvergent_NSCauchy: "NSconvergent X \<Longrightarrow> NSCauchy X"
+apply (simp add: NSconvergent_def NSLIMSEQ_def NSCauchy_def, safe)
+apply (auto intro: approx_trans2)
+done
+
+lemma real_NSCauchy_NSconvergent:
+  fixes X :: "nat \<Rightarrow> real"
+  shows "NSCauchy X \<Longrightarrow> NSconvergent X"
+apply (simp add: NSconvergent_def NSLIMSEQ_def)
+apply (frule NSCauchy_NSBseq)
+apply (simp add: NSBseq_def NSCauchy_def)
+apply (drule HNatInfinite_whn [THEN [2] bspec])
+apply (drule HNatInfinite_whn [THEN [2] bspec])
+apply (auto dest!: st_part_Ex simp add: SReal_iff)
+apply (blast intro: approx_trans3)
+done
+
+lemma NSCauchy_NSconvergent:
+  fixes X :: "nat \<Rightarrow> 'a::banach"
+  shows "NSCauchy X \<Longrightarrow> NSconvergent X"
+apply (drule NSCauchy_Cauchy [THEN Cauchy_convergent])
+apply (erule convergent_NSconvergent_iff [THEN iffD1])
+done
+
+lemma NSCauchy_NSconvergent_iff:
+  fixes X :: "nat \<Rightarrow> 'a::banach"
+  shows "NSCauchy X = NSconvergent X"
+by (fast intro: NSCauchy_NSconvergent NSconvergent_NSCauchy)
+
+
+subsection {* Power Sequences *}
+
+text{*The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
+"x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
+  also fact that bounded and monotonic sequence converges.*}
+
+text{* We now use NS criterion to bring proof of theorem through *}
+
+lemma NSLIMSEQ_realpow_zero:
+  "[| 0 \<le> (x::real); x < 1 |] ==> (%n. x ^ n) ----NS> 0"
+apply (simp add: NSLIMSEQ_def)
+apply (auto dest!: convergent_realpow simp add: convergent_NSconvergent_iff)
+apply (frule NSconvergentD)
+apply (auto simp add: NSLIMSEQ_def NSCauchy_NSconvergent_iff [symmetric] NSCauchy_def starfun_pow)
+apply (frule HNatInfinite_add_one)
+apply (drule bspec, assumption)
+apply (drule bspec, assumption)
+apply (drule_tac x = "N + (1::hypnat) " in bspec, assumption)
+apply (simp add: hyperpow_add)
+apply (drule approx_mult_subst_star_of, assumption)
+apply (drule approx_trans3, assumption)
+apply (auto simp del: star_of_mult simp add: star_of_mult [symmetric])
+done
+
+lemma NSLIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < (1::real) ==> (%n. \<bar>c\<bar> ^ n) ----NS> 0"
+by (simp add: LIMSEQ_rabs_realpow_zero LIMSEQ_NSLIMSEQ_iff [symmetric])
+
+lemma NSLIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < (1::real) ==> (%n. c ^ n) ----NS> 0"
+by (simp add: LIMSEQ_rabs_realpow_zero2 LIMSEQ_NSLIMSEQ_iff [symmetric])
+
+(***---------------------------------------------------------------
+    Theorems proved by Harrison in HOL that we do not need
+    in order to prove equivalence between Cauchy criterion
+    and convergence:
+ -- Show that every sequence contains a monotonic subsequence
+Goal "\<exists>f. subseq f & monoseq (%n. s (f n))"
+ -- Show that a subsequence of a bounded sequence is bounded
+Goal "Bseq X ==> Bseq (%n. X (f n))";
+ -- Show we can take subsequential terms arbitrarily far
+    up a sequence
+Goal "subseq f ==> n \<le> f(n)";
+Goal "subseq f ==> \<exists>n. N1 \<le> n & N2 \<le> f(n)";
+ ---------------------------------------------------------------***)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/HSeries.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,201 @@
+(*  Title       : HSeries.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+
+Converted to Isar and polished by lcp    
+*) 
+
+header{*Finite Summation and Infinite Series for Hyperreals*}
+
+theory HSeries
+imports Series HSEQ
+begin
+
+definition
+  sumhr :: "(hypnat * hypnat * (nat=>real)) => hypreal" where
+  [code func del]: "sumhr = 
+      (%(M,N,f). starfun2 (%m n. setsum f {m..<n}) M N)"
+
+definition
+  NSsums  :: "[nat=>real,real] => bool"     (infixr "NSsums" 80) where
+  "f NSsums s = (%n. setsum f {0..<n}) ----NS> s"
+
+definition
+  NSsummable :: "(nat=>real) => bool" where
+  [code func del]: "NSsummable f = (\<exists>s. f NSsums s)"
+
+definition
+  NSsuminf   :: "(nat=>real) => real" where
+  "NSsuminf f = (THE s. f NSsums s)"
+
+lemma sumhr_app: "sumhr(M,N,f) = ( *f2* (\<lambda>m n. setsum f {m..<n})) M N"
+by (simp add: sumhr_def)
+
+text{*Base case in definition of @{term sumr}*}
+lemma sumhr_zero [simp]: "!!m. sumhr (m,0,f) = 0"
+unfolding sumhr_app by transfer simp
+
+text{*Recursive case in definition of @{term sumr}*}
+lemma sumhr_if: 
+     "!!m n. sumhr(m,n+1,f) = 
+      (if n + 1 \<le> m then 0 else sumhr(m,n,f) + ( *f* f) n)"
+unfolding sumhr_app by transfer simp
+
+lemma sumhr_Suc_zero [simp]: "!!n. sumhr (n + 1, n, f) = 0"
+unfolding sumhr_app by transfer simp
+
+lemma sumhr_eq_bounds [simp]: "!!n. sumhr (n,n,f) = 0"
+unfolding sumhr_app by transfer simp
+
+lemma sumhr_Suc [simp]: "!!m. sumhr (m,m + 1,f) = ( *f* f) m"
+unfolding sumhr_app by transfer simp
+
+lemma sumhr_add_lbound_zero [simp]: "!!k m. sumhr(m+k,k,f) = 0"
+unfolding sumhr_app by transfer simp
+
+lemma sumhr_add:
+  "!!m n. sumhr (m,n,f) + sumhr(m,n,g) = sumhr(m,n,%i. f i + g i)"
+unfolding sumhr_app by transfer (rule setsum_addf [symmetric])
+
+lemma sumhr_mult:
+  "!!m n. hypreal_of_real r * sumhr(m,n,f) = sumhr(m,n,%n. r * f n)"
+unfolding sumhr_app by transfer (rule setsum_right_distrib)
+
+lemma sumhr_split_add:
+  "!!n p. n < p ==> sumhr(0,n,f) + sumhr(n,p,f) = sumhr(0,p,f)"
+unfolding sumhr_app by transfer (simp add: setsum_add_nat_ivl)
+
+lemma sumhr_split_diff: "n<p ==> sumhr(0,p,f) - sumhr(0,n,f) = sumhr(n,p,f)"
+by (drule_tac f = f in sumhr_split_add [symmetric], simp)
+
+lemma sumhr_hrabs: "!!m n. abs(sumhr(m,n,f)) \<le> sumhr(m,n,%i. abs(f i))"
+unfolding sumhr_app by transfer (rule setsum_abs)
+
+text{* other general version also needed *}
+lemma sumhr_fun_hypnat_eq:
+   "(\<forall>r. m \<le> r & r < n --> f r = g r) -->  
+      sumhr(hypnat_of_nat m, hypnat_of_nat n, f) =  
+      sumhr(hypnat_of_nat m, hypnat_of_nat n, g)"
+unfolding sumhr_app by transfer simp
+
+lemma sumhr_const:
+     "!!n. sumhr(0, n, %i. r) = hypreal_of_hypnat n * hypreal_of_real r"
+unfolding sumhr_app by transfer (simp add: real_of_nat_def)
+
+lemma sumhr_less_bounds_zero [simp]: "!!m n. n < m ==> sumhr(m,n,f) = 0"
+unfolding sumhr_app by transfer simp
+
+lemma sumhr_minus: "!!m n. sumhr(m, n, %i. - f i) = - sumhr(m, n, f)"
+unfolding sumhr_app by transfer (rule setsum_negf)
+
+lemma sumhr_shift_bounds:
+  "!!m n. sumhr(m+hypnat_of_nat k,n+hypnat_of_nat k,f) =
+          sumhr(m,n,%i. f(i + k))"
+unfolding sumhr_app by transfer (rule setsum_shift_bounds_nat_ivl)
+
+
+subsection{*Nonstandard Sums*}
+
+text{*Infinite sums are obtained by summing to some infinite hypernatural
+ (such as @{term whn})*}
+lemma sumhr_hypreal_of_hypnat_omega: 
+      "sumhr(0,whn,%i. 1) = hypreal_of_hypnat whn"
+by (simp add: sumhr_const)
+
+lemma sumhr_hypreal_omega_minus_one: "sumhr(0, whn, %i. 1) = omega - 1"
+apply (simp add: sumhr_const)
+(* FIXME: need lemma: hypreal_of_hypnat whn = omega - 1 *)
+(* maybe define omega = hypreal_of_hypnat whn + 1 *)
+apply (unfold star_class_defs omega_def hypnat_omega_def
+              of_hypnat_def star_of_def)
+apply (simp add: starfun_star_n starfun2_star_n real_of_nat_def)
+done
+
+lemma sumhr_minus_one_realpow_zero [simp]: 
+     "!!N. sumhr(0, N + N, %i. (-1) ^ (i+1)) = 0"
+unfolding sumhr_app
+by transfer (simp del: realpow_Suc add: nat_mult_2 [symmetric])
+
+lemma sumhr_interval_const:
+     "(\<forall>n. m \<le> Suc n --> f n = r) & m \<le> na  
+      ==> sumhr(hypnat_of_nat m,hypnat_of_nat na,f) =  
+          (hypreal_of_nat (na - m) * hypreal_of_real r)"
+unfolding sumhr_app by transfer simp
+
+lemma starfunNat_sumr: "!!N. ( *f* (%n. setsum f {0..<n})) N = sumhr(0,N,f)"
+unfolding sumhr_app by transfer (rule refl)
+
+lemma sumhr_hrabs_approx [simp]: "sumhr(0, M, f) @= sumhr(0, N, f)  
+      ==> abs (sumhr(M, N, f)) @= 0"
+apply (cut_tac x = M and y = N in linorder_less_linear)
+apply (auto simp add: approx_refl)
+apply (drule approx_sym [THEN approx_minus_iff [THEN iffD1]])
+apply (auto dest: approx_hrabs 
+            simp add: sumhr_split_diff diff_minus [symmetric])
+done
+
+(*----------------------------------------------------------------
+      infinite sums: Standard and NS theorems
+ ----------------------------------------------------------------*)
+lemma sums_NSsums_iff: "(f sums l) = (f NSsums l)"
+by (simp add: sums_def NSsums_def LIMSEQ_NSLIMSEQ_iff)
+
+lemma summable_NSsummable_iff: "(summable f) = (NSsummable f)"
+by (simp add: summable_def NSsummable_def sums_NSsums_iff)
+
+lemma suminf_NSsuminf_iff: "(suminf f) = (NSsuminf f)"
+by (simp add: suminf_def NSsuminf_def sums_NSsums_iff)
+
+lemma NSsums_NSsummable: "f NSsums l ==> NSsummable f"
+by (simp add: NSsums_def NSsummable_def, blast)
+
+lemma NSsummable_NSsums: "NSsummable f ==> f NSsums (NSsuminf f)"
+apply (simp add: NSsummable_def NSsuminf_def NSsums_def)
+apply (blast intro: theI NSLIMSEQ_unique)
+done
+
+lemma NSsums_unique: "f NSsums s ==> (s = NSsuminf f)"
+by (simp add: suminf_NSsuminf_iff [symmetric] sums_NSsums_iff sums_unique)
+
+lemma NSseries_zero:
+  "\<forall>m. n \<le> Suc m --> f(m) = 0 ==> f NSsums (setsum f {0..<n})"
+by (simp add: sums_NSsums_iff [symmetric] series_zero)
+
+lemma NSsummable_NSCauchy:
+     "NSsummable f =  
+      (\<forall>M \<in> HNatInfinite. \<forall>N \<in> HNatInfinite. abs (sumhr(M,N,f)) @= 0)"
+apply (auto simp add: summable_NSsummable_iff [symmetric] 
+       summable_convergent_sumr_iff convergent_NSconvergent_iff 
+       NSCauchy_NSconvergent_iff [symmetric] NSCauchy_def starfunNat_sumr)
+apply (cut_tac x = M and y = N in linorder_less_linear)
+apply (auto simp add: approx_refl)
+apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])
+apply (rule_tac [2] approx_minus_iff [THEN iffD2])
+apply (auto dest: approx_hrabs_zero_cancel 
+            simp add: sumhr_split_diff diff_minus [symmetric])
+done
+
+
+text{*Terms of a convergent series tend to zero*}
+lemma NSsummable_NSLIMSEQ_zero: "NSsummable f ==> f ----NS> 0"
+apply (auto simp add: NSLIMSEQ_def NSsummable_NSCauchy)
+apply (drule bspec, auto)
+apply (drule_tac x = "N + 1 " in bspec)
+apply (auto intro: HNatInfinite_add_one approx_hrabs_zero_cancel)
+done
+
+text{*Nonstandard comparison test*}
+lemma NSsummable_comparison_test:
+     "[| \<exists>N. \<forall>n. N \<le> n --> abs(f n) \<le> g n; NSsummable g |] ==> NSsummable f"
+apply (fold summable_NSsummable_iff)
+apply (rule summable_comparison_test, simp, assumption)
+done
+
+lemma NSsummable_rabs_comparison_test:
+     "[| \<exists>N. \<forall>n. N \<le> n --> abs(f n) \<le> g n; NSsummable g |]
+      ==> NSsummable (%k. abs (f k))"
+apply (rule NSsummable_comparison_test)
+apply (auto)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/HTranscendental.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,608 @@
+(*  Title       : HTranscendental.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 2001 University of Edinburgh
+
+Converted to Isar and polished by lcp
+*)
+
+header{*Nonstandard Extensions of Transcendental Functions*}
+
+theory HTranscendental
+imports Transcendental HSeries HDeriv
+begin
+
+definition
+  exphr :: "real => hypreal" where
+    --{*define exponential function using standard part *}
+  "exphr x =  st(sumhr (0, whn, %n. inverse(real (fact n)) * (x ^ n)))"
+
+definition
+  sinhr :: "real => hypreal" where
+  "sinhr x = st(sumhr (0, whn, %n. (if even(n) then 0 else
+             ((-1) ^ ((n - 1) div 2))/(real (fact n))) * (x ^ n)))"
+  
+definition
+  coshr :: "real => hypreal" where
+  "coshr x = st(sumhr (0, whn, %n. (if even(n) then
+            ((-1) ^ (n div 2))/(real (fact n)) else 0) * x ^ n))"
+
+
+subsection{*Nonstandard Extension of Square Root Function*}
+
+lemma STAR_sqrt_zero [simp]: "( *f* sqrt) 0 = 0"
+by (simp add: starfun star_n_zero_num)
+
+lemma STAR_sqrt_one [simp]: "( *f* sqrt) 1 = 1"
+by (simp add: starfun star_n_one_num)
+
+lemma hypreal_sqrt_pow2_iff: "(( *f* sqrt)(x) ^ 2 = x) = (0 \<le> x)"
+apply (cases x)
+apply (auto simp add: star_n_le star_n_zero_num starfun hrealpow star_n_eq_iff
+            simp del: hpowr_Suc realpow_Suc)
+done
+
+lemma hypreal_sqrt_gt_zero_pow2: "!!x. 0 < x ==> ( *f* sqrt) (x) ^ 2 = x"
+by (transfer, simp)
+
+lemma hypreal_sqrt_pow2_gt_zero: "0 < x ==> 0 < ( *f* sqrt) (x) ^ 2"
+by (frule hypreal_sqrt_gt_zero_pow2, auto)
+
+lemma hypreal_sqrt_not_zero: "0 < x ==> ( *f* sqrt) (x) \<noteq> 0"
+apply (frule hypreal_sqrt_pow2_gt_zero)
+apply (auto simp add: numeral_2_eq_2)
+done
+
+lemma hypreal_inverse_sqrt_pow2:
+     "0 < x ==> inverse (( *f* sqrt)(x)) ^ 2 = inverse x"
+apply (cut_tac n = 2 and a = "( *f* sqrt) x" in power_inverse [symmetric])
+apply (auto dest: hypreal_sqrt_gt_zero_pow2)
+done
+
+lemma hypreal_sqrt_mult_distrib: 
+    "!!x y. [|0 < x; 0 <y |] ==>
+      ( *f* sqrt)(x*y) = ( *f* sqrt)(x) * ( *f* sqrt)(y)"
+apply transfer
+apply (auto intro: real_sqrt_mult_distrib) 
+done
+
+lemma hypreal_sqrt_mult_distrib2:
+     "[|0\<le>x; 0\<le>y |] ==>  
+     ( *f* sqrt)(x*y) =  ( *f* sqrt)(x) * ( *f* sqrt)(y)"
+by (auto intro: hypreal_sqrt_mult_distrib simp add: order_le_less)
+
+lemma hypreal_sqrt_approx_zero [simp]:
+     "0 < x ==> (( *f* sqrt)(x) @= 0) = (x @= 0)"
+apply (auto simp add: mem_infmal_iff [symmetric])
+apply (rule hypreal_sqrt_gt_zero_pow2 [THEN subst])
+apply (auto intro: Infinitesimal_mult 
+            dest!: hypreal_sqrt_gt_zero_pow2 [THEN ssubst] 
+            simp add: numeral_2_eq_2)
+done
+
+lemma hypreal_sqrt_approx_zero2 [simp]:
+     "0 \<le> x ==> (( *f* sqrt)(x) @= 0) = (x @= 0)"
+by (auto simp add: order_le_less)
+
+lemma hypreal_sqrt_sum_squares [simp]:
+     "(( *f* sqrt)(x*x + y*y + z*z) @= 0) = (x*x + y*y + z*z @= 0)"
+apply (rule hypreal_sqrt_approx_zero2)
+apply (rule add_nonneg_nonneg)+
+apply (auto)
+done
+
+lemma hypreal_sqrt_sum_squares2 [simp]:
+     "(( *f* sqrt)(x*x + y*y) @= 0) = (x*x + y*y @= 0)"
+apply (rule hypreal_sqrt_approx_zero2)
+apply (rule add_nonneg_nonneg)
+apply (auto)
+done
+
+lemma hypreal_sqrt_gt_zero: "!!x. 0 < x ==> 0 < ( *f* sqrt)(x)"
+apply transfer
+apply (auto intro: real_sqrt_gt_zero)
+done
+
+lemma hypreal_sqrt_ge_zero: "0 \<le> x ==> 0 \<le> ( *f* sqrt)(x)"
+by (auto intro: hypreal_sqrt_gt_zero simp add: order_le_less)
+
+lemma hypreal_sqrt_hrabs [simp]: "!!x. ( *f* sqrt)(x ^ 2) = abs(x)"
+by (transfer, simp)
+
+lemma hypreal_sqrt_hrabs2 [simp]: "!!x. ( *f* sqrt)(x*x) = abs(x)"
+by (transfer, simp)
+
+lemma hypreal_sqrt_hyperpow_hrabs [simp]:
+     "!!x. ( *f* sqrt)(x pow (hypnat_of_nat 2)) = abs(x)"
+by (transfer, simp)
+
+lemma star_sqrt_HFinite: "\<lbrakk>x \<in> HFinite; 0 \<le> x\<rbrakk> \<Longrightarrow> ( *f* sqrt) x \<in> HFinite"
+apply (rule HFinite_square_iff [THEN iffD1])
+apply (simp only: hypreal_sqrt_mult_distrib2 [symmetric], simp) 
+done
+
+lemma st_hypreal_sqrt:
+     "[| x \<in> HFinite; 0 \<le> x |] ==> st(( *f* sqrt) x) = ( *f* sqrt)(st x)"
+apply (rule power_inject_base [where n=1])
+apply (auto intro!: st_zero_le hypreal_sqrt_ge_zero)
+apply (rule st_mult [THEN subst])
+apply (rule_tac [3] hypreal_sqrt_mult_distrib2 [THEN subst])
+apply (rule_tac [5] hypreal_sqrt_mult_distrib2 [THEN subst])
+apply (auto simp add: st_hrabs st_zero_le star_sqrt_HFinite)
+done
+
+lemma hypreal_sqrt_sum_squares_ge1 [simp]: "!!x y. x \<le> ( *f* sqrt)(x ^ 2 + y ^ 2)"
+by transfer (rule real_sqrt_sum_squares_ge1)
+
+lemma HFinite_hypreal_sqrt:
+     "[| 0 \<le> x; x \<in> HFinite |] ==> ( *f* sqrt) x \<in> HFinite"
+apply (auto simp add: order_le_less)
+apply (rule HFinite_square_iff [THEN iffD1])
+apply (drule hypreal_sqrt_gt_zero_pow2)
+apply (simp add: numeral_2_eq_2)
+done
+
+lemma HFinite_hypreal_sqrt_imp_HFinite:
+     "[| 0 \<le> x; ( *f* sqrt) x \<in> HFinite |] ==> x \<in> HFinite"
+apply (auto simp add: order_le_less)
+apply (drule HFinite_square_iff [THEN iffD2])
+apply (drule hypreal_sqrt_gt_zero_pow2)
+apply (simp add: numeral_2_eq_2 del: HFinite_square_iff)
+done
+
+lemma HFinite_hypreal_sqrt_iff [simp]:
+     "0 \<le> x ==> (( *f* sqrt) x \<in> HFinite) = (x \<in> HFinite)"
+by (blast intro: HFinite_hypreal_sqrt HFinite_hypreal_sqrt_imp_HFinite)
+
+lemma HFinite_sqrt_sum_squares [simp]:
+     "(( *f* sqrt)(x*x + y*y) \<in> HFinite) = (x*x + y*y \<in> HFinite)"
+apply (rule HFinite_hypreal_sqrt_iff)
+apply (rule add_nonneg_nonneg)
+apply (auto)
+done
+
+lemma Infinitesimal_hypreal_sqrt:
+     "[| 0 \<le> x; x \<in> Infinitesimal |] ==> ( *f* sqrt) x \<in> Infinitesimal"
+apply (auto simp add: order_le_less)
+apply (rule Infinitesimal_square_iff [THEN iffD2])
+apply (drule hypreal_sqrt_gt_zero_pow2)
+apply (simp add: numeral_2_eq_2)
+done
+
+lemma Infinitesimal_hypreal_sqrt_imp_Infinitesimal:
+     "[| 0 \<le> x; ( *f* sqrt) x \<in> Infinitesimal |] ==> x \<in> Infinitesimal"
+apply (auto simp add: order_le_less)
+apply (drule Infinitesimal_square_iff [THEN iffD1])
+apply (drule hypreal_sqrt_gt_zero_pow2)
+apply (simp add: numeral_2_eq_2 del: Infinitesimal_square_iff [symmetric])
+done
+
+lemma Infinitesimal_hypreal_sqrt_iff [simp]:
+     "0 \<le> x ==> (( *f* sqrt) x \<in> Infinitesimal) = (x \<in> Infinitesimal)"
+by (blast intro: Infinitesimal_hypreal_sqrt_imp_Infinitesimal Infinitesimal_hypreal_sqrt)
+
+lemma Infinitesimal_sqrt_sum_squares [simp]:
+     "(( *f* sqrt)(x*x + y*y) \<in> Infinitesimal) = (x*x + y*y \<in> Infinitesimal)"
+apply (rule Infinitesimal_hypreal_sqrt_iff)
+apply (rule add_nonneg_nonneg)
+apply (auto)
+done
+
+lemma HInfinite_hypreal_sqrt:
+     "[| 0 \<le> x; x \<in> HInfinite |] ==> ( *f* sqrt) x \<in> HInfinite"
+apply (auto simp add: order_le_less)
+apply (rule HInfinite_square_iff [THEN iffD1])
+apply (drule hypreal_sqrt_gt_zero_pow2)
+apply (simp add: numeral_2_eq_2)
+done
+
+lemma HInfinite_hypreal_sqrt_imp_HInfinite:
+     "[| 0 \<le> x; ( *f* sqrt) x \<in> HInfinite |] ==> x \<in> HInfinite"
+apply (auto simp add: order_le_less)
+apply (drule HInfinite_square_iff [THEN iffD2])
+apply (drule hypreal_sqrt_gt_zero_pow2)
+apply (simp add: numeral_2_eq_2 del: HInfinite_square_iff)
+done
+
+lemma HInfinite_hypreal_sqrt_iff [simp]:
+     "0 \<le> x ==> (( *f* sqrt) x \<in> HInfinite) = (x \<in> HInfinite)"
+by (blast intro: HInfinite_hypreal_sqrt HInfinite_hypreal_sqrt_imp_HInfinite)
+
+lemma HInfinite_sqrt_sum_squares [simp]:
+     "(( *f* sqrt)(x*x + y*y) \<in> HInfinite) = (x*x + y*y \<in> HInfinite)"
+apply (rule HInfinite_hypreal_sqrt_iff)
+apply (rule add_nonneg_nonneg)
+apply (auto)
+done
+
+lemma HFinite_exp [simp]:
+     "sumhr (0, whn, %n. inverse (real (fact n)) * x ^ n) \<in> HFinite"
+unfolding sumhr_app
+apply (simp only: star_zero_def starfun2_star_of)
+apply (rule NSBseqD2)
+apply (rule NSconvergent_NSBseq)
+apply (rule convergent_NSconvergent_iff [THEN iffD1])
+apply (rule summable_convergent_sumr_iff [THEN iffD1])
+apply (rule summable_exp)
+done
+
+lemma exphr_zero [simp]: "exphr 0 = 1"
+apply (simp add: exphr_def sumhr_split_add
+                   [OF hypnat_one_less_hypnat_omega, symmetric])
+apply (rule st_unique, simp)
+apply (rule subst [where P="\<lambda>x. 1 \<approx> x", OF _ approx_refl])
+apply (rule rev_mp [OF hypnat_one_less_hypnat_omega])
+apply (rule_tac x="whn" in spec)
+apply (unfold sumhr_app, transfer, simp)
+done
+
+lemma coshr_zero [simp]: "coshr 0 = 1"
+apply (simp add: coshr_def sumhr_split_add
+                   [OF hypnat_one_less_hypnat_omega, symmetric]) 
+apply (rule st_unique, simp)
+apply (rule subst [where P="\<lambda>x. 1 \<approx> x", OF _ approx_refl])
+apply (rule rev_mp [OF hypnat_one_less_hypnat_omega])
+apply (rule_tac x="whn" in spec)
+apply (unfold sumhr_app, transfer, simp)
+done
+
+lemma STAR_exp_zero_approx_one [simp]: "( *f* exp) (0::hypreal) @= 1"
+apply (subgoal_tac "( *f* exp) (0::hypreal) = 1", simp)
+apply (transfer, simp)
+done
+
+lemma STAR_exp_Infinitesimal: "x \<in> Infinitesimal ==> ( *f* exp) (x::hypreal) @= 1"
+apply (case_tac "x = 0")
+apply (cut_tac [2] x = 0 in DERIV_exp)
+apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
+apply (drule_tac x = x in bspec, auto)
+apply (drule_tac c = x in approx_mult1)
+apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] 
+            simp add: mult_assoc)
+apply (rule approx_add_right_cancel [where d="-1"])
+apply (rule approx_sym [THEN [2] approx_trans2])
+apply (auto simp add: diff_def mem_infmal_iff)
+done
+
+lemma STAR_exp_epsilon [simp]: "( *f* exp) epsilon @= 1"
+by (auto intro: STAR_exp_Infinitesimal)
+
+lemma STAR_exp_add: "!!x y. ( *f* exp)(x + y) = ( *f* exp) x * ( *f* exp) y"
+by transfer (rule exp_add)
+
+lemma exphr_hypreal_of_real_exp_eq: "exphr x = hypreal_of_real (exp x)"
+apply (simp add: exphr_def)
+apply (rule st_unique, simp)
+apply (subst starfunNat_sumr [symmetric])
+apply (rule NSLIMSEQ_D [THEN approx_sym])
+apply (rule LIMSEQ_NSLIMSEQ)
+apply (subst sums_def [symmetric])
+apply (cut_tac exp_converges [where x=x], simp)
+apply (rule HNatInfinite_whn)
+done
+
+lemma starfun_exp_ge_add_one_self [simp]: "!!x::hypreal. 0 \<le> x ==> (1 + x) \<le> ( *f* exp) x"
+by transfer (rule exp_ge_add_one_self_aux)
+
+(* exp (oo) is infinite *)
+lemma starfun_exp_HInfinite:
+     "[| x \<in> HInfinite; 0 \<le> x |] ==> ( *f* exp) (x::hypreal) \<in> HInfinite"
+apply (frule starfun_exp_ge_add_one_self)
+apply (rule HInfinite_ge_HInfinite, assumption)
+apply (rule order_trans [of _ "1+x"], auto) 
+done
+
+lemma starfun_exp_minus: "!!x. ( *f* exp) (-x) = inverse(( *f* exp) x)"
+by transfer (rule exp_minus)
+
+(* exp (-oo) is infinitesimal *)
+lemma starfun_exp_Infinitesimal:
+     "[| x \<in> HInfinite; x \<le> 0 |] ==> ( *f* exp) (x::hypreal) \<in> Infinitesimal"
+apply (subgoal_tac "\<exists>y. x = - y")
+apply (rule_tac [2] x = "- x" in exI)
+apply (auto intro!: HInfinite_inverse_Infinitesimal starfun_exp_HInfinite
+            simp add: starfun_exp_minus HInfinite_minus_iff)
+done
+
+lemma starfun_exp_gt_one [simp]: "!!x::hypreal. 0 < x ==> 1 < ( *f* exp) x"
+by transfer (rule exp_gt_one)
+
+lemma starfun_ln_exp [simp]: "!!x. ( *f* ln) (( *f* exp) x) = x"
+by transfer (rule ln_exp)
+
+lemma starfun_exp_ln_iff [simp]: "!!x. (( *f* exp)(( *f* ln) x) = x) = (0 < x)"
+by transfer (rule exp_ln_iff)
+
+lemma starfun_exp_ln_eq: "!!u x. ( *f* exp) u = x ==> ( *f* ln) x = u"
+by transfer (rule exp_ln_eq)
+
+lemma starfun_ln_less_self [simp]: "!!x. 0 < x ==> ( *f* ln) x < x"
+by transfer (rule ln_less_self)
+
+lemma starfun_ln_ge_zero [simp]: "!!x. 1 \<le> x ==> 0 \<le> ( *f* ln) x"
+by transfer (rule ln_ge_zero)
+
+lemma starfun_ln_gt_zero [simp]: "!!x .1 < x ==> 0 < ( *f* ln) x"
+by transfer (rule ln_gt_zero)
+
+lemma starfun_ln_not_eq_zero [simp]: "!!x. [| 0 < x; x \<noteq> 1 |] ==> ( *f* ln) x \<noteq> 0"
+by transfer simp
+
+lemma starfun_ln_HFinite: "[| x \<in> HFinite; 1 \<le> x |] ==> ( *f* ln) x \<in> HFinite"
+apply (rule HFinite_bounded)
+apply assumption 
+apply (simp_all add: starfun_ln_less_self order_less_imp_le)
+done
+
+lemma starfun_ln_inverse: "!!x. 0 < x ==> ( *f* ln) (inverse x) = -( *f* ln) x"
+by transfer (rule ln_inverse)
+
+lemma starfun_abs_exp_cancel: "\<And>x. \<bar>( *f* exp) (x::hypreal)\<bar> = ( *f* exp) x"
+by transfer (rule abs_exp_cancel)
+
+lemma starfun_exp_less_mono: "\<And>x y::hypreal. x < y \<Longrightarrow> ( *f* exp) x < ( *f* exp) y"
+by transfer (rule exp_less_mono)
+
+lemma starfun_exp_HFinite: "x \<in> HFinite ==> ( *f* exp) (x::hypreal) \<in> HFinite"
+apply (auto simp add: HFinite_def, rename_tac u)
+apply (rule_tac x="( *f* exp) u" in rev_bexI)
+apply (simp add: Reals_eq_Standard)
+apply (simp add: starfun_abs_exp_cancel)
+apply (simp add: starfun_exp_less_mono)
+done
+
+lemma starfun_exp_add_HFinite_Infinitesimal_approx:
+     "[|x \<in> Infinitesimal; z \<in> HFinite |] ==> ( *f* exp) (z + x::hypreal) @= ( *f* exp) z"
+apply (simp add: STAR_exp_add)
+apply (frule STAR_exp_Infinitesimal)
+apply (drule approx_mult2)
+apply (auto intro: starfun_exp_HFinite)
+done
+
+(* using previous result to get to result *)
+lemma starfun_ln_HInfinite:
+     "[| x \<in> HInfinite; 0 < x |] ==> ( *f* ln) x \<in> HInfinite"
+apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])
+apply (drule starfun_exp_HFinite)
+apply (simp add: starfun_exp_ln_iff [THEN iffD2] HFinite_HInfinite_iff)
+done
+
+lemma starfun_exp_HInfinite_Infinitesimal_disj:
+ "x \<in> HInfinite ==> ( *f* exp) x \<in> HInfinite | ( *f* exp) (x::hypreal) \<in> Infinitesimal"
+apply (insert linorder_linear [of x 0]) 
+apply (auto intro: starfun_exp_HInfinite starfun_exp_Infinitesimal)
+done
+
+(* check out this proof!!! *)
+lemma starfun_ln_HFinite_not_Infinitesimal:
+     "[| x \<in> HFinite - Infinitesimal; 0 < x |] ==> ( *f* ln) x \<in> HFinite"
+apply (rule ccontr, drule HInfinite_HFinite_iff [THEN iffD2])
+apply (drule starfun_exp_HInfinite_Infinitesimal_disj)
+apply (simp add: starfun_exp_ln_iff [symmetric] HInfinite_HFinite_iff
+            del: starfun_exp_ln_iff)
+done
+
+(* we do proof by considering ln of 1/x *)
+lemma starfun_ln_Infinitesimal_HInfinite:
+     "[| x \<in> Infinitesimal; 0 < x |] ==> ( *f* ln) x \<in> HInfinite"
+apply (drule Infinitesimal_inverse_HInfinite)
+apply (frule positive_imp_inverse_positive)
+apply (drule_tac [2] starfun_ln_HInfinite)
+apply (auto simp add: starfun_ln_inverse HInfinite_minus_iff)
+done
+
+lemma starfun_ln_less_zero: "!!x. [| 0 < x; x < 1 |] ==> ( *f* ln) x < 0"
+by transfer (rule ln_less_zero)
+
+lemma starfun_ln_Infinitesimal_less_zero:
+     "[| x \<in> Infinitesimal; 0 < x |] ==> ( *f* ln) x < 0"
+by (auto intro!: starfun_ln_less_zero simp add: Infinitesimal_def)
+
+lemma starfun_ln_HInfinite_gt_zero:
+     "[| x \<in> HInfinite; 0 < x |] ==> 0 < ( *f* ln) x"
+by (auto intro!: starfun_ln_gt_zero simp add: HInfinite_def)
+
+
+(*
+Goalw [NSLIM_def] "(%h. ((x powr h) - 1) / h) -- 0 --NS> ln x"
+*)
+
+lemma HFinite_sin [simp]:
+     "sumhr (0, whn, %n. (if even(n) then 0 else  
+              (-1 ^ ((n - 1) div 2))/(real (fact n))) * x ^ n)  
+              \<in> HFinite"
+unfolding sumhr_app
+apply (simp only: star_zero_def starfun2_star_of)
+apply (rule NSBseqD2)
+apply (rule NSconvergent_NSBseq)
+apply (rule convergent_NSconvergent_iff [THEN iffD1])
+apply (rule summable_convergent_sumr_iff [THEN iffD1])
+apply (simp only: One_nat_def summable_sin)
+done
+
+lemma STAR_sin_zero [simp]: "( *f* sin) 0 = 0"
+by transfer (rule sin_zero)
+
+lemma STAR_sin_Infinitesimal [simp]: "x \<in> Infinitesimal ==> ( *f* sin) x @= x"
+apply (case_tac "x = 0")
+apply (cut_tac [2] x = 0 in DERIV_sin)
+apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
+apply (drule bspec [where x = x], auto)
+apply (drule approx_mult1 [where c = x])
+apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
+           simp add: mult_assoc)
+done
+
+lemma HFinite_cos [simp]:
+     "sumhr (0, whn, %n. (if even(n) then  
+            (-1 ^ (n div 2))/(real (fact n)) else  
+            0) * x ^ n) \<in> HFinite"
+unfolding sumhr_app
+apply (simp only: star_zero_def starfun2_star_of)
+apply (rule NSBseqD2)
+apply (rule NSconvergent_NSBseq)
+apply (rule convergent_NSconvergent_iff [THEN iffD1])
+apply (rule summable_convergent_sumr_iff [THEN iffD1])
+apply (rule summable_cos)
+done
+
+lemma STAR_cos_zero [simp]: "( *f* cos) 0 = 1"
+by transfer (rule cos_zero)
+
+lemma STAR_cos_Infinitesimal [simp]: "x \<in> Infinitesimal ==> ( *f* cos) x @= 1"
+apply (case_tac "x = 0")
+apply (cut_tac [2] x = 0 in DERIV_cos)
+apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
+apply (drule bspec [where x = x])
+apply auto
+apply (drule approx_mult1 [where c = x])
+apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
+            simp add: mult_assoc)
+apply (rule approx_add_right_cancel [where d = "-1"])
+apply (simp add: diff_def)
+done
+
+lemma STAR_tan_zero [simp]: "( *f* tan) 0 = 0"
+by transfer (rule tan_zero)
+
+lemma STAR_tan_Infinitesimal: "x \<in> Infinitesimal ==> ( *f* tan) x @= x"
+apply (case_tac "x = 0")
+apply (cut_tac [2] x = 0 in DERIV_tan)
+apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
+apply (drule bspec [where x = x], auto)
+apply (drule approx_mult1 [where c = x])
+apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
+             simp add: mult_assoc)
+done
+
+lemma STAR_sin_cos_Infinitesimal_mult:
+     "x \<in> Infinitesimal ==> ( *f* sin) x * ( *f* cos) x @= x"
+apply (insert approx_mult_HFinite [of "( *f* sin) x" _ "( *f* cos) x" 1]) 
+apply (simp add: Infinitesimal_subset_HFinite [THEN subsetD])
+done
+
+lemma HFinite_pi: "hypreal_of_real pi \<in> HFinite"
+by simp
+
+(* lemmas *)
+
+lemma lemma_split_hypreal_of_real:
+     "N \<in> HNatInfinite  
+      ==> hypreal_of_real a =  
+          hypreal_of_hypnat N * (inverse(hypreal_of_hypnat N) * hypreal_of_real a)"
+by (simp add: mult_assoc [symmetric] zero_less_HNatInfinite)
+
+lemma STAR_sin_Infinitesimal_divide:
+     "[|x \<in> Infinitesimal; x \<noteq> 0 |] ==> ( *f* sin) x/x @= 1"
+apply (cut_tac x = 0 in DERIV_sin)
+apply (simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
+done
+
+(*------------------------------------------------------------------------*) 
+(* sin* (1/n) * 1/(1/n) @= 1 for n = oo                                   *)
+(*------------------------------------------------------------------------*)
+
+lemma lemma_sin_pi:
+     "n \<in> HNatInfinite  
+      ==> ( *f* sin) (inverse (hypreal_of_hypnat n))/(inverse (hypreal_of_hypnat n)) @= 1"
+apply (rule STAR_sin_Infinitesimal_divide)
+apply (auto simp add: zero_less_HNatInfinite)
+done
+
+lemma STAR_sin_inverse_HNatInfinite:
+     "n \<in> HNatInfinite  
+      ==> ( *f* sin) (inverse (hypreal_of_hypnat n)) * hypreal_of_hypnat n @= 1"
+apply (frule lemma_sin_pi)
+apply (simp add: divide_inverse)
+done
+
+lemma Infinitesimal_pi_divide_HNatInfinite: 
+     "N \<in> HNatInfinite  
+      ==> hypreal_of_real pi/(hypreal_of_hypnat N) \<in> Infinitesimal"
+apply (simp add: divide_inverse)
+apply (auto intro: Infinitesimal_HFinite_mult2)
+done
+
+lemma pi_divide_HNatInfinite_not_zero [simp]:
+     "N \<in> HNatInfinite ==> hypreal_of_real pi/(hypreal_of_hypnat N) \<noteq> 0"
+by (simp add: zero_less_HNatInfinite)
+
+lemma STAR_sin_pi_divide_HNatInfinite_approx_pi:
+     "n \<in> HNatInfinite  
+      ==> ( *f* sin) (hypreal_of_real pi/(hypreal_of_hypnat n)) * hypreal_of_hypnat n  
+          @= hypreal_of_real pi"
+apply (frule STAR_sin_Infinitesimal_divide
+               [OF Infinitesimal_pi_divide_HNatInfinite 
+                   pi_divide_HNatInfinite_not_zero])
+apply (auto)
+apply (rule approx_SReal_mult_cancel [of "inverse (hypreal_of_real pi)"])
+apply (auto intro: Reals_inverse simp add: divide_inverse mult_ac)
+done
+
+lemma STAR_sin_pi_divide_HNatInfinite_approx_pi2:
+     "n \<in> HNatInfinite  
+      ==> hypreal_of_hypnat n *  
+          ( *f* sin) (hypreal_of_real pi/(hypreal_of_hypnat n))  
+          @= hypreal_of_real pi"
+apply (rule mult_commute [THEN subst])
+apply (erule STAR_sin_pi_divide_HNatInfinite_approx_pi)
+done
+
+lemma starfunNat_pi_divide_n_Infinitesimal: 
+     "N \<in> HNatInfinite ==> ( *f* (%x. pi / real x)) N \<in> Infinitesimal"
+by (auto intro!: Infinitesimal_HFinite_mult2 
+         simp add: starfun_mult [symmetric] divide_inverse
+                   starfun_inverse [symmetric] starfunNat_real_of_nat)
+
+lemma STAR_sin_pi_divide_n_approx:
+     "N \<in> HNatInfinite ==>  
+      ( *f* sin) (( *f* (%x. pi / real x)) N) @=  
+      hypreal_of_real pi/(hypreal_of_hypnat N)"
+apply (simp add: starfunNat_real_of_nat [symmetric])
+apply (rule STAR_sin_Infinitesimal)
+apply (simp add: divide_inverse)
+apply (rule Infinitesimal_HFinite_mult2)
+apply (subst starfun_inverse)
+apply (erule starfunNat_inverse_real_of_nat_Infinitesimal)
+apply simp
+done
+
+lemma NSLIMSEQ_sin_pi: "(%n. real n * sin (pi / real n)) ----NS> pi"
+apply (auto simp add: NSLIMSEQ_def starfun_mult [symmetric] starfunNat_real_of_nat)
+apply (rule_tac f1 = sin in starfun_o2 [THEN subst])
+apply (auto simp add: starfun_mult [symmetric] starfunNat_real_of_nat divide_inverse)
+apply (rule_tac f1 = inverse in starfun_o2 [THEN subst])
+apply (auto dest: STAR_sin_pi_divide_HNatInfinite_approx_pi 
+            simp add: starfunNat_real_of_nat mult_commute divide_inverse)
+done
+
+lemma NSLIMSEQ_cos_one: "(%n. cos (pi / real n))----NS> 1"
+apply (simp add: NSLIMSEQ_def, auto)
+apply (rule_tac f1 = cos in starfun_o2 [THEN subst])
+apply (rule STAR_cos_Infinitesimal)
+apply (auto intro!: Infinitesimal_HFinite_mult2 
+            simp add: starfun_mult [symmetric] divide_inverse
+                      starfun_inverse [symmetric] starfunNat_real_of_nat)
+done
+
+lemma NSLIMSEQ_sin_cos_pi:
+     "(%n. real n * sin (pi / real n) * cos (pi / real n)) ----NS> pi"
+by (insert NSLIMSEQ_mult [OF NSLIMSEQ_sin_pi NSLIMSEQ_cos_one], simp)
+
+
+text{*A familiar approximation to @{term "cos x"} when @{term x} is small*}
+
+lemma STAR_cos_Infinitesimal_approx:
+     "x \<in> Infinitesimal ==> ( *f* cos) x @= 1 - x ^ 2"
+apply (rule STAR_cos_Infinitesimal [THEN approx_trans])
+apply (auto simp add: Infinitesimal_approx_minus [symmetric] 
+            diff_minus add_assoc [symmetric] numeral_2_eq_2)
+done
+
+lemma STAR_cos_Infinitesimal_approx2:
+     "x \<in> Infinitesimal ==> ( *f* cos) x @= 1 - (x ^ 2)/2"
+apply (rule STAR_cos_Infinitesimal [THEN approx_trans])
+apply (auto intro: Infinitesimal_SReal_divide 
+            simp add: Infinitesimal_approx_minus [symmetric] numeral_2_eq_2)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/HyperDef.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,570 @@
+(*  Title       : HOL/Hyperreal/HyperDef.thy
+    ID          : $Id$
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
+*)
+
+header{*Construction of Hyperreals Using Ultrafilters*}
+
+theory HyperDef
+imports HyperNat "../Real/Real"
+uses ("hypreal_arith.ML")
+begin
+
+types hypreal = "real star"
+
+abbreviation
+  hypreal_of_real :: "real => real star" where
+  "hypreal_of_real == star_of"
+
+abbreviation
+  hypreal_of_hypnat :: "hypnat \<Rightarrow> hypreal" where
+  "hypreal_of_hypnat \<equiv> of_hypnat"
+
+definition
+  omega :: hypreal where
+   -- {*an infinite number @{text "= [<1,2,3,...>]"} *}
+  "omega = star_n (\<lambda>n. real (Suc n))"
+
+definition
+  epsilon :: hypreal where
+   -- {*an infinitesimal number @{text "= [<1,1/2,1/3,...>]"} *}
+  "epsilon = star_n (\<lambda>n. inverse (real (Suc n)))"
+
+notation (xsymbols)
+  omega  ("\<omega>") and
+  epsilon  ("\<epsilon>")
+
+notation (HTML output)
+  omega  ("\<omega>") and
+  epsilon  ("\<epsilon>")
+
+
+subsection {* Real vector class instances *}
+
+instantiation star :: (scaleR) scaleR
+begin
+
+definition
+  star_scaleR_def [transfer_unfold, code func del]: "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)
+
+lemma star_of_scaleR [simp]: "star_of (scaleR r x) = scaleR r (star_of x)"
+by transfer (rule refl)
+
+instance star :: (real_vector) real_vector
+proof
+  fix a b :: real
+  show "\<And>x y::'a star. scaleR a (x + y) = scaleR a x + scaleR a y"
+    by transfer (rule scaleR_right_distrib)
+  show "\<And>x::'a star. scaleR (a + b) x = scaleR a x + scaleR b x"
+    by transfer (rule scaleR_left_distrib)
+  show "\<And>x::'a star. scaleR a (scaleR b x) = scaleR (a * b) x"
+    by transfer (rule scaleR_scaleR)
+  show "\<And>x::'a star. scaleR 1 x = x"
+    by transfer (rule scaleR_one)
+qed
+
+instance star :: (real_algebra) real_algebra
+proof
+  fix a :: real
+  show "\<And>x y::'a star. scaleR a x * y = scaleR a (x * y)"
+    by transfer (rule mult_scaleR_left)
+  show "\<And>x y::'a star. x * scaleR a y = scaleR a (x * y)"
+    by transfer (rule mult_scaleR_right)
+qed
+
+instance star :: (real_algebra_1) real_algebra_1 ..
+
+instance star :: (real_div_algebra) real_div_algebra ..
+
+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)
+
+lemma Standard_of_real [simp]: "of_real r \<in> Standard"
+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)
+
+lemma of_real_eq_star_of [simp]: "of_real = star_of"
+proof
+  fix r :: real
+  show "of_real r = star_of r"
+    by transfer simp
+qed
+
+lemma Reals_eq_Standard: "(Reals :: hypreal set) = Standard"
+by (simp add: Reals_def Standard_def)
+
+
+subsection {* Injection from @{typ hypreal} *}
+
+definition
+  of_hypreal :: "hypreal \<Rightarrow> 'a::real_algebra_1 star" where
+  [transfer_unfold, code func del]: "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 of_hypreal_0 [simp]: "of_hypreal 0 = 0"
+by transfer (rule of_real_0)
+
+lemma of_hypreal_1 [simp]: "of_hypreal 1 = 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_minus [simp]: "\<And>x. of_hypreal (- x) = - of_hypreal x"
+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_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_by_zero} 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,division_by_zero} 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_0_iff [simp]:
+  "\<And>x. (of_hypreal x = 0) = (x = 0)"
+by transfer (rule of_real_eq_0_iff)
+
+
+subsection{*Properties of @{term starrel}*}
+
+lemma lemma_starrel_refl [simp]: "x \<in> starrel `` {x}"
+by (simp add: starrel_def)
+
+lemma starrel_in_hypreal [simp]: "starrel``{x}:star"
+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{*@{term hypreal_of_real}: 
+            the Injection from @{typ real} to @{typ hypreal}*}
+
+lemma inj_star_of: "inj star_of"
+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 Rep_star_star_n_iff [simp]:
+  "(X \<in> Rep_star (star_n Y)) = ({n. Y n = X n} \<in> \<U>)"
+by (simp add: star_n_def)
+
+lemma Rep_star_star_n: "X \<in> Rep_star (star_n X)"
+by simp
+
+subsection{* Properties of @{term star_n} *}
+
+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)
+
+lemma star_n_minus:
+   "- star_n X = star_n (%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_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_inverse:
+      "inverse (star_n X) = star_n (%n. inverse(X n))"
+by (simp only: star_inverse_def starfun_star_n)
+
+lemma star_n_le:
+      "star_n X \<le> star_n Y =  
+       ({n. X n \<le> Y n} \<in> FreeUltrafilterNat)"
+by (simp only: star_le_def starP2_star_n)
+
+lemma star_n_less:
+      "star_n X < star_n Y = ({n. X n < Y n} \<in> FreeUltrafilterNat)"
+by (simp only: star_less_def starP2_star_n)
+
+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 (%n. 1)"
+by (simp only: star_one_def star_of_def)
+
+lemma star_n_abs:
+     "abs (star_n X) = star_n (%n. abs (X n))"
+by (simp only: star_abs_def starfun_star_n)
+
+subsection{*Misc Others*}
+
+lemma hypreal_not_refl2: "!!(x::hypreal). x < y ==> x \<noteq> y"
+by (auto)
+
+lemma hypreal_eq_minus_iff: "((x::hypreal) = y) = (x + - y = 0)"
+by auto
+
+lemma hypreal_mult_left_cancel: "(c::hypreal) \<noteq> 0 ==> (c*a=c*b) = (a=b)"
+by auto
+    
+lemma hypreal_mult_right_cancel: "(c::hypreal) \<noteq> 0 ==> (a*c=b*c) = (a=b)"
+by auto
+
+lemma hypreal_omega_gt_zero [simp]: "0 < omega"
+by (simp add: omega_def star_n_zero_num star_n_less)
+
+subsection{*Existence of Infinite Hyperreal Number*}
+
+text{*Existence of infinite number not corresponding to any real number.
+Use assumption that member @{term FreeUltrafilterNat} is not finite.*}
+
+
+text{*A few lemmas first*}
+
+lemma lemma_omega_empty_singleton_disj: "{n::nat. x = real n} = {} |  
+      (\<exists>y. {n::nat. x = real n} = {y})"
+by force
+
+lemma lemma_finite_omega_set: "finite {n::nat. x = real n}"
+by (cut_tac x = x in lemma_omega_empty_singleton_disj, auto)
+
+lemma not_ex_hypreal_of_real_eq_omega: 
+      "~ (\<exists>x. hypreal_of_real x = omega)"
+apply (simp add: omega_def)
+apply (simp add: star_of_def star_n_eq_iff)
+apply (auto simp add: real_of_nat_Suc diff_eq_eq [symmetric] 
+            lemma_finite_omega_set [THEN FreeUltrafilterNat.finite])
+done
+
+lemma hypreal_of_real_not_eq_omega: "hypreal_of_real x \<noteq> omega"
+by (insert not_ex_hypreal_of_real_eq_omega, auto)
+
+text{*Existence of infinitesimal number also not corresponding to any
+ real number*}
+
+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
+
+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 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])
+
+lemma hypreal_of_real_not_eq_epsilon: "hypreal_of_real x \<noteq> epsilon"
+by (insert not_ex_hypreal_of_real_eq_epsilon, auto)
+
+lemma hypreal_epsilon_not_zero: "epsilon \<noteq> 0"
+by (simp add: epsilon_def star_zero_def star_of_def star_n_eq_iff
+         del: star_of_zero)
+
+lemma hypreal_epsilon_inverse_omega: "epsilon = inverse(omega)"
+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{*Absolute Value Function for the Hyperreals*}
+
+lemma hrabs_add_less:
+     "[| abs x < r; abs y < s |] ==> abs(x+y) < r + (s::hypreal)"
+by (simp add: abs_if split: split_if_asm)
+
+lemma hrabs_less_gt_zero: "abs x < r ==> (0::hypreal) < r"
+by (blast intro!: order_le_less_trans abs_ge_zero)
+
+lemma hrabs_disj: "abs x = (x::'a::abs_if) | abs x = -x"
+by (simp add: abs_if)
+
+lemma hrabs_add_lemma_disj: "(y::hypreal) + - x + (y + - z) = abs (x + - z) ==> y = z | x = y"
+by (simp add: abs_if split add: split_if_asm)
+
+
+subsection{*Embedding the Naturals into the Hyperreals*}
+
+abbreviation
+  hypreal_of_nat :: "nat => hypreal" where
+  "hypreal_of_nat == of_nat"
+
+lemma SNat_eq: "Nats = {n. \<exists>N. n = hypreal_of_nat N}"
+by (simp add: Nats_def image_def)
+
+(*------------------------------------------------------------*)
+(* naturals embedded in hyperreals                            *)
+(* is a hyperreal c.f. NS extension                           *)
+(*------------------------------------------------------------*)
+
+lemma hypreal_of_nat_eq:
+     "hypreal_of_nat (n::nat) = hypreal_of_real (real n)"
+by (simp add: real_of_nat_def)
+
+lemma hypreal_of_nat:
+     "hypreal_of_nat m = star_n (%n. real m)"
+apply (fold star_of_def)
+apply (simp add: real_of_nat_def)
+done
+
+(*
+FIXME: we should declare this, as for type int, but many proofs would break.
+It replaces x+-y by x-y.
+Addsimps [symmetric hypreal_diff_def]
+*)
+
+use "hypreal_arith.ML"
+declaration {* K hypreal_arith_setup *}
+
+
+subsection {* Exponentials on the Hyperreals *}
+
+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 hrealpow_two: "(r::hypreal) ^ Suc (Suc 0) = r * r"
+by simp
+
+lemma hrealpow_two_le [simp]: "(0::hypreal) \<le> r ^ Suc (Suc 0)"
+by (auto simp add: zero_le_mult_iff)
+
+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_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 hypreal_add_nonneg_eq_0_iff:
+     "[| 0 \<le> x; 0 \<le> y |] ==> (x+y = 0) = (x = 0 & y = (0::hypreal))"
+by arith
+
+
+text{*FIXME: DELETE THESE*}
+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
+
+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)
+
+(*FIXME: This and RealPow.abs_realpow_two should be replaced by an abstract
+  result proved in Ring_and_Field*)
+lemma hrabs_hrealpow_two [simp]:
+     "abs(x ^ Suc (Suc 0)) = (x::hypreal) ^ Suc (Suc 0)"
+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)
+
+lemma two_hrealpow_gt [simp]: "hypreal_of_nat n < 2 ^ n"
+apply (induct n)
+apply (auto simp add: left_distrib)
+apply (cut_tac n = n in two_hrealpow_ge_one, arith)
+done
+
+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_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: right_distrib left_distrib)
+
+lemma power_hypreal_of_real_number_of:
+     "(number_of v :: hypreal) ^ n = hypreal_of_real ((number_of v) ^ n)"
+by simp
+declare power_hypreal_of_real_number_of [of _ "number_of w", standard, simp]
+(*
+lemma hrealpow_HFinite:
+  fixes x :: "'a::{real_normed_algebra,recpower} star"
+  shows "x \<in> HFinite ==> x ^ n \<in> HFinite"
+apply (induct_tac "n")
+apply (auto simp add: power_Suc intro: HFinite_mult)
+done
+*)
+
+subsection{*Powers with Hypernatural Exponents*}
+
+definition pow :: "['a::power star, nat star] \<Rightarrow> 'a star" (infixr "pow" 80) where
+  hyperpow_def [transfer_unfold, code func del]: "R pow N = ( *f2* op ^) R N"
+  (* hypernatural powers of hyperreals *)
+
+lemma Standard_hyperpow [simp]:
+  "\<lbrakk>r \<in> Standard; n \<in> Standard\<rbrakk> \<Longrightarrow> r pow n \<in> Standard"
+unfolding hyperpow_def by simp
+
+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::{recpower,semiring_0} star) pow (n + (1::hypnat)) = 0"
+by transfer simp
+
+lemma hyperpow_not_zero:
+  "\<And>r n. r \<noteq> (0::'a::{recpower,field} star) ==> r pow n \<noteq> 0"
+by transfer (rule field_power_not_zero)
+
+lemma hyperpow_inverse:
+  "\<And>r n. r \<noteq> (0::'a::{recpower,division_by_zero,field} star)
+   \<Longrightarrow> inverse (r pow n) = (inverse r) pow n"
+by transfer (rule power_inverse)
+
+lemma hyperpow_hrabs:
+  "\<And>r n. abs (r::'a::{recpower,ordered_idom} star) pow n = abs (r pow n)"
+by transfer (rule power_abs [symmetric])
+
+lemma hyperpow_add:
+  "\<And>r n m. (r::'a::recpower star) pow (n + m) = (r pow n) * (r pow m)"
+by transfer (rule power_add)
+
+lemma hyperpow_one [simp]:
+  "\<And>r. (r::'a::recpower star) pow (1::hypnat) = r"
+by transfer (rule power_one_right)
+
+lemma hyperpow_two:
+  "\<And>r. (r::'a::recpower star) pow ((1::hypnat) + (1::hypnat)) = r * r"
+by transfer (simp add: power_Suc)
+
+lemma hyperpow_gt_zero:
+  "\<And>r n. (0::'a::{recpower,ordered_semidom} star) < r \<Longrightarrow> 0 < r pow n"
+by transfer (rule zero_less_power)
+
+lemma hyperpow_ge_zero:
+  "\<And>r n. (0::'a::{recpower,ordered_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::{recpower,ordered_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_eq_one [simp]:
+  "\<And>n. 1 pow n = (1::'a::recpower star)"
+by transfer (rule power_one)
+
+lemma hrabs_hyperpow_minus_one [simp]:
+  "\<And>n. abs(-1 pow n) = (1::'a::{number_ring,recpower,ordered_idom} star)"
+by transfer (rule abs_power_minus_one)
+
+lemma hyperpow_mult:
+  "\<And>r s n. (r * s::'a::{comm_monoid_mult,recpower} star) pow n
+   = (r pow n) * (s pow n)"
+by transfer (rule power_mult_distrib)
+
+lemma hyperpow_two_le [simp]:
+  "(0::'a::{recpower,ordered_ring_strict} star) \<le> r pow (1 + 1)"
+by (auto simp add: hyperpow_two zero_le_mult_iff)
+
+lemma hrabs_hyperpow_two [simp]:
+  "abs(x pow (1 + 1)) =
+   (x::'a::{recpower,ordered_ring_strict} star) pow (1 + 1)"
+by (simp only: abs_of_nonneg hyperpow_two_le)
+
+lemma hyperpow_two_hrabs [simp]:
+  "abs(x::'a::{recpower,ordered_idom} star) pow (1 + 1)  = x pow (1 + 1)"
+by (simp add: hyperpow_hrabs)
+
+text{*The precondition could be weakened to @{term "0\<le>x"}*}
+lemma hypreal_mult_less_mono:
+     "[| u<v;  x<y;  (0::hypreal) < v;  0 < x |] ==> u*x < v* y"
+ by (simp add: Ring_and_Field.mult_strict_mono order_less_imp_le)
+
+lemma hyperpow_two_gt_one:
+  "\<And>r::'a::{recpower,ordered_semidom} star. 1 < r \<Longrightarrow> 1 < r pow (1 + 1)"
+by transfer (simp add: power_gt1)
+
+lemma hyperpow_two_ge_one:
+  "\<And>r::'a::{recpower,ordered_semidom} star. 1 \<le> r \<Longrightarrow> 1 \<le> r pow (1 + 1)"
+by transfer (simp add: 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
+
+lemma hyperpow_minus_one2 [simp]:
+     "!!n. -1 pow ((1 + 1)*n) = (1::hypreal)"
+by transfer (subst power_mult, simp)
+
+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_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)
+
+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> Reals"
+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_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_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_hypnat_of_nat: "\<And>x. x pow hypnat_of_nat n = x ^ n"
+by transfer (rule refl)
+
+lemma of_hypreal_hyperpow:
+  "\<And>x n. of_hypreal (x pow n) =
+   (of_hypreal x::'a::{real_algebra_1,recpower} star) pow n"
+by transfer (rule of_real_power)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/HyperNat.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,425 @@
+(*  Title       : HyperNat.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+
+Converted to Isar and polished by lcp    
+*)
+
+header{*Hypernatural numbers*}
+
+theory HyperNat
+imports StarDef
+begin
+
+types hypnat = "nat star"
+
+abbreviation
+  hypnat_of_nat :: "nat => nat star" where
+  "hypnat_of_nat == star_of"
+
+definition
+  hSuc :: "hypnat => hypnat" where
+  hSuc_def [transfer_unfold, code func del]: "hSuc = *f* Suc"
+
+subsection{*Properties Transferred from Naturals*}
+
+lemma hSuc_not_zero [iff]: "\<And>m. hSuc m \<noteq> 0"
+by transfer (rule Suc_not_Zero)
+
+lemma zero_not_hSuc [iff]: "\<And>m. 0 \<noteq> hSuc m"
+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 zero_less_hSuc [iff]: "\<And>n. 0 < hSuc n"
+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_diff_0_eq_0 [simp]: "!!n. (0::hypnat) - 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_diff_diff_left: "!!i j k. (i::hypnat) - 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_add_inverse [simp]: "!!m n. ((n::hypnat) + 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_cancel [simp]: "!!k m n. ((k::hypnat) + 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_add_0 [simp]: "!!m n. (n::hypnat) - (n+m) = (0::hypnat)"
+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_distrib2: "!!k m n. (k::hypnat) * (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_mult_is_0 [simp]: "!!m n. (m*n = (0::hypnat)) = (m=0 | 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_not_less0 [iff]: "!!n. ~ n < (0::hypnat)"
+by transfer (rule not_less0)
+
+lemma hypnat_less_one [iff]:
+      "!!n. (n < (1::hypnat)) = (n=0)"
+by transfer (rule less_one)
+
+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]: "!!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)
+
+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_le_add1 [simp]: "!!x n. (x::hypnat) \<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_one_self_less [simp]: "(x::hypnat) < x + (1::hypnat)"
+by (insert add_strict_left_mono [OF zero_less_one], auto)
+
+lemma hypnat_neq0_conv [iff]: "!!n. (n \<noteq> 0) = (0 < (n::hypnat))"
+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_iff2: "(0 < n) = (\<exists>m. n = m + (1::hypnat))"
+apply safe
+ apply (rule_tac x = "n - (1::hypnat) " in exI)
+ apply (simp add: hypnat_gt_zero_iff) 
+apply (insert add_le_less_mono [OF _ zero_less_one, of 0], auto) 
+done
+
+lemma hypnat_add_self_not_less: "~ (x + y < (x::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))"
+    -- {* elimination of @{text -} on @{text hypnat} *}
+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])
+next
+  case False
+    thus ?thesis
+      by (auto simp add: linorder_not_less dest: order_le_less_trans) 
+qed
+
+subsection{*Properties of the set of embedded natural numbers*}
+
+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
+qed
+
+lemma Nats_eq_Standard: "(Nats :: nat star set) = Standard"
+by (auto simp add: Nats_def Standard_def)
+
+lemma hypnat_of_nat_mem_Nats [simp]: "hypnat_of_nat n \<in> Nats"
+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_Suc [simp]:
+     "hypnat_of_nat (Suc n) = hypnat_of_nat n + (1::hypnat)"
+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 Nats_diff [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> (a-b :: hypnat) \<in> Nats"
+by (simp add: Nats_eq_Standard)
+
+
+subsection{*Infinite Hypernatural Numbers -- @{term HNatInfinite}*}
+
+definition
+  (* the set of infinite hypernatural numbers *)
+  HNatInfinite :: "hypnat set" where
+  "HNatInfinite = {n. n \<notin> Nats}"
+
+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 star_of_neq_HNatInfinite: "N \<in> HNatInfinite \<Longrightarrow> star_of n \<noteq> N"
+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_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
+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)
+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])
+
+subsubsection {* Closure Rules *}
+
+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_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 zero_less_HNatInfinite: "x \<in> HNatInfinite \<Longrightarrow> 0 < x"
+by (simp add: Nats_less_HNatInfinite)
+
+lemma one_less_HNatInfinite: "x \<in> HNatInfinite \<Longrightarrow> 1 < x"
+by (simp add: Nats_less_HNatInfinite)
+
+lemma one_le_HNatInfinite: "x \<in> HNatInfinite \<Longrightarrow> 1 \<le> x"
+by (simp add: Nats_le_HNatInfinite)
+
+lemma zero_not_mem_HNatInfinite [simp]: "0 \<notin> HNatInfinite"
+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 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_add: "x \<in> HNatInfinite \<Longrightarrow> x + y \<in> HNatInfinite"
+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)
+
+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_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
+
+
+subsection{*Existence of an infinite hypernatural number*}
+
+definition
+  (* omega is in fact an infinite hypernatural number = [<1,2,3,...>] *)
+  whn :: hypnat where
+  hypnat_omega_def: "whn = star_n (%n::nat. n)"
+
+lemma hypnat_of_nat_neq_whn: "hypnat_of_nat n \<noteq> whn"
+by (simp add: 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: 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)
+
+lemma HNatInfinite_whn [simp]: "whn \<in> HNatInfinite"
+by (simp add: HNatInfinite_def)
+
+lemma lemma_unbounded_set [simp]: "{n::nat. m < n} \<in> FreeUltrafilterNat"
+apply (insert finite_atMost [of m]) 
+apply (simp add: atMost_def)
+apply (drule FreeUltrafilterNat.finite)
+apply (drule FreeUltrafilterNat.not_memD)
+apply (simp add: Collect_neg_eq [symmetric] linorder_not_le)
+done
+
+lemma Compl_Collect_le: "- {n::nat. N \<le> n} = {n. n < N}"
+by (simp add: Collect_neg_eq [symmetric] linorder_not_le) 
+
+lemma hypnat_of_nat_eq:
+     "hypnat_of_nat m  = star_n (%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)
+
+lemma Nats_less_whn: "n \<in> Nats \<Longrightarrow> n < whn"
+by (simp add: Nats_less_HNatInfinite)
+
+lemma Nats_le_whn: "n \<in> Nats \<Longrightarrow> n \<le> whn"
+by (simp add: Nats_le_HNatInfinite)
+
+lemma hypnat_of_nat_less_whn [simp]: "hypnat_of_nat n < 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)
+
+lemma hypnat_zero_less_hypnat_omega [simp]: "0 < whn"
+by (simp add: Nats_less_whn)
+
+lemma hypnat_one_less_hypnat_omega [simp]: "1 < whn"
+by (simp add: Nats_less_whn)
+
+
+subsubsection{*Alternative characterization of the set of infinite hypernaturals*}
+
+text{* @{term "HNatInfinite = {N. \<forall>n \<in> Nats. n < N}"}*}
+
+(*??delete? similar reasoning in hypnat_omega_gt_SHNat above*)
+lemma HNatInfinite_FreeUltrafilterNat_lemma:
+  assumes "\<forall>N::nat. {n. f n \<noteq> N} \<in> FreeUltrafilterNat"
+  shows "{n. N < f n} \<in> FreeUltrafilterNat"
+apply (induct N)
+using assms
+apply (drule_tac x = 0 in spec, simp)
+using assms
+apply (drule_tac x = "Suc N" in spec)
+apply (elim ultra, auto)
+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
+
+
+subsubsection{*Alternative Characterization of @{term HNatInfinite} using 
+Free Ultrafilter*}
+
+lemma HNatInfinite_FreeUltrafilterNat:
+     "star_n X \<in> HNatInfinite ==> \<forall>u. {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. {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)
+
+lemma HNatInfinite_FreeUltrafilterNat_iff:
+     "(star_n X \<in> HNatInfinite) = (\<forall>u. {n. u < X n}:  FreeUltrafilterNat)"
+by (rule iffI [OF HNatInfinite_FreeUltrafilterNat 
+                 FreeUltrafilterNat_HNatInfinite])
+
+subsection {* Embedding of the Hypernaturals into other types *}
+
+definition
+  of_hypnat :: "hypnat \<Rightarrow> 'a::semiring_1_cancel star" where
+  of_hypnat_def [transfer_unfold, code func del]: "of_hypnat = *f* of_nat"
+
+lemma of_hypnat_0 [simp]: "of_hypnat 0 = 0"
+by transfer (rule of_nat_0)
+
+lemma of_hypnat_1 [simp]: "of_hypnat 1 = 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)
+
+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_less_iff [simp]:
+  "\<And>m n. (of_hypnat m < (of_hypnat n::'a::ordered_semidom star)) = (m < n)"
+by transfer (rule of_nat_less_iff)
+
+lemma of_hypnat_0_less_iff [simp]:
+  "\<And>n. (0 < (of_hypnat n::'a::ordered_semidom star)) = (0 < n)"
+by transfer (rule of_nat_0_less_iff)
+
+lemma of_hypnat_less_0_iff [simp]:
+  "\<And>m. \<not> (of_hypnat m::'a::ordered_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::ordered_semidom star)) = (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::ordered_semidom star)"
+by transfer (rule of_nat_0_le_iff)
+
+lemma of_hypnat_le_0_iff [simp]:
+  "\<And>m. ((of_hypnat m::'a::ordered_semidom star) \<le> 0) = (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::ordered_semidom star)) = (m = n)"
+by transfer (rule of_nat_eq_iff)
+
+lemma of_hypnat_eq_0_iff [simp]:
+  "\<And>m. ((of_hypnat m::'a::ordered_semidom star) = 0) = (m = 0)"
+by transfer (rule of_nat_eq_0_iff)
+
+lemma HNatInfinite_of_hypnat_gt_zero:
+  "N \<in> HNatInfinite \<Longrightarrow> (0::'a::ordered_semidom star) < of_hypnat N"
+by (rule ccontr, simp add: linorder_not_less)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/Hypercomplex.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,5 @@
+theory Hypercomplex
+imports CLim Hyperreal
+begin
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/Hyperreal.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,14 @@
+(*  Title:      HOL/Hyperreal/Hyperreal.thy
+    ID:         $Id$
+    Author:     Jacques Fleuriot, Cambridge University Computer Laboratory
+    Copyright   1998  University of Cambridge
+
+Construction of the Hyperreals by the Ultrapower Construction
+and mechanization of Nonstandard Real Analysis
+*)
+
+theory Hyperreal
+imports Ln Deriv Taylor Integration HLog
+begin
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/NSA.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,2298 @@
+(*  Title       : NSA.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+
+Converted to Isar and polished by lcp
+*)
+
+header{*Infinite Numbers, Infinitesimals, Infinitely Close Relation*}
+
+theory NSA
+imports HyperDef "../Real/RComplete"
+begin
+
+definition
+  hnorm :: "'a::norm star \<Rightarrow> real star" where
+  [transfer_unfold]: "hnorm = *f* norm"
+
+definition
+  Infinitesimal  :: "('a::real_normed_vector) star set" where
+  [code func del]: "Infinitesimal = {x. \<forall>r \<in> Reals. 0 < r --> hnorm x < r}"
+
+definition
+  HFinite :: "('a::real_normed_vector) star set" where
+  [code func del]: "HFinite = {x. \<exists>r \<in> Reals. hnorm x < r}"
+
+definition
+  HInfinite :: "('a::real_normed_vector) star set" where
+  [code func del]: "HInfinite = {x. \<forall>r \<in> Reals. r < hnorm x}"
+
+definition
+  approx :: "['a::real_normed_vector star, 'a star] => bool"  (infixl "@=" 50) where
+    --{*the `infinitely close' relation*}
+  "(x @= y) = ((x - y) \<in> Infinitesimal)"
+
+definition
+  st        :: "hypreal => hypreal" where
+    --{*the standard part of a hyperreal*}
+  "st = (%x. @r. x \<in> HFinite & r \<in> Reals & r @= x)"
+
+definition
+  monad     :: "'a::real_normed_vector star => 'a star set" where
+  "monad x = {y. x @= y}"
+
+definition
+  galaxy    :: "'a::real_normed_vector star => 'a star set" where
+  "galaxy x = {y. (x + -y) \<in> HFinite}"
+
+notation (xsymbols)
+  approx  (infixl "\<approx>" 50)
+
+notation (HTML output)
+  approx  (infixl "\<approx>" 50)
+
+lemma SReal_def: "Reals == {x. \<exists>r. x = hypreal_of_real r}"
+by (simp add: Reals_def image_def)
+
+subsection {* Nonstandard Extension of the Norm Function *}
+
+definition
+  scaleHR :: "real star \<Rightarrow> 'a star \<Rightarrow> 'a::real_normed_vector star" where
+  [transfer_unfold, code func del]: "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,recpower} star) n.
+   hnorm (x pow n) = hnorm x pow n"
+by transfer (rule norm_power)
+
+lemma hnorm_one [simp]:
+  "hnorm (1\<Colon>'a::real_normed_div_algebra star) = 1"
+by transfer (rule norm_one)
+
+lemma hnorm_zero [simp]:
+  "hnorm (0\<Colon>'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_by_zero} star.
+   hnorm (inverse a) = inverse (hnorm a)"
+by transfer (rule norm_inverse)
+
+lemma hnorm_divide:
+  "\<And>a b::'a::{real_normed_field,division_by_zero} star.
+   hnorm (a / b) = hnorm a / hnorm b"
+by transfer (rule norm_divide)
+
+lemma hypreal_hnorm_def [simp]:
+  "\<And>r::hypreal. hnorm r \<equiv> \<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{*Closure Laws for the Standard Reals*}
+
+lemma Reals_minus_iff [simp]: "(-x \<in> Reals) = (x \<in> Reals)"
+apply auto
+apply (drule Reals_minus, auto)
+done
+
+lemma Reals_add_cancel: "\<lbrakk>x + y \<in> Reals; y \<in> Reals\<rbrakk> \<Longrightarrow> x \<in> Reals"
+by (drule (1) Reals_diff, simp)
+
+lemma SReal_hrabs: "(x::hypreal) \<in> Reals ==> abs x \<in> Reals"
+by (simp add: Reals_eq_Standard)
+
+lemma SReal_hypreal_of_real [simp]: "hypreal_of_real x \<in> Reals"
+by (simp add: Reals_eq_Standard)
+
+lemma SReal_divide_number_of: "r \<in> Reals ==> r/(number_of w::hypreal) \<in> Reals"
+by simp
+
+text{*epsilon is not in Reals because it is an infinitesimal*}
+lemma SReal_epsilon_not_mem: "epsilon \<notin> Reals"
+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> Reals"
+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> Reals} = (UNIV::real set)"
+by simp
+
+lemma SReal_iff: "(x \<in> Reals) = (\<exists>y. x = hypreal_of_real y)"
+by (simp add: SReal_def)
+
+lemma hypreal_of_real_image: "hypreal_of_real `(UNIV::real set) = Reals"
+by (simp add: Reals_eq_Standard Standard_def)
+
+lemma inv_hypreal_of_real_image: "inv hypreal_of_real ` Reals = 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> Reals |] ==> \<exists>Q. P = hypreal_of_real ` Q"
+by (simp add: SReal_def image_def, blast)
+
+lemma SReal_dense:
+     "[| (x::hypreal) \<in> Reals; y \<in> Reals;  x<y |] ==> \<exists>r \<in> Reals. x<r & r<y"
+apply (auto simp add: SReal_def)
+apply (drule dense, auto)
+done
+
+text{*Completeness of Reals, but both lemmas are unused.*}
+
+lemma SReal_sup_lemma:
+     "P \<subseteq> Reals ==> ((\<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> Reals; \<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{* Set of Finite Elements is a Subring of the Extended Reals*}
+
+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: "(Reals::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]: "(abs (x::hypreal) \<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_number_of [simp]: "number_of w \<in> HFinite"
+unfolding star_number_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,recpower} star"
+  shows "x \<in> HFinite ==> x ^ n \<in> HFinite"
+apply (induct_tac "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 (case_tac "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{* Set of Infinitesimals is a Subring of the Hyperreals*}
+
+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_number_of 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]:
+  "(abs (x::hypreal) \<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"
+by (simp add: diff_def Infinitesimal_add)
+
+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 Reals_divide)
+apply assumption
+apply (simp only: divide_pos_pos)
+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 (simp only: divide_pos_pos)
+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 Reals_divide)
+apply (simp only: divide_pos_pos)
+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 add: divide_pos_pos)
+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 (auto intro: HInfinite_add_ge_zero)
+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 ==> abs x \<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; abs (x::hypreal) \<le> e |] ==> x \<in> Infinitesimal"
+by (erule hnorm_le_Infinitesimal, simp)
+
+lemma hrabs_less_Infinitesimal:
+      "[| e \<in> Infinitesimal; abs (x::hypreal) < 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 |] ==> abs (x pow N) \<le> abs x"
+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 (drule mp, blast intro: mult_pos_pos)
+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{*The Infinitely Close Relation*}
+
+lemma mem_infmal_iff: "(x \<in> Infinitesimal) = (x @= 0)"
+by (simp add: Infinitesimal_def approx_def)
+
+lemma approx_minus_iff: " (x @= y) = (x - y @= 0)"
+by (simp add: approx_def)
+
+lemma approx_minus_iff2: " (x @= y) = (-y + x @= 0)"
+by (simp add: approx_def diff_minus add_commute)
+
+lemma approx_refl [iff]: "x @= 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 @= y ==> y @= x"
+apply (simp add: approx_def)
+apply (drule Infinitesimal_minus_iff [THEN iffD2])
+apply simp
+done
+
+lemma approx_trans: "[| x @= y; y @= z |] ==> x @= z"
+apply (simp add: approx_def)
+apply (drule (1) Infinitesimal_add)
+apply (simp add: diff_def)
+done
+
+lemma approx_trans2: "[| r @= x; s @= x |] ==> r @= s"
+by (blast intro: approx_sym approx_trans)
+
+lemma approx_trans3: "[| x @= r; x @= s|] ==> r @= s"
+by (blast intro: approx_sym approx_trans)
+
+lemma number_of_approx_reorient: "(number_of w @= x) = (x @= number_of w)"
+by (blast intro: approx_sym)
+
+lemma zero_approx_reorient: "(0 @= x) = (x @= 0)"
+by (blast intro: approx_sym)
+
+lemma one_approx_reorient: "(1 @= x) = (x @= 1)"
+by (blast intro: approx_sym)
+
+
+ML {*
+local
+(*** re-orientation, following HOL/Integ/Bin.ML
+     We re-orient x @=y where x is 0, 1 or a numeral, unless y is as well!
+ ***)
+
+(*reorientation simprules using ==, for the following simproc*)
+val meta_zero_approx_reorient = thm "zero_approx_reorient" RS eq_reflection;
+val meta_one_approx_reorient = thm "one_approx_reorient" RS eq_reflection;
+val meta_number_of_approx_reorient = thm "number_of_approx_reorient" RS eq_reflection
+
+(*reorientation simplification procedure: reorients (polymorphic)
+  0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*)
+fun reorient_proc sg _ (_ $ t $ u) =
+  case u of
+      Const(@{const_name HOL.zero}, _) => NONE
+    | Const(@{const_name HOL.one}, _) => NONE
+    | Const(@{const_name Int.number_of}, _) $ _ => NONE
+    | _ => SOME (case t of
+                Const(@{const_name HOL.zero}, _) => meta_zero_approx_reorient
+              | Const(@{const_name HOL.one}, _) => meta_one_approx_reorient
+              | Const(@{const_name Int.number_of}, _) $ _ =>
+                                 meta_number_of_approx_reorient);
+
+in
+val approx_reorient_simproc =
+  Int_Numeral_Base_Simprocs.prep_simproc
+    ("reorient_simproc", ["0@=x", "1@=x", "number_of w @= x"], reorient_proc);
+end;
+
+Addsimprocs [approx_reorient_simproc];
+*}
+
+lemma Infinitesimal_approx_minus: "(x-y \<in> Infinitesimal) = (x @= y)"
+by (simp add: approx_minus_iff [symmetric] mem_infmal_iff)
+
+lemma approx_monad_iff: "(x @= 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 @= y"
+apply (simp add: mem_infmal_iff)
+apply (blast intro: approx_trans approx_sym)
+done
+
+lemma approx_add: "[| a @= b; c @= d |] ==> a+c @= 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 @= b ==> -a @= -b"
+apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])
+apply (drule approx_minus_iff [THEN iffD1])
+apply (simp add: add_commute diff_def)
+done
+
+lemma approx_minus2: "-a @= -b ==> a @= b"
+by (auto dest: approx_minus)
+
+lemma approx_minus_cancel [simp]: "(-a @= -b) = (a @= b)"
+by (blast intro: approx_minus approx_minus2)
+
+lemma approx_add_minus: "[| a @= b; c @= d |] ==> a + -c @= b + -d"
+by (blast intro!: approx_add approx_minus)
+
+lemma approx_diff: "[| a @= b; c @= d |] ==> a - c @= b - d"
+by (simp only: diff_minus approx_add approx_minus)
+
+lemma approx_mult1:
+  fixes a b c :: "'a::real_normed_algebra star"
+  shows "[| a @= b; c: HFinite|] ==> a*c @= 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 @= b; c: HFinite|] ==> c*a @= 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 @= v*x; x @= y; v \<in> HFinite|] ==> u @= 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 @= x*v; x @= y; v \<in> HFinite |] ==> u @= 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 @= x*star_of v; x @= y |] ==> u @= y*star_of v"
+by (auto intro: approx_mult_subst2)
+
+lemma approx_eq_imp: "a = b ==> a @= b"
+by (simp add: approx_def)
+
+lemma Infinitesimal_minus_approx: "x \<in> Infinitesimal ==> -x @= 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 @= z)"
+by (simp add: approx_def)
+
+lemma bex_Infinitesimal_iff2: "(\<exists>y \<in> Infinitesimal. x = z + y) = (x @= z)"
+by (force simp add: bex_Infinitesimal_iff [symmetric])
+
+lemma Infinitesimal_add_approx: "[| y \<in> Infinitesimal; x + y = z |] ==> x @= 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 @= 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 @= y + x"
+by (auto dest: Infinitesimal_add_approx_self simp add: add_commute)
+
+lemma Infinitesimal_add_minus_approx_self: "y \<in> Infinitesimal ==> x @= x + -y"
+by (blast intro!: Infinitesimal_add_approx_self Infinitesimal_minus_iff [THEN iffD2])
+
+lemma Infinitesimal_add_cancel: "[| y \<in> Infinitesimal; x+y @= z|] ==> x @= 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 @= z + y|] ==> x @= 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  @= d + c ==> b @= c"
+apply (drule approx_minus_iff [THEN iffD1])
+apply (simp add: approx_minus_iff [symmetric] add_ac)
+done
+
+lemma approx_add_right_cancel: "b + d @= c + d ==> b @= c"
+apply (rule approx_add_left_cancel)
+apply (simp add: add_commute)
+done
+
+lemma approx_add_mono1: "b @= c ==> d + b @= d + c"
+apply (rule approx_minus_iff [THEN iffD2])
+apply (simp add: approx_minus_iff [symmetric] add_ac)
+done
+
+lemma approx_add_mono2: "b @= c ==> b + a @= c + a"
+by (simp add: add_commute approx_add_mono1)
+
+lemma approx_add_left_iff [simp]: "(a + b @= a + c) = (b @= c)"
+by (fast elim: approx_add_left_cancel approx_add_mono1)
+
+lemma approx_add_right_iff [simp]: "(b + a @= c + a) = (b @= c)"
+by (simp add: add_commute)
+
+lemma approx_HFinite: "[| x \<in> HFinite; x @= 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 @= 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 @= b; c @= d; b: HFinite; d: HFinite|] ==> a*c @= 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 @= star_of b; c: HFinite|] ==> scaleHR a c @= 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 @= b ==> c *\<^sub>R a @= c *\<^sub>R b"
+by (simp add: approx_def Infinitesimal_scaleR2
+              scaleR_right_diff_distrib [symmetric])
+
+lemma approx_scaleR_HFinite:
+  "[|a @= star_of b; c @= d; d: HFinite|] ==> scaleHR a c @= 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 @= star_of b; c @= star_of d |]
+      ==> a*c @= 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> Reals; a \<noteq> 0; a*x @= 0 |] ==> x @= 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> Reals; x @= 0 |] ==> x*a @= 0"
+by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult1)
+
+lemma approx_mult_SReal2: "[| (a::hypreal) \<in> Reals; x @= 0 |] ==> a*x @= 0"
+by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult2)
+
+lemma approx_mult_SReal_zero_cancel_iff [simp]:
+     "[|(a::hypreal) \<in> Reals; a \<noteq> 0 |] ==> (a*x @= 0) = (x @= 0)"
+by (blast intro: approx_SReal_mult_cancel_zero approx_mult_SReal2)
+
+lemma approx_SReal_mult_cancel:
+     "[| (a::hypreal) \<in> Reals; a \<noteq> 0; a* w @= a*z |] ==> w @= 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> Reals; a \<noteq> 0|] ==> (a* w @= a*z) = (w @= 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 @= g; g \<le> z |] ==> f @= 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{* Zero is the Only Infinitesimal that is also a Real*}
+
+lemma Infinitesimal_less_SReal:
+     "[| (x::hypreal) \<in> Reals; 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> Reals|] ==> 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> Reals |] ==> y \<notin> Infinitesimal"
+apply (subst Infinitesimal_minus_iff [symmetric])
+apply (rule SReal_not_Infinitesimal, auto)
+done
+
+lemma SReal_Int_Infinitesimal_zero: "Reals 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> Reals; x \<in> Infinitesimal|] ==> x = 0"
+by (cut_tac SReal_Int_Infinitesimal_zero, blast)
+
+lemma SReal_HFinite_diff_Infinitesimal:
+     "[| (x::hypreal) \<in> Reals; 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 number_of_not_Infinitesimal [simp]:
+     "number_of w \<noteq> (0::hypreal) ==> (number_of w :: hypreal) \<notin> Infinitesimal"
+by (fast dest: Reals_number_of [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> Reals; x @= 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 @= 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> Reals |] ==> 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{* Uniqueness: Two Infinitely Close Reals are Equal*}
+
+lemma star_of_approx_iff [simp]: "(star_of x @= 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> Reals; y \<in> Reals|] ==> (x @= 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 number_of_approx_iff [simp]:
+     "(number_of v @= (number_of w :: 'a::{number,real_normed_vector} star)) =
+      (number_of v = (number_of w :: 'a))"
+apply (unfold star_number_def)
+apply (rule star_of_approx_iff)
+done
+
+(*And also for 0 @= #nn and 1 @= #nn, #nn @= 0 and #nn @= 1.*)
+lemma [simp]:
+  "(number_of w @= (0::'a::{number,real_normed_vector} star)) =
+   (number_of w = (0::'a))"
+  "((0::'a::{number,real_normed_vector} star) @= number_of w) =
+   (number_of w = (0::'a))"
+  "(number_of w @= (1::'b::{number,one,real_normed_vector} star)) =
+   (number_of w = (1::'b))"
+  "((1::'b::{number,one,real_normed_vector} star) @= number_of w) =
+   (number_of w = (1::'b))"
+  "~ (0 @= (1::'c::{zero_neq_one,real_normed_vector} star))"
+  "~ (1 @= (0::'c::{zero_neq_one,real_normed_vector} star))"
+apply (unfold star_number_def star_zero_def star_one_def)
+apply (unfold star_of_approx_iff)
+by (auto intro: sym)
+
+lemma star_of_approx_number_of_iff [simp]:
+     "(star_of k @= number_of w) = (k = number_of w)"
+by (subst star_of_approx_iff [symmetric], auto)
+
+lemma star_of_approx_zero_iff [simp]: "(star_of k @= 0) = (k = 0)"
+by (simp_all add: star_of_approx_iff [symmetric])
+
+lemma star_of_approx_one_iff [simp]: "(star_of k @= 1) = (k = 1)"
+by (simp_all add: star_of_approx_iff [symmetric])
+
+lemma approx_unique_real:
+     "[| (r::hypreal) \<in> Reals; s \<in> Reals; r @= x; s @= x|] ==> r = s"
+by (blast intro: SReal_approx_iff [THEN iffD1] approx_trans2)
+
+
+subsection{* Existence of Unique Real Infinitely Close*}
+
+subsubsection{*Lifting of the Ub and Lub Properties*}
+
+lemma hypreal_of_real_isUb_iff:
+      "(isUb (Reals) (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 Reals (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 Reals (hypreal_of_real ` Q) (hypreal_of_real Y)"
+apply (simp add: isLub_def leastP_def)
+apply (auto simp add: hypreal_of_real_isUb_iff setge_def)
+apply (frule_tac x2 = x in isUbD2a [THEN SReal_iff [THEN iffD1], THEN exE])
+ prefer 2 apply assumption
+apply (drule_tac x = xa in spec)
+apply (auto simp add: hypreal_of_real_isUb_iff)
+done
+
+lemma hypreal_of_real_isLub_iff:
+     "(isLub Reals (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 Reals P Y ==> \<exists>Yo. isUb Reals P (hypreal_of_real Yo)"
+by (auto simp add: SReal_iff isUb_def)
+
+lemma lemma_isLub_hypreal_of_real:
+     "isLub Reals P Y ==> \<exists>Yo. isLub Reals 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 Reals P (hypreal_of_real Yo) ==> \<exists>Y. isLub Reals P Y"
+by (auto simp add: isLub_def leastP_def isUb_def)
+
+lemma SReal_complete:
+     "[| P \<subseteq> Reals;  \<exists>x. x \<in> P;  \<exists>Y. isUb Reals P Y |]
+      ==> \<exists>t::hypreal. isLub Reals 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 hypreal_isLub_unique:
+     "[| isLub R S x; isLub R S y |] ==> x = (y::hypreal)"
+apply (frule isLub_isUb)
+apply (frule_tac x = y in isLub_isUb)
+apply (blast intro!: order_antisym dest!: isLub_le_isUb)
+done
+
+lemma lemma_st_part_ub:
+     "(x::hypreal) \<in> HFinite ==> \<exists>u. isUb Reals {s. s \<in> Reals & 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> Reals & 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_subset: "{s. s \<in> Reals & s < x} \<subseteq> Reals"
+by auto
+
+lemma lemma_st_part_lub:
+     "(x::hypreal) \<in> HFinite ==> \<exists>t. isLub Reals {s. s \<in> Reals & s < x} t"
+by (blast intro!: SReal_complete lemma_st_part_ub lemma_st_part_nonempty lemma_st_part_subset)
+
+lemma lemma_hypreal_le_left_cancel: "((t::hypreal) + r \<le> t) = (r \<le> 0)"
+apply safe
+apply (drule_tac c = "-t" in add_left_mono)
+apply (drule_tac [2] c = t in add_left_mono)
+apply (auto simp add: add_assoc [symmetric])
+done
+
+lemma lemma_st_part_le1:
+     "[| (x::hypreal) \<in> HFinite;  isLub Reals {s. s \<in> Reals & s < x} t;
+         r \<in> Reals;  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> Reals |]
+      ==> isUb Reals {s. s \<in> Reals & 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 Reals {s. s \<in> Reals & s < x} t;
+         r \<in> Reals; 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 Reals {s. s \<in> Reals & s < x} t;
+         r \<in> Reals; 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 Reals {s. s \<in> Reals & s < x} t;
+         r \<in> Reals;  0 < r |]
+      ==> -(x + -t) \<le> r"
+apply (subgoal_tac "(t + -r \<le> x)") 
+apply (auto intro: lemma_st_part_le2)
+done
+
+lemma lemma_SReal_ub:
+     "(x::hypreal) \<in> Reals ==> isUb Reals {s. s \<in> Reals & s < x} x"
+by (auto intro: isUbI setleI order_less_imp_le)
+
+lemma lemma_SReal_lub:
+     "(x::hypreal) \<in> Reals ==> isLub Reals {s. s \<in> Reals & 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 Reals {s. s \<in> Reals & s < x} t;
+         r \<in> Reals; 0 < r |]
+      ==> x + -t \<noteq> r"
+apply auto
+apply (frule isLubD1a [THEN Reals_minus])
+apply (drule Reals_add_cancel, assumption)
+apply (drule_tac x = x in lemma_SReal_lub)
+apply (drule hypreal_isLub_unique, assumption, auto)
+done
+
+lemma lemma_st_part_not_eq2:
+     "[| (x::hypreal) \<in> HFinite;
+         isLub Reals {s. s \<in> Reals & s < x} t;
+         r \<in> Reals; 0 < r |]
+      ==> -(x + -t) \<noteq> r"
+apply (auto)
+apply (frule isLubD1a)
+apply (drule Reals_add_cancel, assumption)
+apply (drule_tac a = "-x" in Reals_minus, simp)
+apply (drule_tac x = x in lemma_SReal_lub)
+apply (drule hypreal_isLub_unique, assumption, auto)
+done
+
+lemma lemma_st_part_major:
+     "[| (x::hypreal) \<in> HFinite;
+         isLub Reals {s. s \<in> Reals & s < x} t;
+         r \<in> Reals; 0 < r |]
+      ==> abs (x - t) < r"
+apply (frule lemma_st_part1a)
+apply (frule_tac [4] lemma_st_part2a, auto)
+apply (drule order_le_imp_less_or_eq)+
+apply (auto dest: lemma_st_part_not_eq1 lemma_st_part_not_eq2 simp add: abs_less_iff)
+done
+
+lemma lemma_st_part_major2:
+     "[| (x::hypreal) \<in> HFinite; isLub Reals {s. s \<in> Reals & s < x} t |]
+      ==> \<forall>r \<in> Reals. 0 < r --> abs (x - t) < r"
+by (blast dest!: lemma_st_part_major)
+
+
+text{*Existence of real and Standard Part Theorem*}
+lemma lemma_st_part_Ex:
+     "(x::hypreal) \<in> HFinite
+       ==> \<exists>t \<in> Reals. \<forall>r \<in> Reals. 0 < r --> abs (x - t) < 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 @= t"
+apply (simp add: approx_def Infinitesimal_def)
+apply (drule lemma_st_part_Ex, auto)
+done
+
+text{*There is a unique real infinitely close*}
+lemma st_part_Ex1: "x \<in> HFinite ==> EX! t::hypreal. t \<in> Reals & x @= 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{* Finite, Infinite and Infinitesimal*}
+
+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 right_inverse)
+done
+
+lemma approx_inverse:
+  fixes x y :: "'a::real_normed_div_algebra star"
+  shows
+     "[| x @= y; y \<in>  HFinite - Infinitesimal |]
+      ==> inverse x @= 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) @= 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) @= 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 @= 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 @= a*z |] ==> w @= 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 @= a * z) = (w @= 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> Reals |] ==> 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: "abs (x::hypreal) @= x | abs x @= -x"
+by (cut_tac x = x in hrabs_disj, auto)
+
+
+subsection{*Theorems about Monads*}
+
+lemma monad_hrabs_Un_subset: "monad (abs x) \<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) = (abs x \<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{*Proof that @{term "x @= y"} implies @{term"\<bar>x\<bar> @= \<bar>y\<bar>"}*}
+
+lemma approx_subset_monad: "x @= y ==> {x,y} \<le> monad x"
+apply (simp (no_asm))
+apply (simp add: approx_monad_iff)
+done
+
+lemma approx_subset_monad2: "x @= 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 @= u"
+by (simp add: monad_def)
+
+lemma approx_mem_monad: "x @= u ==> u \<in> monad x"
+by (simp add: monad_def)
+
+lemma approx_mem_monad2: "x @= u ==> x \<in> monad u"
+apply (simp add: monad_def)
+apply (blast intro!: approx_sym)
+done
+
+lemma approx_mem_monad_zero: "[| x @= 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 @= y; (x::hypreal) \<in> Infinitesimal |] ==> abs x @= abs y"
+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 @= 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 @= y|] ==> y < 0"
+by (blast dest: Ball_mem_monad_less_zero approx_subset_monad)
+
+theorem approx_hrabs: "(x::hypreal) @= y ==> abs x @= abs y"
+by (drule approx_hnorm, simp)
+
+lemma approx_hrabs_zero_cancel: "abs(x::hypreal) @= 0 ==> x @= 0"
+apply (cut_tac x = x in hrabs_disj)
+apply (auto dest: approx_minus)
+done
+
+lemma approx_hrabs_add_Infinitesimal:
+  "(e::hypreal) \<in> Infinitesimal ==> abs x @= abs(x+e)"
+by (fast intro: approx_hrabs Infinitesimal_add_approx_self)
+
+lemma approx_hrabs_add_minus_Infinitesimal:
+     "(e::hypreal) \<in> Infinitesimal ==> abs x @= abs(x + -e)"
+by (fast intro: approx_hrabs Infinitesimal_add_minus_approx_self)
+
+lemma hrabs_add_Infinitesimal_cancel:
+     "[| (e::hypreal) \<in> Infinitesimal; e' \<in> Infinitesimal;
+         abs(x+e) = abs(y+e')|] ==> abs x @= abs y"
+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;
+         abs(x + -e) = abs(y + -e')|] ==> abs x @= abs y"
+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 {* More @{term HFinite} and @{term Infinitesimal} Theorems *}
+
+(* 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; abs(hypreal_of_real r) < hypreal_of_real y |]
+      ==> abs (hypreal_of_real r + x) < 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;  abs(hypreal_of_real r) < hypreal_of_real y |]
+      ==> abs (x + hypreal_of_real r) < 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: equals_zero_I [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: add_ac)
+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: add_ac)
+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: add_ac)
+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: add_ac)
+done
+
+lemma monad_hrabs_less:
+     "[| y \<in> monad x; 0 < hypreal_of_real e |]
+      ==> abs (y - x) < 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{* Theorems about Standard Part*}
+
+lemma st_approx_self: "x \<in> HFinite ==> st x @= 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> Reals"
+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> Reals ==> st x = x"
+apply (erule st_unique)
+apply (rule approx_refl)
+done
+
+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 @= y"
+by (auto dest!: st_approx_self elim!: approx_trans3)
+
+lemma approx_st_eq: 
+  assumes "x \<in> HFinite" and "y \<in> HFinite" and "x @= y" 
+  shows "st x = st y"
+proof -
+  have "st x @= x" "st y @= y" "st x \<in> Reals" "st y \<in> Reals"
+    by (simp_all add: st_approx_self st_SReal prems) 
+  with prems 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 @= y) = (st x = st y)"
+by (blast intro: approx_st_eq st_eq_approx)
+
+lemma st_Infinitesimal_add_SReal:
+     "[| x \<in> Reals; 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> Reals; 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_number_of [simp]: "st (number_of w) = number_of w"
+by (rule Reals_number_of [THEN st_SReal_eq])
+
+(*the theorem above for the special cases of zero and one*)
+lemma [simp]: "st 0 = 0" "st 1 = 1"
+by (simp_all 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 hypreal_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)"
+apply (frule HFinite_st_Infinitesimal_add)
+apply (rotate_tac 1)
+apply (frule HFinite_st_Infinitesimal_add, safe)
+apply (rule Infinitesimal_add_st_le_cancel)
+apply (rule_tac [3] x = ea and y = e in Infinitesimal_diff)
+apply (auto simp add: add_assoc [symmetric])
+done
+
+lemma st_zero_le: "[| 0 \<le> x;  x \<in> HFinite |] ==> 0 \<le> st x"
+apply (subst numeral_0_eq_0 [symmetric])
+apply (rule st_number_of [THEN subst])
+apply (rule st_le, auto)
+done
+
+lemma st_zero_ge: "[| x \<le> 0;  x \<in> HFinite |] ==> st x \<le> 0"
+apply (subst numeral_0_eq_0 [symmetric])
+apply (rule st_number_of [THEN subst])
+apply (rule st_le, auto)
+done
+
+lemma st_hrabs: "x \<in> HFinite ==> abs(st x) = st(abs x)"
+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 {* Alternative Definitions using Free Ultrafilter *}
+
+subsubsection {* @{term HFinite} *}
+
+lemma HFinite_FreeUltrafilterNat:
+    "star_n X \<in> HFinite 
+     ==> \<exists>u. {n. norm (X n) < u} \<in> 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. {n. norm (X n) < u} \<in> 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. {n. norm (X n) < u} \<in> FreeUltrafilterNat)"
+by (blast intro!: HFinite_FreeUltrafilterNat FreeUltrafilterNat_HFinite)
+
+subsubsection {* @{term HInfinite} *}
+
+lemma lemma_Compl_eq: "- {n. u < norm (xa n)} = {n. norm (xa n) \<le> u}"
+by auto
+
+lemma lemma_Compl_eq2: "- {n. norm (xa n) < u} = {n. u \<le> norm (xa n)}"
+by auto
+
+lemma lemma_Int_eq1:
+     "{n. norm (xa n) \<le> u} Int {n. u \<le> norm (xa n)}
+          = {n. norm(xa n) = u}"
+by auto
+
+lemma lemma_FreeUltrafilterNat_one:
+     "{n. norm (xa n) = u} \<le> {n. norm (xa n) < u + (1::real)}"
+by auto
+
+(*-------------------------------------
+  Exclude this type of sets from free
+  ultrafilter for Infinite numbers!
+ -------------------------------------*)
+lemma FreeUltrafilterNat_const_Finite:
+     "{n. norm (X n) = u} \<in> FreeUltrafilterNat ==> star_n X \<in> HFinite"
+apply (rule FreeUltrafilterNat_HFinite)
+apply (rule_tac x = "u + 1" in exI)
+apply (erule ultra, simp)
+done
+
+lemma HInfinite_FreeUltrafilterNat:
+     "star_n X \<in> HInfinite ==> \<forall>u. {n. u < norm (X n)} \<in> 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 (drule FreeUltrafilterNat.not_memD)
+apply (simp add: Collect_neg_eq [symmetric] linorder_not_less)
+apply (erule ultra, simp)
+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. {n. u < norm (X n)} \<in> 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)
+apply (drule (1) FreeUltrafilterNat.Int)
+apply (simp add: Collect_conj_eq [symmetric])
+apply (subgoal_tac "\<forall>n. \<not> (norm (X n) < u \<and> u < norm (X n))", auto)
+done
+
+lemma HInfinite_FreeUltrafilterNat_iff:
+     "(star_n X \<in> HInfinite) = (\<forall>u. {n. u < norm (X n)} \<in> FreeUltrafilterNat)"
+by (blast intro!: HInfinite_FreeUltrafilterNat FreeUltrafilterNat_HInfinite)
+
+subsubsection {* @{term Infinitesimal} *}
+
+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. {n. norm (X n) < u} \<in> \<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. {n. norm (X n) < u} \<in> \<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. {n. norm (X n) < u} \<in> \<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 add: real_of_nat_Suc_gt_zero)
+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 (no_asm_use))
+apply (rule real_of_nat_Suc_gt_zero [THEN positive_imp_inverse_positive, THEN star_of_less [THEN iffD2], THEN [2] impE])
+prefer 2 apply assumption
+apply (simp add: real_of_nat_def)
+apply (auto dest!: reals_Archimedean simp add: SReal_iff)
+apply (drule star_of_less [THEN iffD2])
+apply (simp add: real_of_nat_def)
+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{*Proof that @{term omega} is an infinite number*}
+
+text{*It will follow that epsilon is an infinitesimal number.*}
+
+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_nat_segment: "finite {n::nat. n < m}"
+apply (induct "m")
+apply (auto simp add: Suc_Un_eq)
+done
+
+lemma finite_real_of_nat_segment: "finite {n::nat. real n < real (m::nat)}"
+by (auto intro: finite_nat_segment)
+
+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. abs(real n) \<le> u}"
+apply (simp (no_asm) add: real_of_nat_Suc_gt_zero finite_real_of_nat_le_real)
+done
+
+lemma rabs_real_of_nat_le_real_FreeUltrafilterNat:
+     "{n. abs(real n) \<le> u} \<notin> FreeUltrafilterNat"
+by (blast intro!: FreeUltrafilterNat.finite finite_rabs_real_of_nat_le_real)
+
+lemma FreeUltrafilterNat_nat_gt_real: "{n. u < real n} \<in> FreeUltrafilterNat"
+apply (rule ccontr, drule FreeUltrafilterNat.not_memD)
+apply (subgoal_tac "- {n::nat. u < real n} = {n. real n \<le> u}")
+prefer 2 apply force
+apply (simp add: finite_real_of_nat_le_real [THEN FreeUltrafilterNat.finite])
+done
+
+(*--------------------------------------------------------------
+ The complement of {n. abs(real n) \<le> u} =
+ {n. u < abs (real n)} 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{*@{term omega} is a member of @{term HInfinite}*}
+
+lemma FreeUltrafilterNat_omega: "{n. u < real n} \<in> FreeUltrafilterNat"
+apply (cut_tac u = u in rabs_real_of_nat_le_real_FreeUltrafilterNat)
+apply (auto dest: FreeUltrafilterNat.not_memD simp add: Compl_real_le_eq)
+done
+
+theorem HInfinite_omega [simp]: "omega \<in> HInfinite"
+apply (simp add: omega_def)
+apply (rule FreeUltrafilterNat_HInfinite)
+apply (simp (no_asm) add: real_norm_def real_of_nat_Suc diff_less_eq [symmetric] FreeUltrafilterNat_omega)
+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 @= 0"
+apply (simp (no_asm) add: mem_infmal_iff [symmetric])
+done
+
+(*------------------------------------------------------------------------
+  Needed for proof that we define a hyperreal [<X(n)] @= 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 add: real_of_nat_Suc_gt_zero)
+apply (simp add: real_mult_commute)
+done
+
+lemma finite_inverse_real_of_posnat_gt_real:
+     "0 < u ==> finite {n. u < inverse(real(Suc n))}"
+apply (simp (no_asm_simp) add: real_of_nat_less_inverse_iff)
+apply (simp (no_asm_simp) add: real_of_nat_Suc less_diff_eq [symmetric])
+apply (rule finite_real_of_nat_less_real)
+done
+
+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))}"
+apply (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le)
+done
+
+lemma real_of_nat_inverse_eq_iff:
+     "(u = inverse (real(Suc n))) = (real(Suc n) = inverse u)"
+by (auto simp add: real_of_nat_Suc_gt_zero less_imp_neq [THEN not_sym])
+
+lemma lemma_finite_omega_set2: "finite {n::nat. u = inverse(real(Suc n))}"
+apply (simp (no_asm_simp) add: real_of_nat_inverse_eq_iff)
+apply (cut_tac x = "inverse u - 1" in lemma_finite_omega_set)
+apply (simp add: real_of_nat_Suc diff_eq_eq [symmetric] eq_commute)
+done
+
+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_omega_set2 finite_inverse_real_of_posnat_gt_real)
+
+lemma inverse_real_of_posnat_ge_real_FreeUltrafilterNat:
+     "0 < u ==> {n. u \<le> inverse(real(Suc n))} \<notin> 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}"
+apply (auto dest!: order_le_less_trans simp add: linorder_not_le)
+done
+
+lemma FreeUltrafilterNat_inverse_real_of_posnat:
+     "0 < u ==>
+      {n. inverse(real(Suc n)) < u} \<in> FreeUltrafilterNat"
+apply (cut_tac u = u in inverse_real_of_posnat_ge_real_FreeUltrafilterNat)
+apply (auto dest: FreeUltrafilterNat.not_memD simp add: Compl_le_inverse_eq)
+done
+
+text{* 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*}
+
+lemma SEQ_Infinitesimal:
+      "( *f* (%n::nat. inverse(real(Suc n)))) whn : Infinitesimal"
+apply (simp add: hypnat_omega_def starfun_star_n star_n_inverse)
+apply (simp add: Infinitesimal_FreeUltrafilterNat_iff)
+apply (simp add: real_of_nat_Suc_gt_zero FreeUltrafilterNat_inverse_real_of_posnat)
+done
+
+text{* 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.*}
+
+(*-----------------------------------------------------
+    |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"
+apply (auto intro!: bexI dest: FreeUltrafilterNat_inverse_real_of_posnat FreeUltrafilterNat.Int intro: order_less_trans FreeUltrafilterNat.subset simp add: star_n_diff star_of_def Infinitesimal_FreeUltrafilterNat_iff star_n_inverse)
+done
+
+lemma real_seq_to_hypreal_approx:
+     "\<forall>n. norm(X n - x) < inverse(real(Suc n))
+      ==> star_n X @= star_of x"
+apply (subst approx_minus_iff)
+apply (rule mem_infmal_iff [THEN subst])
+apply (erule real_seq_to_hypreal_Infinitesimal)
+done
+
+lemma real_seq_to_hypreal_approx2:
+     "\<forall>n. norm(x - X n) < inverse(real(Suc n))
+               ==> star_n X @= star_of x"
+apply (rule real_seq_to_hypreal_approx)
+apply (subst norm_minus_cancel [symmetric])
+apply (simp del: norm_minus_cancel)
+done
+
+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"
+by (auto intro!: bexI
+	 dest: FreeUltrafilterNat_inverse_real_of_posnat 
+	       FreeUltrafilterNat.Int
+	 intro: order_less_trans FreeUltrafilterNat.subset 
+	 simp add: Infinitesimal_FreeUltrafilterNat_iff star_n_diff 
+                   star_n_inverse)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/NSCA.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,525 @@
+(*  Title       : NSCA.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 2001,2002 University of Edinburgh
+*)
+
+header{*Non-Standard Complex Analysis*}
+
+theory NSCA
+imports NSComplex "../Hyperreal/HTranscendental"
+begin
+
+abbreviation
+   (* standard complex numbers reagarded as an embedded subset of NS complex *)
+   SComplex  :: "hcomplex set" where
+   "SComplex \<equiv> Standard"
+
+definition --{* standard part map*}
+  stc :: "hcomplex => hcomplex" where 
+  [code func del]: "stc x = (SOME r. x \<in> HFinite & r:SComplex & r @= x)"
+
+
+subsection{*Closure Laws for SComplex, the Standard Complex Numbers*}
+
+lemma SComplex_minus_iff [simp]: "(-x \<in> SComplex) = (x \<in> SComplex)"
+by (auto, drule Standard_minus, auto)
+
+lemma SComplex_add_cancel:
+     "[| x + y \<in> SComplex; y \<in> SComplex |] ==> x \<in> SComplex"
+by (drule (1) Standard_diff, simp)
+
+lemma SReal_hcmod_hcomplex_of_complex [simp]:
+     "hcmod (hcomplex_of_complex r) \<in> Reals"
+by (simp add: Reals_eq_Standard)
+
+lemma SReal_hcmod_number_of [simp]: "hcmod (number_of w ::hcomplex) \<in> Reals"
+by (simp add: Reals_eq_Standard)
+
+lemma SReal_hcmod_SComplex: "x \<in> SComplex ==> hcmod x \<in> Reals"
+by (simp add: Reals_eq_Standard)
+
+lemma SComplex_divide_number_of:
+     "r \<in> SComplex ==> r/(number_of w::hcomplex) \<in> SComplex"
+by simp
+
+lemma SComplex_UNIV_complex:
+     "{x. hcomplex_of_complex x \<in> SComplex} = (UNIV::complex set)"
+by simp
+
+lemma SComplex_iff: "(x \<in> SComplex) = (\<exists>y. x = hcomplex_of_complex y)"
+by (simp add: Standard_def image_def)
+
+lemma hcomplex_of_complex_image:
+     "hcomplex_of_complex `(UNIV::complex set) = SComplex"
+by (simp add: Standard_def)
+
+lemma inv_hcomplex_of_complex_image: "inv hcomplex_of_complex `SComplex = UNIV"
+apply (auto simp add: Standard_def image_def)
+apply (rule inj_hcomplex_of_complex [THEN inv_f_f, THEN subst], blast)
+done
+
+lemma SComplex_hcomplex_of_complex_image: 
+      "[| \<exists>x. x: P; P \<le> SComplex |] ==> \<exists>Q. P = hcomplex_of_complex ` Q"
+apply (simp add: Standard_def, blast)
+done
+
+lemma SComplex_SReal_dense:
+     "[| x \<in> SComplex; y \<in> SComplex; hcmod x < hcmod y  
+      |] ==> \<exists>r \<in> Reals. hcmod x< r & r < hcmod y"
+apply (auto intro: SReal_dense simp add: SReal_hcmod_SComplex)
+done
+
+lemma SComplex_hcmod_SReal: 
+      "z \<in> SComplex ==> hcmod z \<in> Reals"
+by (simp add: Reals_eq_Standard)
+
+
+subsection{*The Finite Elements form a Subring*}
+
+lemma HFinite_hcmod_hcomplex_of_complex [simp]:
+     "hcmod (hcomplex_of_complex r) \<in> HFinite"
+by (auto intro!: SReal_subset_HFinite [THEN subsetD])
+
+lemma HFinite_hcmod_iff: "(x \<in> HFinite) = (hcmod x \<in> HFinite)"
+by (simp add: HFinite_def)
+
+lemma HFinite_bounded_hcmod:
+  "[|x \<in> HFinite; y \<le> hcmod x; 0 \<le> y |] ==> y: HFinite"
+by (auto intro: HFinite_bounded simp add: HFinite_hcmod_iff)
+
+
+subsection{*The Complex Infinitesimals form a Subring*}
+
+lemma hcomplex_sum_of_halves: "x/(2::hcomplex) + x/(2::hcomplex) = x"
+by auto
+
+lemma Infinitesimal_hcmod_iff: 
+   "(z \<in> Infinitesimal) = (hcmod z \<in> Infinitesimal)"
+by (simp add: Infinitesimal_def)
+
+lemma HInfinite_hcmod_iff: "(z \<in> HInfinite) = (hcmod z \<in> HInfinite)"
+by (simp add: HInfinite_def)
+
+lemma HFinite_diff_Infinitesimal_hcmod:
+     "x \<in> HFinite - Infinitesimal ==> hcmod x \<in> HFinite - Infinitesimal"
+by (simp add: HFinite_hcmod_iff Infinitesimal_hcmod_iff)
+
+lemma hcmod_less_Infinitesimal:
+     "[| e \<in> Infinitesimal; hcmod x < hcmod e |] ==> x \<in> Infinitesimal"
+by (auto elim: hrabs_less_Infinitesimal simp add: Infinitesimal_hcmod_iff)
+
+lemma hcmod_le_Infinitesimal:
+     "[| e \<in> Infinitesimal; hcmod x \<le> hcmod e |] ==> x \<in> Infinitesimal"
+by (auto elim: hrabs_le_Infinitesimal simp add: Infinitesimal_hcmod_iff)
+
+lemma Infinitesimal_interval_hcmod:
+     "[| e \<in> Infinitesimal;  
+          e' \<in> Infinitesimal;  
+          hcmod e' < hcmod x ; hcmod x < hcmod e  
+       |] ==> x \<in> Infinitesimal"
+by (auto intro: Infinitesimal_interval simp add: Infinitesimal_hcmod_iff)
+
+lemma Infinitesimal_interval2_hcmod:
+     "[| e \<in> Infinitesimal;  
+         e' \<in> Infinitesimal;  
+         hcmod e' \<le> hcmod x ; hcmod x \<le> hcmod e  
+      |] ==> x \<in> Infinitesimal"
+by (auto intro: Infinitesimal_interval2 simp add: Infinitesimal_hcmod_iff)
+
+
+subsection{*The ``Infinitely Close'' Relation*}
+
+(*
+Goalw [capprox_def,approx_def] "(z @c= w) = (hcmod z @= hcmod w)"
+by (auto_tac (claset(),simpset() addsimps [Infinitesimal_hcmod_iff]));
+*)
+
+lemma approx_SComplex_mult_cancel_zero:
+     "[| a \<in> SComplex; a \<noteq> 0; a*x @= 0 |] ==> x @= 0"
+apply (drule Standard_inverse [THEN Standard_subset_HFinite [THEN subsetD]])
+apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric])
+done
+
+lemma approx_mult_SComplex1: "[| a \<in> SComplex; x @= 0 |] ==> x*a @= 0"
+by (auto dest: Standard_subset_HFinite [THEN subsetD] approx_mult1)
+
+lemma approx_mult_SComplex2: "[| a \<in> SComplex; x @= 0 |] ==> a*x @= 0"
+by (auto dest: Standard_subset_HFinite [THEN subsetD] approx_mult2)
+
+lemma approx_mult_SComplex_zero_cancel_iff [simp]:
+     "[|a \<in> SComplex; a \<noteq> 0 |] ==> (a*x @= 0) = (x @= 0)"
+by (blast intro: approx_SComplex_mult_cancel_zero approx_mult_SComplex2)
+
+lemma approx_SComplex_mult_cancel:
+     "[| a \<in> SComplex; a \<noteq> 0; a* w @= a*z |] ==> w @= z"
+apply (drule Standard_inverse [THEN Standard_subset_HFinite [THEN subsetD]])
+apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric])
+done
+
+lemma approx_SComplex_mult_cancel_iff1 [simp]:
+     "[| a \<in> SComplex; a \<noteq> 0|] ==> (a* w @= a*z) = (w @= z)"
+by (auto intro!: approx_mult2 Standard_subset_HFinite [THEN subsetD]
+            intro: approx_SComplex_mult_cancel)
+
+(* TODO: generalize following theorems: hcmod -> hnorm *)
+
+lemma approx_hcmod_approx_zero: "(x @= y) = (hcmod (y - x) @= 0)"
+apply (subst hnorm_minus_commute)
+apply (simp add: approx_def Infinitesimal_hcmod_iff diff_minus)
+done
+
+lemma approx_approx_zero_iff: "(x @= 0) = (hcmod x @= 0)"
+by (simp add: approx_hcmod_approx_zero)
+
+lemma approx_minus_zero_cancel_iff [simp]: "(-x @= 0) = (x @= 0)"
+by (simp add: approx_def)
+
+lemma Infinitesimal_hcmod_add_diff:
+     "u @= 0 ==> hcmod(x + u) - hcmod x \<in> Infinitesimal"
+apply (drule approx_approx_zero_iff [THEN iffD1])
+apply (rule_tac e = "hcmod u" and e' = "- hcmod u" in Infinitesimal_interval2)
+apply (auto simp add: mem_infmal_iff [symmetric] diff_def)
+apply (rule_tac c1 = "hcmod x" in add_le_cancel_left [THEN iffD1])
+apply (auto simp add: diff_minus [symmetric])
+done
+
+lemma approx_hcmod_add_hcmod: "u @= 0 ==> hcmod(x + u) @= hcmod x"
+apply (rule approx_minus_iff [THEN iffD2])
+apply (auto intro: Infinitesimal_hcmod_add_diff simp add: mem_infmal_iff [symmetric] diff_minus [symmetric])
+done
+
+
+subsection{*Zero is the Only Infinitesimal Complex Number*}
+
+lemma Infinitesimal_less_SComplex:
+   "[| x \<in> SComplex; y \<in> Infinitesimal; 0 < hcmod x |] ==> hcmod y < hcmod x"
+by (auto intro: Infinitesimal_less_SReal SComplex_hcmod_SReal simp add: Infinitesimal_hcmod_iff)
+
+lemma SComplex_Int_Infinitesimal_zero: "SComplex Int Infinitesimal = {0}"
+by (auto simp add: Standard_def Infinitesimal_hcmod_iff)
+
+lemma SComplex_Infinitesimal_zero:
+     "[| x \<in> SComplex; x \<in> Infinitesimal|] ==> x = 0"
+by (cut_tac SComplex_Int_Infinitesimal_zero, blast)
+
+lemma SComplex_HFinite_diff_Infinitesimal:
+     "[| x \<in> SComplex; x \<noteq> 0 |] ==> x \<in> HFinite - Infinitesimal"
+by (auto dest: SComplex_Infinitesimal_zero Standard_subset_HFinite [THEN subsetD])
+
+lemma hcomplex_of_complex_HFinite_diff_Infinitesimal:
+     "hcomplex_of_complex x \<noteq> 0 
+      ==> hcomplex_of_complex x \<in> HFinite - Infinitesimal"
+by (rule SComplex_HFinite_diff_Infinitesimal, auto)
+
+lemma number_of_not_Infinitesimal [simp]:
+     "number_of w \<noteq> (0::hcomplex) ==> (number_of w::hcomplex) \<notin> Infinitesimal"
+by (fast dest: Standard_number_of [THEN SComplex_Infinitesimal_zero])
+
+lemma approx_SComplex_not_zero:
+     "[| y \<in> SComplex; x @= y; y\<noteq> 0 |] ==> x \<noteq> 0"
+by (auto dest: SComplex_Infinitesimal_zero approx_sym [THEN mem_infmal_iff [THEN iffD2]])
+
+lemma SComplex_approx_iff:
+     "[|x \<in> SComplex; y \<in> SComplex|] ==> (x @= y) = (x = y)"
+by (auto simp add: Standard_def)
+
+lemma number_of_Infinitesimal_iff [simp]:
+     "((number_of w :: hcomplex) \<in> Infinitesimal) =
+      (number_of w = (0::hcomplex))"
+apply (rule iffI)
+apply (fast dest: Standard_number_of [THEN SComplex_Infinitesimal_zero])
+apply (simp (no_asm_simp))
+done
+
+lemma approx_unique_complex:
+     "[| r \<in> SComplex; s \<in> SComplex; r @= x; s @= x|] ==> r = s"
+by (blast intro: SComplex_approx_iff [THEN iffD1] approx_trans2)
+
+subsection {* Properties of @{term hRe}, @{term hIm} and @{term HComplex} *}
+
+
+lemma abs_hRe_le_hcmod: "\<And>x. \<bar>hRe x\<bar> \<le> hcmod x"
+by transfer (rule abs_Re_le_cmod)
+
+lemma abs_hIm_le_hcmod: "\<And>x. \<bar>hIm x\<bar> \<le> hcmod x"
+by transfer (rule abs_Im_le_cmod)
+
+lemma Infinitesimal_hRe: "x \<in> Infinitesimal \<Longrightarrow> hRe x \<in> Infinitesimal"
+apply (rule InfinitesimalI2, simp)
+apply (rule order_le_less_trans [OF abs_hRe_le_hcmod])
+apply (erule (1) InfinitesimalD2)
+done
+
+lemma Infinitesimal_hIm: "x \<in> Infinitesimal \<Longrightarrow> hIm x \<in> Infinitesimal"
+apply (rule InfinitesimalI2, simp)
+apply (rule order_le_less_trans [OF abs_hIm_le_hcmod])
+apply (erule (1) InfinitesimalD2)
+done
+
+lemma real_sqrt_lessI: "\<lbrakk>0 < u; x < u\<twosuperior>\<rbrakk> \<Longrightarrow> sqrt x < u"
+(* TODO: this belongs somewhere else *)
+by (frule real_sqrt_less_mono) simp
+
+lemma hypreal_sqrt_lessI:
+  "\<And>x u. \<lbrakk>0 < u; x < u\<twosuperior>\<rbrakk> \<Longrightarrow> ( *f* sqrt) x < u"
+by transfer (rule real_sqrt_lessI)
+ 
+lemma hypreal_sqrt_ge_zero: "\<And>x. 0 \<le> x \<Longrightarrow> 0 \<le> ( *f* sqrt) x"
+by transfer (rule real_sqrt_ge_zero)
+
+lemma Infinitesimal_sqrt:
+  "\<lbrakk>x \<in> Infinitesimal; 0 \<le> x\<rbrakk> \<Longrightarrow> ( *f* sqrt) x \<in> Infinitesimal"
+apply (rule InfinitesimalI2)
+apply (drule_tac r="r\<twosuperior>" in InfinitesimalD2, simp)
+apply (simp add: hypreal_sqrt_ge_zero)
+apply (rule hypreal_sqrt_lessI, simp_all)
+done
+
+lemma Infinitesimal_HComplex:
+  "\<lbrakk>x \<in> Infinitesimal; y \<in> Infinitesimal\<rbrakk> \<Longrightarrow> HComplex x y \<in> Infinitesimal"
+apply (rule Infinitesimal_hcmod_iff [THEN iffD2])
+apply (simp add: hcmod_i)
+apply (rule Infinitesimal_sqrt)
+apply (rule Infinitesimal_add)
+apply (erule Infinitesimal_hrealpow, simp)
+apply (erule Infinitesimal_hrealpow, simp)
+apply (rule add_nonneg_nonneg)
+apply (rule zero_le_power2)
+apply (rule zero_le_power2)
+done
+
+lemma hcomplex_Infinitesimal_iff:
+  "(x \<in> Infinitesimal) = (hRe x \<in> Infinitesimal \<and> hIm x \<in> Infinitesimal)"
+apply (safe intro!: Infinitesimal_hRe Infinitesimal_hIm)
+apply (drule (1) Infinitesimal_HComplex, simp)
+done
+
+lemma hRe_diff [simp]: "\<And>x y. hRe (x - y) = hRe x - hRe y"
+by transfer (rule complex_Re_diff)
+
+lemma hIm_diff [simp]: "\<And>x y. hIm (x - y) = hIm x - hIm y"
+by transfer (rule complex_Im_diff)
+
+lemma approx_hRe: "x \<approx> y \<Longrightarrow> hRe x \<approx> hRe y"
+unfolding approx_def by (drule Infinitesimal_hRe) simp
+
+lemma approx_hIm: "x \<approx> y \<Longrightarrow> hIm x \<approx> hIm y"
+unfolding approx_def by (drule Infinitesimal_hIm) simp
+
+lemma approx_HComplex:
+  "\<lbrakk>a \<approx> b; c \<approx> d\<rbrakk> \<Longrightarrow> HComplex a c \<approx> HComplex b d"
+unfolding approx_def by (simp add: Infinitesimal_HComplex)
+
+lemma hcomplex_approx_iff:
+  "(x \<approx> y) = (hRe x \<approx> hRe y \<and> hIm x \<approx> hIm y)"
+unfolding approx_def by (simp add: hcomplex_Infinitesimal_iff)
+
+lemma HFinite_hRe: "x \<in> HFinite \<Longrightarrow> hRe x \<in> HFinite"
+apply (auto simp add: HFinite_def SReal_def)
+apply (rule_tac x="star_of r" in exI, simp)
+apply (erule order_le_less_trans [OF abs_hRe_le_hcmod])
+done
+
+lemma HFinite_hIm: "x \<in> HFinite \<Longrightarrow> hIm x \<in> HFinite"
+apply (auto simp add: HFinite_def SReal_def)
+apply (rule_tac x="star_of r" in exI, simp)
+apply (erule order_le_less_trans [OF abs_hIm_le_hcmod])
+done
+
+lemma HFinite_HComplex:
+  "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> HComplex x y \<in> HFinite"
+apply (subgoal_tac "HComplex x 0 + HComplex 0 y \<in> HFinite", simp)
+apply (rule HFinite_add)
+apply (simp add: HFinite_hcmod_iff hcmod_i)
+apply (simp add: HFinite_hcmod_iff hcmod_i)
+done
+
+lemma hcomplex_HFinite_iff:
+  "(x \<in> HFinite) = (hRe x \<in> HFinite \<and> hIm x \<in> HFinite)"
+apply (safe intro!: HFinite_hRe HFinite_hIm)
+apply (drule (1) HFinite_HComplex, simp)
+done
+
+lemma hcomplex_HInfinite_iff:
+  "(x \<in> HInfinite) = (hRe x \<in> HInfinite \<or> hIm x \<in> HInfinite)"
+by (simp add: HInfinite_HFinite_iff hcomplex_HFinite_iff)
+
+lemma hcomplex_of_hypreal_approx_iff [simp]:
+     "(hcomplex_of_hypreal x @= hcomplex_of_hypreal z) = (x @= z)"
+by (simp add: hcomplex_approx_iff)
+
+lemma Standard_HComplex:
+  "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> HComplex x y \<in> Standard"
+by (simp add: HComplex_def)
+
+(* Here we go - easy proof now!! *)
+lemma stc_part_Ex: "x:HFinite ==> \<exists>t \<in> SComplex. x @= t"
+apply (simp add: hcomplex_HFinite_iff hcomplex_approx_iff)
+apply (rule_tac x="HComplex (st (hRe x)) (st (hIm x))" in bexI)
+apply (simp add: st_approx_self [THEN approx_sym])
+apply (simp add: Standard_HComplex st_SReal [unfolded Reals_eq_Standard])
+done
+
+lemma stc_part_Ex1: "x:HFinite ==> EX! t. t \<in> SComplex &  x @= t"
+apply (drule stc_part_Ex, safe)
+apply (drule_tac [2] approx_sym, drule_tac [2] approx_sym, drule_tac [2] approx_sym)
+apply (auto intro!: approx_unique_complex)
+done
+
+lemmas hcomplex_of_complex_approx_inverse =
+  hcomplex_of_complex_HFinite_diff_Infinitesimal [THEN [2] approx_inverse]
+
+
+subsection{*Theorems About Monads*}
+
+lemma monad_zero_hcmod_iff: "(x \<in> monad 0) = (hcmod x:monad 0)"
+by (simp add: Infinitesimal_monad_zero_iff [symmetric] Infinitesimal_hcmod_iff)
+
+
+subsection{*Theorems About Standard Part*}
+
+lemma stc_approx_self: "x \<in> HFinite ==> stc x @= x"
+apply (simp add: stc_def)
+apply (frule stc_part_Ex, safe)
+apply (rule someI2)
+apply (auto intro: approx_sym)
+done
+
+lemma stc_SComplex: "x \<in> HFinite ==> stc x \<in> SComplex"
+apply (simp add: stc_def)
+apply (frule stc_part_Ex, safe)
+apply (rule someI2)
+apply (auto intro: approx_sym)
+done
+
+lemma stc_HFinite: "x \<in> HFinite ==> stc x \<in> HFinite"
+by (erule stc_SComplex [THEN Standard_subset_HFinite [THEN subsetD]])
+
+lemma stc_unique: "\<lbrakk>y \<in> SComplex; y \<approx> x\<rbrakk> \<Longrightarrow> stc x = y"
+apply (frule Standard_subset_HFinite [THEN subsetD])
+apply (drule (1) approx_HFinite)
+apply (unfold stc_def)
+apply (rule some_equality)
+apply (auto intro: approx_unique_complex)
+done
+
+lemma stc_SComplex_eq [simp]: "x \<in> SComplex ==> stc x = x"
+apply (erule stc_unique)
+apply (rule approx_refl)
+done
+
+lemma stc_hcomplex_of_complex:
+     "stc (hcomplex_of_complex x) = hcomplex_of_complex x"
+by auto
+
+lemma stc_eq_approx:
+     "[| x \<in> HFinite; y \<in> HFinite; stc x = stc y |] ==> x @= y"
+by (auto dest!: stc_approx_self elim!: approx_trans3)
+
+lemma approx_stc_eq:
+     "[| x \<in> HFinite; y \<in> HFinite; x @= y |] ==> stc x = stc y"
+by (blast intro: approx_trans approx_trans2 SComplex_approx_iff [THEN iffD1]
+          dest: stc_approx_self stc_SComplex)
+
+lemma stc_eq_approx_iff:
+     "[| x \<in> HFinite; y \<in> HFinite|] ==> (x @= y) = (stc x = stc y)"
+by (blast intro: approx_stc_eq stc_eq_approx)
+
+lemma stc_Infinitesimal_add_SComplex:
+     "[| x \<in> SComplex; e \<in> Infinitesimal |] ==> stc(x + e) = x"
+apply (erule stc_unique)
+apply (erule Infinitesimal_add_approx_self)
+done
+
+lemma stc_Infinitesimal_add_SComplex2:
+     "[| x \<in> SComplex; e \<in> Infinitesimal |] ==> stc(e + x) = x"
+apply (erule stc_unique)
+apply (erule Infinitesimal_add_approx_self2)
+done
+
+lemma HFinite_stc_Infinitesimal_add:
+     "x \<in> HFinite ==> \<exists>e \<in> Infinitesimal. x = stc(x) + e"
+by (blast dest!: stc_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2])
+
+lemma stc_add:
+     "[| x \<in> HFinite; y \<in> HFinite |] ==> stc (x + y) = stc(x) + stc(y)"
+by (simp add: stc_unique stc_SComplex stc_approx_self approx_add)
+
+lemma stc_number_of [simp]: "stc (number_of w) = number_of w"
+by (rule Standard_number_of [THEN stc_SComplex_eq])
+
+lemma stc_zero [simp]: "stc 0 = 0"
+by simp
+
+lemma stc_one [simp]: "stc 1 = 1"
+by simp
+
+lemma stc_minus: "y \<in> HFinite ==> stc(-y) = -stc(y)"
+by (simp add: stc_unique stc_SComplex stc_approx_self approx_minus)
+
+lemma stc_diff: 
+     "[| x \<in> HFinite; y \<in> HFinite |] ==> stc (x-y) = stc(x) - stc(y)"
+by (simp add: stc_unique stc_SComplex stc_approx_self approx_diff)
+
+lemma stc_mult:
+     "[| x \<in> HFinite; y \<in> HFinite |]  
+               ==> stc (x * y) = stc(x) * stc(y)"
+by (simp add: stc_unique stc_SComplex stc_approx_self approx_mult_HFinite)
+
+lemma stc_Infinitesimal: "x \<in> Infinitesimal ==> stc x = 0"
+by (simp add: stc_unique mem_infmal_iff)
+
+lemma stc_not_Infinitesimal: "stc(x) \<noteq> 0 ==> x \<notin> Infinitesimal"
+by (fast intro: stc_Infinitesimal)
+
+lemma stc_inverse:
+     "[| x \<in> HFinite; stc x \<noteq> 0 |]  
+      ==> stc(inverse x) = inverse (stc x)"
+apply (drule stc_not_Infinitesimal)
+apply (simp add: stc_unique stc_SComplex stc_approx_self approx_inverse)
+done
+
+lemma stc_divide [simp]:
+     "[| x \<in> HFinite; y \<in> HFinite; stc y \<noteq> 0 |]  
+      ==> stc(x/y) = (stc x) / (stc y)"
+by (simp add: divide_inverse stc_mult stc_not_Infinitesimal HFinite_inverse stc_inverse)
+
+lemma stc_idempotent [simp]: "x \<in> HFinite ==> stc(stc(x)) = stc(x)"
+by (blast intro: stc_HFinite stc_approx_self approx_stc_eq)
+
+lemma HFinite_HFinite_hcomplex_of_hypreal:
+     "z \<in> HFinite ==> hcomplex_of_hypreal z \<in> HFinite"
+by (simp add: hcomplex_HFinite_iff)
+
+lemma SComplex_SReal_hcomplex_of_hypreal:
+     "x \<in> Reals ==>  hcomplex_of_hypreal x \<in> SComplex"
+apply (rule Standard_of_hypreal)
+apply (simp add: Reals_eq_Standard)
+done
+
+lemma stc_hcomplex_of_hypreal: 
+ "z \<in> HFinite ==> stc(hcomplex_of_hypreal z) = hcomplex_of_hypreal (st z)"
+apply (rule stc_unique)
+apply (rule SComplex_SReal_hcomplex_of_hypreal)
+apply (erule st_SReal)
+apply (simp add: hcomplex_of_hypreal_approx_iff st_approx_self)
+done
+
+(*
+Goal "x \<in> HFinite ==> hcmod(stc x) = st(hcmod x)"
+by (dtac stc_approx_self 1)
+by (auto_tac (claset(),simpset() addsimps [bex_Infinitesimal_iff2 RS sym]));
+
+
+approx_hcmod_add_hcmod
+*)
+
+lemma Infinitesimal_hcnj_iff [simp]:
+     "(hcnj z \<in> Infinitesimal) = (z \<in> Infinitesimal)"
+by (simp add: Infinitesimal_hcmod_iff)
+
+lemma Infinitesimal_hcomplex_of_hypreal_epsilon [simp]:
+     "hcomplex_of_hypreal epsilon \<in> Infinitesimal"
+by (simp add: Infinitesimal_hcmod_iff)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/NSComplex.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,674 @@
+(*  Title:       NSComplex.thy
+    ID:      $Id$
+    Author:      Jacques D. Fleuriot
+    Copyright:   2001  University of Edinburgh
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
+*)
+
+header{*Nonstandard Complex Numbers*}
+
+theory NSComplex
+imports Complex "../Hyperreal/NSA"
+begin
+
+types hcomplex = "complex star"
+
+abbreviation
+  hcomplex_of_complex :: "complex => complex star" where
+  "hcomplex_of_complex == star_of"
+
+abbreviation
+  hcmod :: "complex star => real star" where
+  "hcmod == hnorm"
+
+
+  (*--- real and Imaginary parts ---*)
+
+definition
+  hRe :: "hcomplex => hypreal" where
+  [code func del]: "hRe = *f* Re"
+
+definition
+  hIm :: "hcomplex => hypreal" where
+  [code func del]: "hIm = *f* Im"
+
+
+  (*------ imaginary unit ----------*)
+
+definition
+  iii :: hcomplex where
+  "iii = star_of ii"
+
+  (*------- complex conjugate ------*)
+
+definition
+  hcnj :: "hcomplex => hcomplex" where
+  [code func del]: "hcnj = *f* cnj"
+
+  (*------------ Argand -------------*)
+
+definition
+  hsgn :: "hcomplex => hcomplex" where
+  [code func del]: "hsgn = *f* sgn"
+
+definition
+  harg :: "hcomplex => hypreal" where
+  [code func del]: "harg = *f* arg"
+
+definition
+  (* abbreviation for (cos a + i sin a) *)
+  hcis :: "hypreal => hcomplex" where
+  [code func del]:"hcis = *f* cis"
+
+  (*----- injection from hyperreals -----*)
+
+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
+  [code func del]: "hrcis = *f2* rcis"
+
+  (*------------ e ^ (x + iy) ------------*)
+definition
+  hexpi :: "hcomplex => hcomplex" where
+  [code func del]: "hexpi = *f* expi"
+
+definition
+  HComplex :: "[hypreal,hypreal] => hcomplex" where
+  [code func del]: "HComplex = *f2* Complex"
+
+lemmas hcomplex_defs [transfer_unfold] =
+  hRe_def hIm_def iii_def hcnj_def hsgn_def harg_def hcis_def
+  hrcis_def hexpi_def HComplex_def
+
+lemma Standard_hRe [simp]: "x \<in> Standard \<Longrightarrow> hRe x \<in> Standard"
+by (simp add: hcomplex_defs)
+
+lemma Standard_hIm [simp]: "x \<in> Standard \<Longrightarrow> hIm x \<in> Standard"
+by (simp add: hcomplex_defs)
+
+lemma Standard_iii [simp]: "iii \<in> Standard"
+by (simp add: hcomplex_defs)
+
+lemma Standard_hcnj [simp]: "x \<in> Standard \<Longrightarrow> hcnj x \<in> Standard"
+by (simp add: hcomplex_defs)
+
+lemma Standard_hsgn [simp]: "x \<in> Standard \<Longrightarrow> hsgn x \<in> Standard"
+by (simp add: hcomplex_defs)
+
+lemma Standard_harg [simp]: "x \<in> Standard \<Longrightarrow> harg x \<in> Standard"
+by (simp add: hcomplex_defs)
+
+lemma Standard_hcis [simp]: "r \<in> Standard \<Longrightarrow> hcis r \<in> Standard"
+by (simp add: hcomplex_defs)
+
+lemma Standard_hexpi [simp]: "x \<in> Standard \<Longrightarrow> hexpi x \<in> Standard"
+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_HComplex [simp]:
+  "\<lbrakk>r \<in> Standard; s \<in> Standard\<rbrakk> \<Longrightarrow> HComplex r s \<in> Standard"
+by (simp add: hcomplex_defs)
+
+lemma hcmod_def: "hcmod = *f* cmod"
+by (rule hnorm_def)
+
+
+subsection{*Properties of Nonstandard Real and Imaginary Parts*}
+
+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_equality [intro?]:
+  "!!z w. hRe z = hRe w ==> hIm z = hIm w ==> z = w"
+by transfer (rule complex_equality)
+
+lemma hcomplex_hRe_zero [simp]: "hRe 0 = 0"
+by transfer (rule complex_Re_zero)
+
+lemma hcomplex_hIm_zero [simp]: "hIm 0 = 0"
+by transfer (rule complex_Im_zero)
+
+lemma hcomplex_hRe_one [simp]: "hRe 1 = 1"
+by transfer (rule complex_Re_one)
+
+lemma hcomplex_hIm_one [simp]: "hIm 1 = 0"
+by transfer (rule complex_Im_one)
+
+
+subsection{*Addition for Nonstandard Complex Numbers*}
+
+lemma hRe_add: "!!x y. hRe(x + y) = hRe(x) + hRe(y)"
+by transfer (rule complex_Re_add)
+
+lemma hIm_add: "!!x y. hIm(x + y) = hIm(x) + hIm(y)"
+by transfer (rule complex_Im_add)
+
+subsection{*More Minus Laws*}
+
+lemma hRe_minus: "!!z. hRe(-z) = - hRe(z)"
+by transfer (rule complex_Re_minus)
+
+lemma hIm_minus: "!!z. hIm(-z) = - hIm(z)"
+by transfer (rule complex_Im_minus)
+
+lemma hcomplex_add_minus_eq_minus:
+      "x + y = (0::hcomplex) ==> x = -y"
+apply (drule OrderedGroup.equals_zero_I)
+apply (simp add: minus_equation_iff [of x y])
+done
+
+lemma hcomplex_i_mult_eq [simp]: "iii * iii = - 1"
+by transfer (rule i_mult_eq2)
+
+lemma hcomplex_i_mult_left [simp]: "!!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)
+
+
+subsection{*More Multiplication Laws*}
+
+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_left_cancel:
+     "(c::hcomplex) \<noteq> (0::hcomplex) ==> (c*a=c*b) = (a=b)"
+by simp
+
+lemma hcomplex_mult_right_cancel:
+     "(c::hcomplex) \<noteq> (0::hcomplex) ==> (a*c=b*c) = (a=b)"
+by simp
+
+
+subsection{*Subraction and Division*}
+
+lemma hcomplex_diff_eq_eq [simp]: "((x::hcomplex) - y = z) = (x = z + y)"
+(* TODO: delete *)
+by (rule OrderedGroup.diff_eq_eq)
+
+
+subsection{*Embedding Properties for @{term hcomplex_of_hypreal} Map*}
+
+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]: "!!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)
+
+subsection{*HComplex theorems*}
+
+lemma hRe_HComplex [simp]: "!!x y. hRe (HComplex x y) = x"
+by transfer (rule Re)
+
+lemma hIm_HComplex [simp]: "!!x y. hIm (HComplex x y) = y"
+by transfer (rule Im)
+
+lemma hcomplex_surj [simp]: "!!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)
+
+
+subsection{*Modulus (Absolute Value) of Nonstandard Complex Number*}
+
+lemma hcomplex_of_hypreal_abs:
+     "hcomplex_of_hypreal (abs x) =
+      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_add [simp]:
+  "!!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_diff [simp]:
+  "!!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)
+
+(*HComplex_inverse is proved below*)
+
+lemma hcomplex_of_hypreal_eq: "!!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)
+
+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)
+
+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)
+
+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)
+
+lemma i_hcomplex_of_hypreal [simp]:
+     "!!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)
+
+
+subsection{*Conjugation*}
+
+lemma hcomplex_hcnj_cancel_iff [iff]: "!!x y. (hcnj x = hcnj y) = (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_hcomplex_of_hypreal [simp]:
+     "!!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_hcnj_minus: "!!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_add: "!!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_mult: "!!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 hcnj_one [simp]: "hcnj 1 = 1"
+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) ^ 2 + hIm(z) ^ 2)"
+by transfer (rule complex_mult_cnj)
+
+
+subsection{*More Theorems about the Function @{term hcmod}*}
+
+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: "!!z. hcmod(z * hcnj(z)) = hcmod(z) ^ 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)
+
+
+subsection{*Exponentiation*}
+
+lemma hcomplexpow_0 [simp]:   "z ^ 0       = (1::hcomplex)"
+by (rule power_0)
+
+lemma hcomplexpow_Suc [simp]: "z ^ (Suc n) = (z::hcomplex) * (z ^ n)"
+by (rule power_Suc)
+
+lemma hcomplexpow_i_squared [simp]: "iii ^ 2 = -1"
+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_hcnj_pow: "!!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 hcpow_minus:
+     "!!x n. (-x::hcomplex) pow n =
+      (if ( *p* even) n then (x pow n) else -(x pow n))"
+by transfer (rule neg_power_if)
+
+lemma hcpow_mult:
+  "!!r s n. ((r::hcomplex) * s) pow n = (r pow n) * (s pow n)"
+by transfer (rule power_mult_distrib)
+
+lemma hcpow_zero2 [simp]:
+  "\<And>n. 0 pow (hSuc n) = (0::'a::{recpower,semiring_0} star)"
+by transfer (rule power_0_Suc)
+
+lemma hcpow_not_zero [simp,intro]:
+  "!!r n. r \<noteq> 0 ==> r pow n \<noteq> (0::hcomplex)"
+by (rule hyperpow_not_zero)
+
+lemma hcpow_zero_zero: "r pow n = (0::hcomplex) ==> r = 0"
+by (blast intro: ccontr dest: hcpow_not_zero)
+
+subsection{*The Function @{term hsgn}*}
+
+lemma hsgn_zero [simp]: "hsgn 0 = 0"
+by transfer (rule sgn_zero)
+
+lemma hsgn_one [simp]: "hsgn 1 = 1"
+by transfer (rule sgn_one)
+
+lemma hsgn_minus: "!!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 hcmod_i: "!!x y. hcmod (HComplex x y) = ( *f* sqrt) (x ^ 2 + y ^ 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)
+
+lemma hcomplex_eq_cancel_iff2 [simp]:
+     "(HComplex x y = hcomplex_of_hypreal xa) = (x = xa & 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_1 [simp]: "!!x y. (HComplex x y = 1) = (x = 1 & y = 0)"
+by transfer (rule Complex_eq_1)
+
+lemma i_eq_HComplex_0_1: "iii = HComplex 0 1"
+by transfer (rule i_def [THEN meta_eq_to_obj_eq])
+
+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 hIm_hsgn [simp]: "!!z. hIm(hsgn z) = hIm(z)/hcmod z"
+by transfer (rule Im_sgn)
+
+lemma hcomplex_inverse_complex_split:
+     "!!x y. inverse(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y) =
+      hcomplex_of_hypreal(x/(x ^ 2 + y ^ 2)) -
+      iii * hcomplex_of_hypreal(y/(x ^ 2 + y ^ 2))"
+by transfer (rule complex_inverse_complex_split)
+
+lemma HComplex_inverse:
+     "!!x y. inverse (HComplex x y) =
+      HComplex (x/(x ^ 2 + y ^ 2)) (-y/(x ^ 2 + y ^ 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_mult_i_eq [simp]:
+    "!!y. hIm (iii * hcomplex_of_hypreal y) = y"
+by transfer simp
+
+lemma hcmod_mult_i [simp]: "!!y. hcmod (iii * hcomplex_of_hypreal y) = abs y"
+by transfer simp
+
+lemma hcmod_mult_i2 [simp]: "!!y. hcmod (hcomplex_of_hypreal y * iii) = abs y"
+by transfer simp
+
+(*---------------------------------------------------------------------------*)
+(*  harg                                                                     *)
+(*---------------------------------------------------------------------------*)
+
+lemma cos_harg_i_mult_zero_pos:
+     "!!y. 0 < y ==> ( *f* cos) (harg(HComplex 0 y)) = 0"
+by transfer (rule cos_arg_i_mult_zero_pos)
+
+lemma cos_harg_i_mult_zero_neg:
+     "!!y. y < 0 ==> ( *f* cos) (harg(HComplex 0 y)) = 0"
+by transfer (rule cos_arg_i_mult_zero_neg)
+
+lemma cos_harg_i_mult_zero [simp]:
+     "!!y. y \<noteq> 0 ==> ( *f* cos) (harg(HComplex 0 y)) = 0"
+by transfer (rule cos_arg_i_mult_zero)
+
+lemma hcomplex_of_hypreal_zero_iff [simp]:
+     "!!y. (hcomplex_of_hypreal y = 0) = (y = 0)"
+by transfer (rule of_real_eq_0_iff)
+
+
+subsection{*Polar Form for Nonstandard Complex Numbers*}
+
+lemma complex_split_polar2:
+     "\<forall>n. \<exists>r a. (z n) =  complex_of_real r * (Complex (cos a) (sin a))"
+by (blast 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 (rule 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: cis_def)
+
+lemma hrcis_Ex: "!!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
+
+lemma hRe_hrcis [simp]: "!!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
+
+lemma hIm_hrcis [simp]: "!!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 (rule cmod_unit_one)
+
+lemma hcmod_complex_polar [simp]:
+  "!!r a. hcmod (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) =
+      abs r"
+by transfer (rule cmod_complex_polar)
+
+lemma hcmod_hrcis [simp]: "!!r a. hcmod(hrcis r a) = abs r"
+by transfer (rule complex_mod_rcis)
+
+(*---------------------------------------------------------------------------*)
+(*  (r1 * hrcis a) * (r2 * hrcis b) = r1 * r2 * hrcis (a + b)                *)
+(*---------------------------------------------------------------------------*)
+
+lemma hcis_hrcis_eq: "!!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 hcis_mult: "!!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)
+
+lemma hrcis_zero_mod [simp]: "!!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 hcomplex_i_mult_minus [simp]: "!!x. iii * (iii * x) = - x"
+by transfer (rule complex_i_mult_minus)
+
+lemma hcomplex_i_mult_minus2 [simp]: "iii * iii * x = - x"
+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 (fold real_of_nat_def)
+apply (rule cis_real_of_nat_Suc_mult)
+done
+
+lemma NSDeMoivre: "!!a. (hcis a) ^ n = hcis (hypreal_of_nat n * a)"
+apply transfer
+apply (fold real_of_nat_def)
+apply (rule DeMoivre)
+done
+
+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 (fold real_of_nat_def, simp add: cis_real_of_nat_Suc_mult)
+
+lemma NSDeMoivre_ext:
+  "!!a n. (hcis a) pow n = hcis (hypreal_of_hypnat n * a)"
+by transfer (fold real_of_nat_def, rule DeMoivre)
+
+lemma NSDeMoivre2:
+  "!!a r. (hrcis r a) ^ n = hrcis (r ^ n) (hypreal_of_nat n * a)"
+by transfer (fold real_of_nat_def, rule DeMoivre2)
+
+lemma DeMoivre2_ext:
+  "!! a r n. (hrcis r a) pow n = hrcis (r pow n) (hypreal_of_hypnat n * a)"
+by transfer (fold real_of_nat_def, rule DeMoivre2)
+
+lemma hcis_inverse [simp]: "!!a. inverse(hcis a) = hcis (-a)"
+by transfer (rule cis_inverse)
+
+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 (rule Re_cis)
+
+lemma hIm_hcis [simp]: "!!a. hIm(hcis a) = ( *f* sin) a"
+by transfer (rule Im_cis)
+
+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 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 hexpi_add: "!!a b. hexpi(a + b) = hexpi(a) * hexpi(b)"
+by transfer (rule expi_add)
+
+
+subsection{*@{term hcomplex_of_complex}: the Injection from
+  type @{typ complex} to to @{typ hcomplex}*}
+
+lemma inj_hcomplex_of_complex: "inj(hcomplex_of_complex)"
+(* TODO: delete *)
+by (rule inj_star_of)
+
+lemma hcomplex_of_complex_i: "iii = hcomplex_of_complex ii"
+by (rule iii_def)
+
+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 hcmod_hcomplex_of_complex:
+     "hcmod (hcomplex_of_complex x) = hypreal_of_real (cmod x)"
+by transfer (rule refl)
+
+
+subsection{*Numerals and Arithmetic*}
+
+lemma hcomplex_number_of_def: "(number_of w :: hcomplex) == of_int w"
+by transfer (rule number_of_eq [THEN eq_reflection])
+
+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)
+
+lemma hcomplex_hypreal_number_of: 
+  "hcomplex_of_complex (number_of w) = hcomplex_of_hypreal(number_of w)"
+by transfer (rule of_real_number_of_eq [symmetric])
+
+lemma hcomplex_number_of_hcnj [simp]:
+     "hcnj (number_of v :: hcomplex) = number_of v"
+by transfer (rule complex_cnj_number_of)
+
+lemma hcomplex_number_of_hcmod [simp]: 
+      "hcmod(number_of v :: hcomplex) = abs (number_of v :: hypreal)"
+by transfer (rule norm_number_of)
+
+lemma hcomplex_number_of_hRe [simp]: 
+      "hRe(number_of v :: hcomplex) = number_of v"
+by transfer (rule complex_Re_number_of)
+
+lemma hcomplex_number_of_hIm [simp]: 
+      "hIm(number_of v :: hcomplex) = 0"
+by transfer (rule complex_Im_number_of)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/NatStar.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,241 @@
+(*  Title       : NatStar.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+
+Converted to Isar and polished by lcp
+*)
+
+header{*Star-transforms for the Hypernaturals*}
+
+theory NatStar
+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)
+
+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 InternalSets_Un:
+     "[| X \<in> InternalSets; Y \<in> InternalSets |]
+      ==> (X Un 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 InternalSets_Int:
+     "[| X \<in> InternalSets; Y \<in> InternalSets |]
+      ==> (X Int 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 InternalSets_Compl: "X \<in> InternalSets ==> -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 InternalSets_diff:
+     "[| X \<in> InternalSets; Y \<in> InternalSets |]
+      ==> (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
+
+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 InternalSets_starset_n [simp]: "( *s* X) \<in> InternalSets"
+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)
+
+
+subsection{*Nonstandard Extensions of Functions*}
+
+text{* Example of transfer of a property from reals to hyperreals
+    --- used for limit comparison of sequences*}
+
+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
+
+(*****----- and another -----*****)
+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
+
+text{*Nonstandard extension when we increment the argument by one*}
+
+lemma starfun_shift_one:
+     "!!N. ( *f* (%n. f (Suc n))) N = ( *f* f) (N + (1::hypnat))"
+by (transfer, simp)
+
+text{*Nonstandard extension with absolute value*}
+
+lemma starfun_abs: "!!N. ( *f* (%n. abs (f n))) N = abs(( *f* f) N)"
+by (transfer, rule refl)
+
+text{*The hyperpow function as a nonstandard extension of realpow*}
+
+lemma starfun_pow: "!!N. ( *f* (%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_pow3: "!!R. ( *f* (%r. r ^ n)) R = (R) pow hypnat_of_nat n"
+by (transfer, rule refl)
+
+text{*The @{term hypreal_of_hypnat} function as a nonstandard extension of
+  @{term real_of_nat} *}
+
+lemma starfunNat_real_of_nat: "( *f* real) = hypreal_of_hypnat"
+by transfer (simp add: expand_fun_eq real_of_nat_def)
+
+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
+
+text{*Internal functions - some redundancy with *f* now*}
+
+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{*Multiplication: @{text "( *fn) x ( *gn) = *(fn x gn)"}*}
+
+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{*Addition: @{text "( *fn) + ( *gn) = *(fn + gn)"}*}
+
+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{*Subtraction: @{text "( *fn) - ( *gn) = *(fn + - gn)"}*}
+
+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{*Composition: @{text "( *fn) o ( *gn) = *(fn o gn)"}*}
+
+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_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_eq [simp]:
+     "( *fn* f) (star_of n) = star_n (%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 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
+
+
+subsection{*Nonstandard Characterization of Induction*}
+
+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)
+
+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)
+
+lemma starP2_eq_iff: "( *p2* (op =)) = (op =)"
+by transfer (rule refl)
+
+lemma starP2_eq_iff2: "( *p2* (%x y. x = y)) X Y = (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_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
+
+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
+
+text{* Goldblatt page 129 Thm 11.3.2*}
+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
+
+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
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/ROOT.ML	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,2 @@
+(* $ID$ *)
+use_thy "Hypercomplex";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/Star.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,348 @@
+(*  Title       : Star.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
+*)
+
+header{*Star-Transforms in Non-Standard Analysis*}
+
+theory Star
+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
+  InternalSets :: "'a star set set" where
+  [code func del]: "InternalSets = {X. \<exists>As. X = *sn* As}"
+
+definition
+  (* nonstandard extension of function *)
+  is_starext  :: "['a star => 'a star, 'a => 'a] => bool" where
+  [code func del]: "is_starext F f = (\<forall>x y. \<exists>X \<in> Rep_star(x). \<exists>Y \<in> Rep_star(y).
+                        ((y = (F x)) = ({n. Y n = f(X n)} : FreeUltrafilterNat)))"
+
+definition
+  (* internal functions *)
+  starfun_n :: "(nat => ('a => 'b)) => 'a star => 'b star"   ("*fn* _" [80] 80) where
+  "*fn* F = Ifun (star_n F)"
+
+definition
+  InternalFuns :: "('a star => 'b star) set" where
+  [code func del]:"InternalFuns = {X. \<exists>F. X = *fn* F}"
+
+
+(*--------------------------------------------------------
+   Preamble - Pulling "EX" over "ALL"
+ ---------------------------------------------------------*)
+
+(* 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
+
+subsection{*Properties of the Star-transform Applied to Sets of Reals*}
+
+lemma STAR_star_of_image_subset: "star_of ` A <= *s* A"
+by auto
+
+lemma STAR_hypreal_of_real_Int: "*s* X Int Reals = 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 lemma_not_hyprealA: "x \<notin> hypreal_of_real ` A ==> \<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_Compl_eq: "- {n. X n = xa} = {n. X n \<noteq> xa}"
+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)
+done
+
+lemma STAR_singleton: "*s* {x} = {star_of x}"
+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)
+
+text{*Nonstandard extension of a set (defined using a constant
+   sequence) as a special case of an internal set*}
+
+lemma starset_n_starset: "\<forall>n. (As n = A) ==> *sn* As = *s* A"
+apply (drule expand_fun_eq [THEN iffD2])
+apply (simp add: starset_n_def starset_def star_of_def)
+done
+
+
+(*----------------------------------------------------------------*)
+(* Theorems about nonstandard extensions of functions             *)
+(*----------------------------------------------------------------*)
+
+(*----------------------------------------------------------------*)
+(* Nonstandard extension of a function (defined using a           *)
+(* constant sequence) as a special case of an internal function   *)
+(*----------------------------------------------------------------*)
+
+lemma starfun_n_starfun: "\<forall>n. (F n = f) ==> *fn* F = *f* f"
+apply (drule expand_fun_eq [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). 
+
+   Proof now Uses the ultrafilter tactic!
+*)
+
+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
+
+text{*Nonstandard extension of functions*}
+
+lemma starfun:
+      "( *f* f) (star_n X) = star_n (%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)
+
+(*-------------------------------------------
+  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)
+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)
+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)
+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)
+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)
+declare starfun_diff [symmetric, simp]
+
+(*--------------------------------------
+  composition: ( *f) o ( *g) = *(f o g)
+ ---------------------------------------*)
+
+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)
+
+text{*NS extension of constant function*}
+lemma starfun_const_fun [simp]: "!!x. ( *f* (%x. k)) x = star_of k"
+by (transfer, rule refl)
+
+text{*the NS extension of the identity function*}
+
+lemma starfun_Id [simp]: "!!x. ( *f* (%x. x)) x = x"
+by (transfer, rule refl)
+
+(* this is trivial, given starfun_Id *)
+lemma starfun_Idfun_approx:
+  "x @= star_of a ==> ( *f* (%x. x)) x @= star_of a"
+by (simp only: starfun_Id)
+
+text{*The Star-function is a (nonstandard) extension of the function*}
+
+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{*Any nonstandard extension is in fact the Star-function*}
+
+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
+
+lemma is_starext_starfun_iff: "(is_starext F f) = (F = *f* f)"
+by (blast intro: is_starfun_starext is_starext_starfun)
+
+text{*extented function has same solution as its standard
+   version for real arguments. i.e they are the same
+   for all real arguments*}
+lemma starfun_eq: "( *f* f) (star_of a) = star_of (f a)"
+by (rule starfun_star_of)
+
+lemma starfun_approx: "( *f* f) (star_of a) @= star_of (f a)"
+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)
+
+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_mult_HFinite_approx:
+  fixes l m :: "'a::real_normed_algebra star"
+  shows "[| ( *f* f) x @= l; ( *f* g) x @= m;
+                  l: HFinite; m: HFinite
+               |] ==>  ( *f* (%x. f x * g x)) x @= l * m"
+apply (drule (3) approx_mult_HFinite)
+apply (auto intro: approx_HFinite [OF _ approx_sym])
+done
+
+lemma starfun_add_approx: "[| ( *f* f) x @= l; ( *f* g) x @= m
+               |] ==>  ( *f* (%x. f x + g x)) x @= l + m"
+by (auto intro: approx_add)
+
+text{*Examples: hrabs is nonstandard extension of rabs
+              inverse is nonstandard extension of inverse*}
+
+(* can be proved easily using theorem "starfun" and *)
+(* properties of ultrafilter as for inverse below we  *)
+(* use the theorem we proved above instead          *)
+
+lemma starfun_rabs_hrabs: "*f* abs = abs"
+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: "!!x. inverse (( *f* f) x) = ( *f* (%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)
+declare starfun_divide [symmetric, simp]
+
+lemma starfun_inverse2: "!!x. inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x"
+by (transfer, rule refl)
+
+text{*General lemma/theorem needed for proofs in elementary
+    topology of the reals*}
+lemma starfun_mem_starset:
+      "!!x. ( *f* f) x : *s* A ==> x : *s* {x. f x  \<in> A}"
+by (transfer, simp)
+
+text{*Alternative definition for hrabs with rabs function
+   applied entrywise to equivalence class representative.
+   This is easily proved using starfun and ns extension thm*}
+lemma hypreal_hrabs:
+     "abs (star_n X) = star_n (%n. abs (X n))"
+by (simp only: starfun_rabs_hrabs [symmetric] starfun)
+
+text{*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
+   by its NS extenson. See second NS set extension below.*}
+lemma STAR_rabs_add_minus:
+   "*s* {x. abs (x + - y) < r} =
+     {x. abs(x + -star_of y) < star_of r}"
+by (transfer, rule refl)
+
+lemma STAR_starfun_rabs_add_minus:
+  "*s* {x. abs (f x + - y) < r} =
+       {x. abs(( *f* f) x + -star_of y) < star_of r}"
+by (transfer, rule refl)
+
+text{*Another characterization of Infinitesimal and one of @= relation.
+   In this theory since @{text hypreal_hrabs} proved here. Maybe
+   move both theorems??*}
+lemma Infinitesimal_FreeUltrafilterNat_iff2:
+     "(star_n X \<in> Infinitesimal) =
+      (\<forall>m. {n. norm(X n) < inverse(real(Suc m))}
+                \<in>  FreeUltrafilterNat)"
+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 real_of_nat_def)
+
+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 real_of_nat_def [symmetric] star_n_inverse real_norm_def
+      HNatInfinite_FreeUltrafilterNat_iff
+      Infinitesimal_FreeUltrafilterNat_iff2)
+apply (drule_tac x="Suc m" in spec)
+apply (erule ultra, simp)
+done
+
+lemma approx_FreeUltrafilterNat_iff: "star_n X @= star_n Y =
+      (\<forall>r>0. {n. norm (X n - Y n) < r} : FreeUltrafilterNat)"
+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 @= star_n Y =
+      (\<forall>m. {n. norm (X n - Y n) <
+                  inverse(real(Suc m))} : FreeUltrafilterNat)"
+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)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/StarDef.thy	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,1029 @@
+(*  Title       : HOL/Hyperreal/StarDef.thy
+    ID          : $Id$
+    Author      : Jacques D. Fleuriot and Brian Huffman
+*)
+
+header {* Construction of Star Types Using Ultrafilters *}
+
+theory StarDef
+imports Filter
+uses ("transfer.ML")
+begin
+
+subsection {* A Free Ultrafilter over the Naturals *}
+
+definition
+  FreeUltrafilterNat :: "nat set set"  ("\<U>") where
+  "\<U> = (SOME U. freeultrafilter U)"
+
+lemma freeultrafilter_FreeUltrafilterNat: "freeultrafilter \<U>"
+apply (unfold FreeUltrafilterNat_def)
+apply (rule someI_ex [where P=freeultrafilter])
+apply (rule freeultrafilter_Ex)
+apply (rule nat_infinite)
+done
+
+interpretation FreeUltrafilterNat: freeultrafilter [FreeUltrafilterNat]
+by (rule freeultrafilter_FreeUltrafilterNat)
+
+text {* This rule takes the place of the old ultra tactic *}
+
+lemma ultra:
+  "\<lbrakk>{n. P n} \<in> \<U>; {n. P n \<longrightarrow> Q n} \<in> \<U>\<rbrakk> \<Longrightarrow> {n. Q n} \<in> \<U>"
+by (simp add: Collect_imp_eq
+    FreeUltrafilterNat.Un_iff FreeUltrafilterNat.Compl_iff)
+
+
+subsection {* Definition of @{text star} type constructor *}
+
+definition
+  starrel :: "((nat \<Rightarrow> 'a) \<times> (nat \<Rightarrow> 'a)) set" where
+  "starrel = {(X,Y). {n. X n = Y n} \<in> \<U>}"
+
+typedef 'a star = "(UNIV :: (nat \<Rightarrow> 'a) set) // starrel"
+by (auto intro: quotientI)
+
+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)
+
+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 ex_star_eq: "(\<exists>x. P x) = (\<exists>X. P (star_n X))"
+by (auto, rule_tac x=x in star_cases, auto)
+
+text {* Proving that @{term starrel} is an equivalence relation *}
+
+lemma starrel_iff [iff]: "((X,Y) \<in> starrel) = ({n. X n = Y n} \<in> \<U>)"
+by (simp add: starrel_def)
+
+lemma equiv_starrel: "equiv UNIV starrel"
+proof (rule equiv.intro)
+  show "reflexive starrel" by (simp add: refl_def)
+  show "sym starrel" by (simp add: sym_def eq_commute)
+  show "trans starrel" by (auto intro: transI elim!: ultra)
+qed
+
+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)
+
+lemma star_n_eq_iff: "(star_n X = star_n Y) = ({n. X n = Y n} \<in> \<U>)"
+by (simp add: star_n_def Abs_star_inject starrel_in_star equiv_starrel_iff)
+
+
+subsection {* Transfer principle *}
+
+text {* This introduction rule starts each transfer proof. *}
+lemma transfer_start:
+  "P \<equiv> {n. Q} \<in> \<U> \<Longrightarrow> Trueprop P \<equiv> Trueprop Q"
+by (subgoal_tac "P \<equiv> Q", simp, simp add: atomize_eq)
+
+text {*Initialize transfer tactic.*}
+use "transfer.ML"
+setup Transfer.setup
+
+text {* Transfer introduction rules. *}
+
+lemma transfer_ex [transfer_intro]:
+  "\<lbrakk>\<And>X. p (star_n X) \<equiv> {n. P n (X n)} \<in> \<U>\<rbrakk>
+    \<Longrightarrow> \<exists>x::'a star. p x \<equiv> {n. \<exists>x. P n x} \<in> \<U>"
+by (simp only: ex_star_eq FreeUltrafilterNat.Collect_ex)
+
+lemma transfer_all [transfer_intro]:
+  "\<lbrakk>\<And>X. p (star_n X) \<equiv> {n. P n (X n)} \<in> \<U>\<rbrakk>
+    \<Longrightarrow> \<forall>x::'a star. p x \<equiv> {n. \<forall>x. P n x} \<in> \<U>"
+by (simp only: all_star_eq FreeUltrafilterNat.Collect_all)
+
+lemma transfer_not [transfer_intro]:
+  "\<lbrakk>p \<equiv> {n. P n} \<in> \<U>\<rbrakk> \<Longrightarrow> \<not> p \<equiv> {n. \<not> P n} \<in> \<U>"
+by (simp only: FreeUltrafilterNat.Collect_not)
+
+lemma transfer_conj [transfer_intro]:
+  "\<lbrakk>p \<equiv> {n. P n} \<in> \<U>; q \<equiv> {n. Q n} \<in> \<U>\<rbrakk>
+    \<Longrightarrow> p \<and> q \<equiv> {n. P n \<and> Q n} \<in> \<U>"
+by (simp only: FreeUltrafilterNat.Collect_conj)
+
+lemma transfer_disj [transfer_intro]:
+  "\<lbrakk>p \<equiv> {n. P n} \<in> \<U>; q \<equiv> {n. Q n} \<in> \<U>\<rbrakk>
+    \<Longrightarrow> p \<or> q \<equiv> {n. P n \<or> Q n} \<in> \<U>"
+by (simp only: FreeUltrafilterNat.Collect_disj)
+
+lemma transfer_imp [transfer_intro]:
+  "\<lbrakk>p \<equiv> {n. P n} \<in> \<U>; q \<equiv> {n. Q n} \<in> \<U>\<rbrakk>
+    \<Longrightarrow> p \<longrightarrow> q \<equiv> {n. P n \<longrightarrow> Q n} \<in> \<U>"
+by (simp only: imp_conv_disj transfer_disj transfer_not)
+
+lemma transfer_iff [transfer_intro]:
+  "\<lbrakk>p \<equiv> {n. P n} \<in> \<U>; q \<equiv> {n. Q n} \<in> \<U>\<rbrakk>
+    \<Longrightarrow> p = q \<equiv> {n. P n = Q n} \<in> \<U>"
+by (simp only: iff_conv_conj_imp transfer_conj transfer_imp)
+
+lemma transfer_if_bool [transfer_intro]:
+  "\<lbrakk>p \<equiv> {n. P n} \<in> \<U>; x \<equiv> {n. X n} \<in> \<U>; y \<equiv> {n. Y n} \<in> \<U>\<rbrakk>
+    \<Longrightarrow> (if p then x else y) \<equiv> {n. if P n then X n else Y n} \<in> \<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> {n. X n = Y n} \<in> \<U>"
+by (simp only: star_n_eq_iff)
+
+lemma transfer_if [transfer_intro]:
+  "\<lbrakk>p \<equiv> {n. P n} \<in> \<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!: ultra)
+done
+
+lemma transfer_fun_eq [transfer_intro]:
+  "\<lbrakk>\<And>X. f (star_n X) = g (star_n X) 
+    \<equiv> {n. F n (X n) = G n (X n)} \<in> \<U>\<rbrakk>
+      \<Longrightarrow> f = g \<equiv> {n. F n = G n} \<in> \<U>"
+by (simp only: expand_fun_eq transfer_all)
+
+lemma transfer_star_n [transfer_intro]: "star_n X \<equiv> star_n (\<lambda>n. X n)"
+by (rule reflexive)
+
+lemma transfer_bool [transfer_intro]: "p \<equiv> {n. p} \<in> \<U>"
+by (simp add: atomize_eq)
+
+
+subsection {* Standard elements *}
+
+definition
+  star_of :: "'a \<Rightarrow> 'a star" where
+  "star_of x == star_n (\<lambda>n. x)"
+
+definition
+  Standard :: "'a star set" where
+  "Standard = range star_of"
+
+text {* Transfer tactic should remove occurrences of @{term star_of} *}
+setup {* Transfer.add_const "StarDef.star_of" *}
+
+declare star_of_def [transfer_intro]
+
+lemma star_of_inject: "(star_of x = star_of y) = (x = y)"
+by (transfer, rule refl)
+
+lemma Standard_star_of [simp]: "star_of x \<in> Standard"
+by (simp add: Standard_def)
+
+
+subsection {* Internal functions *}
+
+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!: ultra)
+
+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])
+
+text {* Transfer tactic should remove occurrences of @{term Ifun} *}
+setup {* Transfer.add_const "StarDef.Ifun" *}
+
+lemma transfer_Ifun [transfer_intro]:
+  "\<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 Ifun_star_of [simp]: "star_of f \<star> star_of x = star_of (f x)"
+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)
+
+text {* Nonstandard extensions of functions *}
+
+definition
+  starfun :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a star \<Rightarrow> 'b star)"  ("*f* _" [80] 80) where
+  "starfun f == \<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 == \<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)
+
+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)
+
+lemma starfun2_star_of [simp]: "( *f2* f) (star_of x) = *f* f x"
+by (transfer, rule refl)
+
+lemma Standard_starfun [simp]: "x \<in> Standard \<Longrightarrow> starfun f x \<in> Standard"
+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_starfun_iff:
+  assumes inj: "\<And>x y. f x = f y \<Longrightarrow> x = y"
+  shows "(starfun f x \<in> Standard) = (x \<in> Standard)"
+proof
+  assume "x \<in> Standard"
+  thus "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
+  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 obtain a where "f a = b" ..
+  hence "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
+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)"
+proof
+  assume "x \<in> Standard \<and> y \<in> Standard"
+  thus "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 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
+qed
+
+
+subsection {* Internal predicates *}
+
+definition unstar :: "bool star \<Rightarrow> bool" where
+  [code func del]: "unstar b \<longleftrightarrow> b = star_of True"
+
+lemma unstar_star_n: "unstar (star_n P) = ({n. P n} \<in> \<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)
+
+text {* Transfer tactic should remove occurrences of @{term unstar} *}
+setup {* Transfer.add_const "StarDef.unstar" *}
+
+lemma transfer_unstar [transfer_intro]:
+  "p \<equiv> star_n P \<Longrightarrow> unstar p \<equiv> {n. P n} \<in> \<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
+  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) = ({n. P (X n)} \<in> \<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) = ({n. P (X n) (Y n)} \<in> \<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)
+
+lemma starP2_star_of [simp]: "( *p2* P) (star_of x) = *p* P x"
+by (transfer, rule refl)
+
+
+subsection {* Internal sets *}
+
+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)) = ({n. X n \<in> A n} \<in> \<U>)"
+by (simp add: Iset_def starP2_star_n)
+
+text {* Transfer tactic should remove occurrences of @{term Iset} *}
+setup {* Transfer.add_const "StarDef.Iset" *}
+
+lemma transfer_mem [transfer_intro]:
+  "\<lbrakk>x \<equiv> star_n X; a \<equiv> Iset (star_n A)\<rbrakk>
+    \<Longrightarrow> x \<in> a \<equiv> {n. X n \<in> A n} \<in> \<U>"
+by (simp only: Iset_star_n)
+
+lemma transfer_Collect [transfer_intro]:
+  "\<lbrakk>\<And>X. p (star_n X) \<equiv> {n. P n (X n)} \<in> \<U>\<rbrakk>
+    \<Longrightarrow> Collect p \<equiv> Iset (star_n (\<lambda>n. Collect (P n)))"
+by (simp add: atomize_eq expand_set_eq 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> {n. A n = B n} \<in> \<U>"
+by (simp only: expand_set_eq 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> {n. P n (X n)} \<in> \<U>\<rbrakk>
+    \<Longrightarrow> \<forall>x\<in>a. p x \<equiv> {n. \<forall>x\<in>A n. P n x} \<in> \<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> {n. P n (X n)} \<in> \<U>\<rbrakk>
+    \<Longrightarrow> \<exists>x\<in>a. p x \<equiv> {n. \<exists>x\<in>A n. P n x} \<in> \<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
+
+text {* Nonstandard extensions of sets. *}
+
+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_UNIV: "*s* (UNIV::'a set) = (UNIV::'a star set)"
+by (transfer UNIV_def, rule refl)
+
+lemma starset_empty: "*s* {} = {}"
+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)
+
+lemma starset_Un: "*s* (A \<union> B) = *s* A \<union> *s* B"
+by (transfer Un_def, rule refl)
+
+lemma starset_Int: "*s* (A \<inter> B) = *s* A \<inter> *s* B"
+by (transfer Int_def, rule refl)
+
+lemma starset_Compl: "*s* -A = -( *s* A)"
+by (transfer Compl_eq, rule refl)
+
+lemma starset_diff: "*s* (A - B) = *s* A - *s* B"
+by (transfer set_diff_eq, rule refl)
+
+lemma starset_image: "*s* (f ` A) = ( *f* f) ` ( *s* A)"
+by (transfer image_def, rule refl)
+
+lemma starset_vimage: "*s* (f -` A) = ( *f* f) -` ( *s* A)"
+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_eq: "( *s* A = *s* B) = (A = B)"
+by (transfer, rule refl)
+
+lemmas starset_simps [simp] =
+  starset_mem     starset_UNIV
+  starset_empty   starset_insert
+  starset_Un      starset_Int
+  starset_Compl   starset_diff
+  starset_image   starset_vimage
+  starset_subset  starset_eq
+
+
+subsection {* Syntactic classes *}
+
+instantiation star :: (zero) zero
+begin
+
+definition
+  star_zero_def [code func del]:    "0 \<equiv> star_of 0"
+
+instance ..
+
+end
+
+instantiation star :: (one) one
+begin
+
+definition
+  star_one_def [code func del]:     "1 \<equiv> star_of 1"
+
+instance ..
+
+end
+
+instantiation star :: (plus) plus
+begin
+
+definition
+  star_add_def [code func del]:     "(op +) \<equiv> *f2* (op +)"
+
+instance ..
+
+end
+
+instantiation star :: (times) times
+begin
+
+definition
+  star_mult_def [code func del]:    "(op *) \<equiv> *f2* (op *)"
+
+instance ..
+
+end
+
+instantiation star :: (uminus) uminus
+begin
+
+definition
+  star_minus_def [code func del]:   "uminus \<equiv> *f* uminus"
+
+instance ..
+
+end
+
+instantiation star :: (minus) minus
+begin
+
+definition
+  star_diff_def [code func del]:    "(op -) \<equiv> *f2* (op -)"
+
+instance ..
+
+end
+
+instantiation star :: (abs) abs
+begin
+
+definition
+  star_abs_def:     "abs \<equiv> *f* abs"
+
+instance ..
+
+end
+
+instantiation star :: (sgn) sgn
+begin
+
+definition
+  star_sgn_def:     "sgn \<equiv> *f* sgn"
+
+instance ..
+
+end
+
+instantiation star :: (inverse) inverse
+begin
+
+definition
+  star_divide_def:  "(op /) \<equiv> *f2* (op /)"
+
+definition
+  star_inverse_def: "inverse \<equiv> *f* inverse"
+
+instance ..
+
+end
+
+instantiation star :: (number) number
+begin
+
+definition
+  star_number_def:  "number_of b \<equiv> star_of (number_of b)"
+
+instance ..
+
+end
+
+instantiation star :: (Divides.div) Divides.div
+begin
+
+definition
+  star_div_def:     "(op div) \<equiv> *f2* (op div)"
+
+definition
+  star_mod_def:     "(op mod) \<equiv> *f2* (op mod)"
+
+instance ..
+
+end
+
+instantiation star :: (power) power
+begin
+
+definition
+  star_power_def:   "(op ^) \<equiv> \<lambda>x n. ( *f* (\<lambda>x. x ^ n)) x"
+
+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 ..
+
+end
+
+lemmas star_class_defs [transfer_unfold] =
+  star_zero_def     star_one_def      star_number_def
+  star_add_def      star_diff_def     star_minus_def
+  star_mult_def     star_divide_def   star_inverse_def
+  star_le_def       star_less_def     star_abs_def       star_sgn_def
+  star_div_def      star_mod_def      star_power_def
+
+text {* Class operations preserve standard elements *}
+
+lemma Standard_zero: "0 \<in> Standard"
+by (simp add: star_zero_def)
+
+lemma Standard_one: "1 \<in> Standard"
+by (simp add: star_one_def)
+
+lemma Standard_number_of: "number_of b \<in> Standard"
+by (simp add: star_number_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_diff: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<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)
+
+lemma Standard_mult: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<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_inverse: "x \<in> Standard \<Longrightarrow> inverse x \<in> Standard"
+by (simp add: star_inverse_def)
+
+lemma Standard_abs: "x \<in> Standard \<Longrightarrow> abs x \<in> Standard"
+by (simp add: star_abs_def)
+
+lemma Standard_div: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x div y \<in> Standard"
+by (simp add: star_div_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_power: "x \<in> Standard \<Longrightarrow> x ^ n \<in> Standard"
+by (simp add: star_power_def)
+
+lemmas Standard_simps [simp] =
+  Standard_zero  Standard_one  Standard_number_of
+  Standard_add  Standard_diff  Standard_minus
+  Standard_mult  Standard_divide  Standard_inverse
+  Standard_abs  Standard_div  Standard_mod
+  Standard_power
+
+text {* @{term star_of} preserves class operations *}
+
+lemma star_of_add: "star_of (x + y) = star_of x + star_of y"
+by transfer (rule refl)
+
+lemma star_of_diff: "star_of (x - y) = star_of x - star_of y"
+by transfer (rule refl)
+
+lemma star_of_minus: "star_of (-x) = - star_of x"
+by transfer (rule refl)
+
+lemma star_of_mult: "star_of (x * y) = star_of x * star_of y"
+by transfer (rule refl)
+
+lemma star_of_divide: "star_of (x / y) = star_of x / star_of y"
+by transfer (rule refl)
+
+lemma star_of_inverse: "star_of (inverse x) = inverse (star_of x)"
+by transfer (rule refl)
+
+lemma star_of_div: "star_of (x div y) = star_of x div star_of y"
+by transfer (rule refl)
+
+lemma star_of_mod: "star_of (x mod y) = star_of x mod star_of y"
+by transfer (rule refl)
+
+lemma star_of_power: "star_of (x ^ n) = star_of x ^ n"
+by transfer (rule refl)
+
+lemma star_of_abs: "star_of (abs x) = abs (star_of x)"
+by transfer (rule refl)
+
+text {* @{term star_of} preserves numerals *}
+
+lemma star_of_zero: "star_of 0 = 0"
+by transfer (rule refl)
+
+lemma star_of_one: "star_of 1 = 1"
+by transfer (rule refl)
+
+lemma star_of_number_of: "star_of (number_of x) = number_of x"
+by transfer (rule refl)
+
+text {* @{term star_of} preserves orderings *}
+
+lemma star_of_less: "(star_of x < star_of y) = (x < y)"
+by transfer (rule refl)
+
+lemma star_of_le: "(star_of x \<le> star_of y) = (x \<le> y)"
+by transfer (rule refl)
+
+lemma star_of_eq: "(star_of x = star_of y) = (x = y)"
+by transfer (rule refl)
+
+text{*As above, for 0*}
+
+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]
+lemmas star_of_0_eq   = star_of_eq   [of 0, simplified star_of_zero]
+
+lemmas star_of_less_0 = star_of_less [of _ 0, simplified star_of_zero]
+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{*As above, for 1*}
+
+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]
+lemmas star_of_1_eq   = star_of_eq   [of 1, simplified star_of_one]
+
+lemmas star_of_less_1 = star_of_less [of _ 1, simplified star_of_one]
+lemmas star_of_le_1   = star_of_le   [of _ 1, simplified star_of_one]
+lemmas star_of_eq_1   = star_of_eq   [of _ 1, simplified star_of_one]
+
+text{*As above, for numerals*}
+
+lemmas star_of_number_less =
+  star_of_less [of "number_of w", standard, simplified star_of_number_of]
+lemmas star_of_number_le   =
+  star_of_le   [of "number_of w", standard, simplified star_of_number_of]
+lemmas star_of_number_eq   =
+  star_of_eq   [of "number_of w", standard, simplified star_of_number_of]
+
+lemmas star_of_less_number =
+  star_of_less [of _ "number_of w", standard, simplified star_of_number_of]
+lemmas star_of_le_number   =
+  star_of_le   [of _ "number_of w", standard, simplified star_of_number_of]
+lemmas star_of_eq_number   =
+  star_of_eq   [of _ "number_of w", standard, simplified star_of_number_of]
+
+lemmas star_of_simps [simp] =
+  star_of_add     star_of_diff    star_of_minus
+  star_of_mult    star_of_divide  star_of_inverse
+  star_of_div     star_of_mod
+  star_of_power   star_of_abs
+  star_of_zero    star_of_one     star_of_number_of
+  star_of_less    star_of_le      star_of_eq
+  star_of_0_less  star_of_0_le    star_of_0_eq
+  star_of_less_0  star_of_le_0    star_of_eq_0
+  star_of_1_less  star_of_1_le    star_of_1_eq
+  star_of_less_1  star_of_le_1    star_of_eq_1
+  star_of_number_less star_of_number_le star_of_number_eq
+  star_of_less_number star_of_le_number star_of_eq_number
+
+subsection {* Ordering and lattice classes *}
+
+instance star :: (order) order
+apply (intro_classes)
+apply (transfer, rule order_less_le)
+apply (transfer, rule order_refl)
+apply (transfer, erule (1) order_trans)
+apply (transfer, erule (1) order_antisym)
+done
+
+instantiation star :: (lower_semilattice) lower_semilattice
+begin
+
+definition
+  star_inf_def [transfer_unfold]: "inf \<equiv> *f2* inf"
+
+instance
+  by default (transfer star_inf_def, auto)+
+
+end
+
+instantiation star :: (upper_semilattice) upper_semilattice
+begin
+
+definition
+  star_sup_def [transfer_unfold]: "sup \<equiv> *f2* sup"
+
+instance
+  by default (transfer star_sup_def, auto)+
+
+end
+
+instance star :: (lattice) lattice ..
+
+instance star :: (distrib_lattice) distrib_lattice
+  by default (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_sup [simp]:
+  "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<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)
+
+lemma star_of_sup [simp]: "star_of (sup x y) = sup (star_of x) (star_of y)"
+by transfer (rule refl)
+
+instance star :: (linorder) linorder
+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
+
+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
+
+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_min [simp]:
+  "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<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)
+
+lemma star_of_min [simp]: "star_of (min x y) = min (star_of x) (star_of y)"
+by transfer (rule refl)
+
+
+subsection {* Ordered group classes *}
+
+instance star :: (semigroup_add) semigroup_add
+by (intro_classes, transfer, rule add_assoc)
+
+instance star :: (ab_semigroup_add) ab_semigroup_add
+by (intro_classes, transfer, rule add_commute)
+
+instance star :: (semigroup_mult) semigroup_mult
+by (intro_classes, transfer, rule mult_assoc)
+
+instance star :: (ab_semigroup_mult) ab_semigroup_mult
+by (intro_classes, transfer, rule mult_commute)
+
+instance star :: (comm_monoid_add) comm_monoid_add
+by (intro_classes, transfer, rule comm_monoid_add_class.zero_plus.add_0)
+
+instance star :: (monoid_mult) monoid_mult
+apply (intro_classes)
+apply (transfer, rule mult_1_left)
+apply (transfer, rule mult_1_right)
+done
+
+instance star :: (comm_monoid_mult) comm_monoid_mult
+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
+
+instance star :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add
+by (intro_classes, transfer, rule add_imp_eq)
+
+instance star :: (ab_group_add) ab_group_add
+apply (intro_classes)
+apply (transfer, rule left_minus)
+apply (transfer, rule diff_minus)
+done
+
+instance star :: (pordered_ab_semigroup_add) pordered_ab_semigroup_add
+by (intro_classes, transfer, rule add_left_mono)
+
+instance star :: (pordered_cancel_ab_semigroup_add) pordered_cancel_ab_semigroup_add ..
+
+instance star :: (pordered_ab_semigroup_add_imp_le) pordered_ab_semigroup_add_imp_le
+by (intro_classes, transfer, rule add_le_imp_le_left)
+
+instance star :: (pordered_comm_monoid_add) pordered_comm_monoid_add ..
+instance star :: (pordered_ab_group_add) pordered_ab_group_add ..
+
+instance star :: (pordered_ab_group_add_abs) pordered_ab_group_add_abs 
+  by intro_classes (transfer,
+    simp add: abs_ge_self abs_leI abs_triangle_ineq)+
+
+instance star :: (ordered_cancel_ab_semigroup_add) ordered_cancel_ab_semigroup_add ..
+instance star :: (lordered_ab_group_add_meet) lordered_ab_group_add_meet ..
+instance star :: (lordered_ab_group_add_meet) lordered_ab_group_add_meet ..
+instance star :: (lordered_ab_group_add) lordered_ab_group_add ..
+
+instance star :: (lordered_ab_group_add_abs) lordered_ab_group_add_abs
+by (intro_classes, transfer, rule abs_lattice)
+
+subsection {* Ring and field classes *}
+
+instance star :: (semiring) semiring
+apply (intro_classes)
+apply (transfer, rule left_distrib)
+apply (transfer, rule right_distrib)
+done
+
+instance star :: (semiring_0) semiring_0 
+by intro_classes (transfer, simp)+
+
+instance star :: (semiring_0_cancel) semiring_0_cancel ..
+
+instance star :: (comm_semiring) comm_semiring 
+by (intro_classes, transfer, rule left_distrib)
+
+instance star :: (comm_semiring_0) comm_semiring_0 ..
+instance star :: (comm_semiring_0_cancel) comm_semiring_0_cancel ..
+
+instance star :: (zero_neq_one) zero_neq_one
+by (intro_classes, transfer, rule zero_neq_one)
+
+instance star :: (semiring_1) semiring_1 ..
+instance star :: (comm_semiring_1) comm_semiring_1 ..
+
+instance star :: (no_zero_divisors) no_zero_divisors
+by (intro_classes, transfer, rule no_zero_divisors)
+
+instance star :: (semiring_1_cancel) semiring_1_cancel ..
+instance star :: (comm_semiring_1_cancel) comm_semiring_1_cancel ..
+instance star :: (ring) ring ..
+instance star :: (comm_ring) comm_ring ..
+instance star :: (ring_1) ring_1 ..
+instance star :: (comm_ring_1) comm_ring_1 ..
+instance star :: (ring_no_zero_divisors) ring_no_zero_divisors ..
+instance star :: (ring_1_no_zero_divisors) ring_1_no_zero_divisors ..
+instance star :: (idom) idom .. 
+
+instance star :: (division_ring) division_ring
+apply (intro_classes)
+apply (transfer, erule left_inverse)
+apply (transfer, erule right_inverse)
+done
+
+instance star :: (field) field
+apply (intro_classes)
+apply (transfer, erule left_inverse)
+apply (transfer, rule divide_inverse)
+done
+
+instance star :: (division_by_zero) division_by_zero
+by (intro_classes, transfer, rule inverse_zero)
+
+instance star :: (pordered_semiring) pordered_semiring
+apply (intro_classes)
+apply (transfer, erule (1) mult_left_mono)
+apply (transfer, erule (1) mult_right_mono)
+done
+
+instance star :: (pordered_cancel_semiring) pordered_cancel_semiring ..
+
+instance star :: (ordered_semiring_strict) ordered_semiring_strict
+apply (intro_classes)
+apply (transfer, erule (1) mult_strict_left_mono)
+apply (transfer, erule (1) mult_strict_right_mono)
+done
+
+instance star :: (pordered_comm_semiring) pordered_comm_semiring
+by (intro_classes, transfer, rule mult_mono1_class.less_eq_less_times_zero.mult_mono1)
+
+instance star :: (pordered_cancel_comm_semiring) pordered_cancel_comm_semiring ..
+
+instance star :: (ordered_comm_semiring_strict) ordered_comm_semiring_strict
+by (intro_classes, transfer, rule ordered_comm_semiring_strict_class.plus_less_eq_less_zero_times.mult_strict_left_mono_comm)
+
+instance star :: (pordered_ring) pordered_ring ..
+instance star :: (pordered_ring_abs) pordered_ring_abs
+  by intro_classes  (transfer, rule abs_eq_mult)
+instance star :: (lordered_ring) lordered_ring ..
+
+instance star :: (abs_if) abs_if
+by (intro_classes, transfer, rule abs_if)
+
+instance star :: (sgn_if) sgn_if
+by (intro_classes, transfer, rule sgn_if)
+
+instance star :: (ordered_ring_strict) ordered_ring_strict ..
+instance star :: (pordered_comm_ring) pordered_comm_ring ..
+
+instance star :: (ordered_semidom) ordered_semidom
+by (intro_classes, transfer, rule zero_less_one)
+
+instance star :: (ordered_idom) ordered_idom ..
+instance star :: (ordered_field) ordered_field ..
+
+subsection {* Power classes *}
+
+text {*
+  Proving the class axiom @{thm [source] power_Suc} for type
+  @{typ "'a star"} is a little tricky, because it quantifies
+  over values of type @{typ nat}. The transfer principle does
+  not handle quantification over non-star types in general,
+  but we can work around this by fixing an arbitrary @{typ nat}
+  value, and then applying the transfer principle.
+*}
+
+instance star :: (recpower) recpower
+proof
+  show "\<And>a::'a star. a ^ 0 = 1"
+    by transfer (rule power_0)
+next
+  fix n show "\<And>a::'a star. a ^ Suc n = a * a ^ n"
+    by transfer (rule power_Suc)
+qed
+
+subsection {* Number classes *}
+
+lemma star_of_nat_def [transfer_unfold]: "of_nat n = star_of (of_nat n)"
+by (induct n, simp_all)
+
+lemma Standard_of_nat [simp]: "of_nat n \<in> Standard"
+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)
+
+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)
+
+lemma Standard_of_int [simp]: "of_int z \<in> Standard"
+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)
+
+instance star :: (semiring_char_0) semiring_char_0
+by intro_classes (simp only: star_of_nat_def star_of_eq of_nat_eq_iff)
+
+instance star :: (ring_char_0) ring_char_0 ..
+
+instance star :: (number_ring) number_ring
+by (intro_classes, simp only: star_number_def star_of_int_def number_of_eq)
+
+subsection {* Finite class *}
+
+lemma starset_finite: "finite A \<Longrightarrow> *s* A = star_of ` A"
+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
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/document/root.tex	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,32 @@
+
+% $Id$
+
+\documentclass[11pt,a4paper]{article}
+\usepackage{graphicx,isabelle,isabellesym,latexsym}
+\usepackage[latin1]{inputenc}
+\usepackage{pdfsetup}
+
+\urlstyle{rm}
+\isabellestyle{it}
+\pagestyle{myheadings}
+
+\begin{document}
+
+\title{Isabelle/HOL-NSA --- Non-Standard Analysis}
+\maketitle
+
+\tableofcontents
+
+\begin{center}
+  \includegraphics[width=\textwidth,height=\textheight,keepaspectratio]{session_graph}
+\end{center}
+
+\newpage
+
+\renewcommand{\isamarkupheader}[1]%
+{\section{\isabellecontext: #1}\markright{THEORY~``\isabellecontext''}}
+
+\parindent 0pt\parskip 0.5ex
+\input{session}
+
+\end{document}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/hypreal_arith.ML	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,46 @@
+(*  Title:      HOL/Hyperreal/hypreal_arith.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow, TU Muenchen
+    Copyright   1999 TU Muenchen
+
+Simprocs for common factor cancellation & Rational coefficient handling
+
+Instantiation of the generic linear arithmetic package for type hypreal.
+*)
+
+local
+
+val simps = [thm "star_of_zero",
+             thm "star_of_one",
+             thm "star_of_number_of",
+             thm "star_of_add",
+             thm "star_of_minus",
+             thm "star_of_diff",
+             thm "star_of_mult"]
+
+val real_inj_thms = [thm "star_of_le" RS iffD2,
+                     thm "star_of_less" RS iffD2,
+                     thm "star_of_eq" RS iffD2]
+
+in
+
+val hyprealT = Type ("StarDef.star", [HOLogic.realT]);
+
+val fast_hypreal_arith_simproc =
+    Simplifier.simproc (the_context ())
+      "fast_hypreal_arith" 
+      ["(m::hypreal) < n", "(m::hypreal) <= n", "(m::hypreal) = n"]
+    (K LinArith.lin_arith_simproc);
+
+val hypreal_arith_setup =
+  LinArith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+   {add_mono_thms = add_mono_thms,
+    mult_mono_thms = mult_mono_thms,
+    inj_thms = real_inj_thms @ inj_thms,
+    lessD = lessD,  (*Can't change lessD: the hypreals are dense!*)
+    neqE = neqE,
+    simpset = simpset addsimps simps}) #>
+  arith_inj_const ("StarDef.star_of", HOLogic.realT --> hyprealT) #>
+  Simplifier.map_ss (fn ss => ss addsimprocs [fast_hypreal_arith_simproc]);
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NSA/transfer.ML	Thu Jul 03 17:47:22 2008 +0200
@@ -0,0 +1,121 @@
+(*  Title       : HOL/Hyperreal/transfer.ML
+    ID          : $Id$
+    Author      : Brian Huffman
+
+Transfer principle tactic for nonstandard analysis.
+*)
+
+signature TRANSFER =
+sig
+  val transfer_tac: Proof.context -> thm list -> int -> tactic
+  val add_const: string -> theory -> theory
+  val setup: theory -> theory
+end;
+
+structure Transfer: TRANSFER =
+struct
+
+structure TransferData = GenericDataFun
+(
+  type T = {
+    intros: thm list,
+    unfolds: thm list,
+    refolds: thm list,
+    consts: string list
+  };
+  val empty = {intros = [], unfolds = [], refolds = [], consts = []};
+  val extend = I;
+  fun merge _
+    ({intros = intros1, unfolds = unfolds1,
+      refolds = refolds1, consts = consts1},
+     {intros = intros2, unfolds = unfolds2,
+      refolds = refolds2, consts = consts2}) =
+   {intros = Thm.merge_thms (intros1, intros2),
+    unfolds = Thm.merge_thms (unfolds1, unfolds2),
+    refolds = Thm.merge_thms (refolds1, refolds2),
+    consts = Library.merge (op =) (consts1, consts2)} : T;
+);
+
+fun unstar_typ (Type ("StarDef.star",[t])) = unstar_typ t
+  | unstar_typ (Type (a, Ts)) = Type (a, map unstar_typ Ts)
+  | unstar_typ T = T
+
+fun unstar_term consts term =
+  let
+    fun unstar (Const(a,T) $ t) = if member (op =) consts a then (unstar t)
+          else (Const(a,unstar_typ T) $ unstar t)
+      | unstar (f $ t) = unstar f $ unstar t
+      | unstar (Const(a,T)) = Const(a,unstar_typ T)
+      | unstar (Abs(a,T,t)) = Abs(a,unstar_typ T,unstar t)
+      | unstar t = t
+  in
+    unstar term
+  end
+
+fun transfer_thm_of ctxt ths t =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val {intros,unfolds,refolds,consts} = TransferData.get (Context.Proof ctxt);
+    val meta = LocalDefs.meta_rewrite_rule ctxt;
+    val ths' = map meta ths;
+    val unfolds' = map meta unfolds and refolds' = map meta refolds;
+    val (_$_$t') = concl_of (MetaSimplifier.rewrite true unfolds' (cterm_of thy t))
+    val u = unstar_term consts t'
+    val tac =
+      rewrite_goals_tac (ths' @ refolds' @ unfolds') THEN
+      ALLGOALS ObjectLogic.full_atomize_tac THEN
+      match_tac [transitive_thm] 1 THEN
+      resolve_tac [@{thm transfer_start}] 1 THEN
+      REPEAT_ALL_NEW (resolve_tac intros) 1 THEN
+      match_tac [reflexive_thm] 1
+  in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end
+
+fun transfer_tac ctxt ths =
+    SUBGOAL (fn (t,i) =>
+        (fn th =>
+            let
+              val tr = transfer_thm_of ctxt ths t
+              val (_$l$r) = concl_of tr;
+              val trs = if l aconv r then [] else [tr];
+            in rewrite_goals_tac trs th end))
+
+local
+
+fun map_intros f = TransferData.map
+  (fn {intros,unfolds,refolds,consts} =>
+    {intros=f intros, unfolds=unfolds, refolds=refolds, consts=consts})
+
+fun map_unfolds f = TransferData.map
+  (fn {intros,unfolds,refolds,consts} =>
+    {intros=intros, unfolds=f unfolds, refolds=refolds, consts=consts})
+
+fun map_refolds f = TransferData.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);
+
+val unfold_add = Thm.declaration_attribute (map_unfolds o Thm.add_thm);
+val unfold_del = Thm.declaration_attribute (map_unfolds o Thm.del_thm);
+
+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
+  (fn {intros,unfolds,refolds,consts} =>
+    {intros=intros, unfolds=unfolds, refolds=refolds, consts=c::consts}))
+
+val setup =
+  Attrib.add_attributes
+    [("transfer_intro", Attrib.add_del_args intro_add intro_del,
+      "declaration of transfer introduction rule"),
+     ("transfer_unfold", Attrib.add_del_args unfold_add unfold_del,
+      "declaration of transfer unfolding rule"),
+     ("transfer_refold", Attrib.add_del_args refold_add refold_del,
+      "declaration of transfer refolding rule")] #>
+  Method.add_method ("transfer", Method.thms_ctxt_args (fn ths => fn ctxt =>
+    Method.SIMPLE_METHOD' (transfer_tac ctxt ths)), "transfer principle");
+
+end;