separation of HOL-Hyperreal from HOL-Real
authorpaulson
Sat, 30 Dec 2000 22:03:47 +0100
changeset 10751 a81ea5d3dd41
parent 10750 a681d3df1a39
child 10752 c4f1bf2acf4c
separation of HOL-Hyperreal from HOL-Real
src/HOL/Hyperreal/HSeries.ML
src/HOL/Hyperreal/HSeries.thy
src/HOL/Hyperreal/HyperArith.thy
src/HOL/Hyperreal/HyperArith0.ML
src/HOL/Hyperreal/HyperArith0.thy
src/HOL/Hyperreal/HyperBin.ML
src/HOL/Hyperreal/HyperBin.thy
src/HOL/Hyperreal/HyperDef.ML
src/HOL/Hyperreal/HyperDef.thy
src/HOL/Hyperreal/HyperNat.ML
src/HOL/Hyperreal/HyperNat.thy
src/HOL/Hyperreal/HyperOrd.ML
src/HOL/Hyperreal/HyperOrd.thy
src/HOL/Hyperreal/HyperPow.ML
src/HOL/Hyperreal/HyperPow.thy
src/HOL/Hyperreal/Hyperreal.thy
src/HOL/Hyperreal/Lim.ML
src/HOL/Hyperreal/Lim.thy
src/HOL/Hyperreal/NSA.ML
src/HOL/Hyperreal/NSA.thy
src/HOL/Hyperreal/NatStar.ML
src/HOL/Hyperreal/NatStar.thy
src/HOL/Hyperreal/ROOT.ML
src/HOL/Hyperreal/SEQ.ML
src/HOL/Hyperreal/SEQ.thy
src/HOL/Hyperreal/Series.ML
src/HOL/Hyperreal/Series.thy
src/HOL/Hyperreal/Star.ML
src/HOL/Hyperreal/Star.thy
src/HOL/Hyperreal/Zorn.ML
src/HOL/Hyperreal/Zorn.thy
src/HOL/Hyperreal/fuf.ML
src/HOL/Hyperreal/hypreal_arith.ML
src/HOL/Hyperreal/hypreal_arith0.ML
src/HOL/IsaMakefile
src/HOL/Real/Hyperreal/Filter.ML
src/HOL/Real/Hyperreal/Filter.thy
src/HOL/Real/Hyperreal/HRealAbs.ML
src/HOL/Real/Hyperreal/HRealAbs.thy
src/HOL/Real/Hyperreal/HyperDef.ML
src/HOL/Real/Hyperreal/HyperDef.thy
src/HOL/Real/Hyperreal/HyperNat.ML
src/HOL/Real/Hyperreal/HyperNat.thy
src/HOL/Real/Hyperreal/HyperOrd.ML
src/HOL/Real/Hyperreal/HyperOrd.thy
src/HOL/Real/Hyperreal/HyperPow.ML
src/HOL/Real/Hyperreal/HyperPow.thy
src/HOL/Real/Hyperreal/Hyperreal.thy
src/HOL/Real/Hyperreal/Lim.ML
src/HOL/Real/Hyperreal/Lim.thy
src/HOL/Real/Hyperreal/NSA.ML
src/HOL/Real/Hyperreal/NSA.thy
src/HOL/Real/Hyperreal/NatStar.ML
src/HOL/Real/Hyperreal/NatStar.thy
src/HOL/Real/Hyperreal/SEQ.ML
src/HOL/Real/Hyperreal/SEQ.thy
src/HOL/Real/Hyperreal/Series.ML
src/HOL/Real/Hyperreal/Series.thy
src/HOL/Real/Hyperreal/Star.ML
src/HOL/Real/Hyperreal/Star.thy
src/HOL/Real/Hyperreal/Zorn.ML
src/HOL/Real/Hyperreal/Zorn.thy
src/HOL/Real/Hyperreal/fuf.ML
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HSeries.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,321 @@
+(*  Title       : HSeries.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Finite summation and infinite series
+                  for hyperreals
+*) 
+
+Goalw [sumhr_def]
+     "sumhr(M,N,f) =  \
+\       Abs_hypreal(UN X:Rep_hypnat(M). UN Y: Rep_hypnat(N). \
+\         hyprel ^^{%n::nat. sumr (X n) (Y n) f})";
+by (Auto_tac);
+qed "sumhr_iff";
+
+Goalw [sumhr_def]
+     "sumhr(Abs_hypnat(hypnatrel^^{%n. M n}), \
+\           Abs_hypnat(hypnatrel^^{%n. N n}), f) = \
+\     Abs_hypreal(hyprel ^^ {%n. sumr (M n) (N n) f})";
+by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
+by (Auto_tac THEN Ultra_tac 1);
+qed "sumhr";
+
+(*------------------------------------------------------- 
+  lcp's suggestion: exploit pattern matching 
+  facilities and use as definition instead (to do)
+ -------------------------------------------------------*)
+Goalw [sumhr_def]
+      "sumhr p = (%(M,N,f). Abs_hypreal(UN X:Rep_hypnat(M). \
+\                     UN Y: Rep_hypnat(N). \
+\           hyprel ^^{%n::nat. sumr (X n) (Y n) f})) p";
+by (res_inst_tac [("p","p")] PairE 1);
+by (res_inst_tac [("p","y")] PairE 1);
+by (Auto_tac);
+qed "sumhr_iff2";
+
+(* Theorem corresponding to base case in def of sumr *)
+Goalw [hypnat_zero_def]
+     "sumhr (m,0,f) = #0";
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), 
+              simpset() addsimps [sumhr, symmetric hypreal_zero_def]));
+qed "sumhr_zero";
+Addsimps [sumhr_zero];
+
+(* Theorem corresponding to recursive case in def of sumr *)
+Goalw [hypnat_one_def]
+     "sumhr(m,n+1hn,f) = (if n + 1hn <= m then #0 \
+\                         else sumhr(m,n,f) + (*fNat* f) n)";
+by (simp_tac (HOL_ss addsimps
+             [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1); 
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+     simpset() addsimps [sumhr, hypnat_add,hypnat_le,starfunNat,hypreal_add]));
+by (ALLGOALS(Ultra_tac));
+qed "sumhr_if";
+
+Goalw [hypnat_one_def] "sumhr (n + 1hn, n, f) = #0";
+by (simp_tac (HOL_ss addsimps
+             [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1); 
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+              simpset() addsimps [sumhr, hypnat_add]));
+qed "sumhr_Suc_zero";
+Addsimps [sumhr_Suc_zero];
+
+Goal "sumhr (n,n,f) = #0";
+by (simp_tac (HOL_ss addsimps
+             [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1); 
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [sumhr]));
+qed "sumhr_eq_bounds";
+Addsimps [sumhr_eq_bounds];
+
+Goalw [hypnat_one_def] 
+     "sumhr (m,m + 1hn,f) = (*fNat* f) m";
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+              simpset() addsimps [sumhr, hypnat_add,starfunNat]));
+qed "sumhr_Suc";
+Addsimps [sumhr_Suc];
+
+Goal "sumhr(m+k,k,f) = #0";
+by (simp_tac (HOL_ss addsimps
+             [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1); 
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","k")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+              simpset() addsimps [sumhr, hypnat_add]));
+qed "sumhr_add_lbound_zero";
+Addsimps [sumhr_add_lbound_zero];
+
+Goal "sumhr (m,n,f) + sumhr(m,n,g) = sumhr(m,n,%i. f i + g i)";
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+              simpset() addsimps [sumhr, hypreal_add,sumr_add]));
+qed "sumhr_add";
+
+Goalw [hypreal_of_real_def]
+      "hypreal_of_real r * sumhr(m,n,f) = sumhr(m,n,%n. r * f n)";
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+              simpset() addsimps [sumhr, hypreal_mult,sumr_mult]));
+qed "sumhr_mult";
+
+Goalw [hypnat_zero_def]
+     "n < p ==> sumhr (0,n,f) + sumhr (n,p,f) = sumhr (0,p,f)";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","p")] eq_Abs_hypnat 1);
+by (auto_tac (claset() addSEs [FreeUltrafilterNat_subset],
+       simpset() addsimps [sumhr,hypreal_add,hypnat_less, sumr_split_add]));
+qed "sumhr_split_add";
+
+(*FIXME delete*)
+Goal "n < p ==> sumhr (0, p, f) + - sumhr (0, n, f) = sumhr (n,p,f)";
+by (dres_inst_tac [("f1","f")] (sumhr_split_add RS sym) 1);
+by (Asm_simp_tac 1);
+qed "sumhr_split_add_minus";
+
+Goal "n < p ==> sumhr (0, p, f) - sumhr (0, n, f) = sumhr (n,p,f)";
+by (dres_inst_tac [("f1","f")] (sumhr_split_add RS sym) 1);
+by (Asm_simp_tac 1);
+qed "sumhr_split_diff";
+
+Goal "abs(sumhr(m,n,f)) <= sumhr(m,n,%i. abs(f i))";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+              simpset() addsimps [sumhr, hypreal_le,hypreal_hrabs,sumr_rabs]));
+qed "sumhr_hrabs";
+
+(* other general version also needed *)
+Goalw [hypnat_of_nat_def]
+     "(ALL r. m <= 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)";
+by (Step_tac 1 THEN dtac sumr_fun_eq 1);
+by (auto_tac (claset(), simpset() addsimps [sumhr]));
+qed "sumhr_fun_hypnat_eq";
+
+Goalw [hypnat_zero_def,hypreal_of_real_def]
+      "sumhr(0,n,%i. r) = hypreal_of_hypnat n*hypreal_of_real r";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (asm_simp_tac
+    (simpset() addsimps [sumhr, hypreal_of_hypnat,hypreal_mult]) 1);
+qed "sumhr_const";
+
+Goalw [hypnat_zero_def,hypreal_of_real_def]
+     "sumhr(0,n,f) + -(hypreal_of_hypnat n*hypreal_of_real r) = \
+\     sumhr(0,n,%i. f i + -r)";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (asm_simp_tac (simpset() addsimps [sumhr,
+		    hypreal_of_hypnat,hypreal_mult,hypreal_add,
+		    hypreal_minus,sumr_add RS sym]) 1);
+qed "sumhr_add_mult_const";
+
+Goal "n < m ==> sumhr (m,n,f) = #0";
+by (simp_tac (HOL_ss addsimps
+             [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1); 
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset() addEs [FreeUltrafilterNat_subset],
+              simpset() addsimps [sumhr,hypnat_less]));
+qed "sumhr_less_bounds_zero";
+Addsimps [sumhr_less_bounds_zero];
+
+Goal "sumhr(m, n, %i. - f i) = - sumhr(m, n, f)";
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [sumhr, hypreal_minus,sumr_minus]));
+qed "sumhr_minus";
+
+Goalw [hypnat_of_nat_def]
+     "sumhr(m+hypnat_of_nat k,n+hypnat_of_nat k,f) = sumhr(m,n,%i. f(i + k))";
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+              simpset() addsimps [sumhr, hypnat_add,sumr_shift_bounds]));
+qed "sumhr_shift_bounds";
+
+(*------------------------------------------------------------------
+      Theorems about NS sums - infinite sums are obtained
+      by summing to some infinite hypernatural (such as whn)
+ -----------------------------------------------------------------*)
+Goalw [hypnat_omega_def,hypnat_zero_def] 
+      "sumhr(0,whn,%i. #1) = hypreal_of_hypnat whn";
+by (auto_tac (claset(),
+              simpset() addsimps [sumhr, hypreal_of_hypnat]));
+qed "sumhr_hypreal_of_hypnat_omega";
+
+Goalw [hypnat_omega_def,hypnat_zero_def,omega_def]  
+     "sumhr(0, whn, %i. #1) = whr - #1";
+by (simp_tac (HOL_ss addsimps
+             [one_eq_numeral_1 RS sym, hypreal_one_def]) 1); 
+by (auto_tac (claset(),
+              simpset() addsimps [sumhr, hypreal_diff, real_of_nat_def]));
+qed "sumhr_hypreal_omega_minus_one";
+
+Goalw [hypnat_zero_def, hypnat_omega_def]
+     "sumhr(0, whn + whn, %i. (-#1) ^ (i+1)) = #0";
+by (simp_tac (HOL_ss addsimps
+             [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1); 
+by (simp_tac (simpset() addsimps [sumhr,hypnat_add] delsimps [realpow_Suc]) 1);
+qed "sumhr_minus_one_realpow_zero";
+Addsimps [sumhr_minus_one_realpow_zero];
+
+Goalw [hypnat_of_nat_def,hypreal_of_real_def]
+     "(ALL n. m <= Suc n --> f n = r) & m <= na \
+\          ==> sumhr(hypnat_of_nat m,hypnat_of_nat na,f) = \
+\          (hypreal_of_nat (na - m) * hypreal_of_real r)";
+by (auto_tac (claset() addSDs [sumr_interval_const],
+    simpset() addsimps [sumhr,hypreal_of_nat_real_of_nat,
+              hypreal_of_real_def,hypreal_mult]));
+qed "sumhr_interval_const";
+
+Goalw [hypnat_zero_def]
+     "(*fNat* (%n. sumr 0 n f)) N = sumhr(0,N,f)";
+by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
+by (asm_full_simp_tac (simpset() addsimps [starfunNat,sumhr]) 1);
+qed "starfunNat_sumr";
+
+Goal "sumhr (0, M, f) @= sumhr (0, N, f) \
+\     ==> abs (sumhr (M, N, f)) @= #0";
+by (cut_inst_tac [("x","M"),("y","N")] hypnat_linear 1);
+by (auto_tac (claset(), simpset() addsimps [inf_close_refl]));
+by (dtac (inf_close_sym RS (inf_close_minus_iff RS iffD1)) 1);
+by (auto_tac (claset() addDs [inf_close_hrabs],
+              simpset() addsimps [sumhr_split_add_minus]));
+qed "sumhr_hrabs_inf_close";
+Addsimps [sumhr_hrabs_inf_close];
+
+(*----------------------------------------------------------------
+      infinite sums: Standard and NS theorems
+ ----------------------------------------------------------------*)
+Goalw [sums_def,NSsums_def] "(f sums l) = (f NSsums l)";
+by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff]) 1);
+qed "sums_NSsums_iff";
+
+Goalw [summable_def,NSsummable_def] 
+      "(summable f) = (NSsummable f)";
+by (simp_tac (simpset() addsimps [sums_NSsums_iff]) 1);
+qed "summable_NSsummable_iff";
+
+Goalw [suminf_def,NSsuminf_def] 
+      "(suminf f) = (NSsuminf f)";
+by (simp_tac (simpset() addsimps [sums_NSsums_iff]) 1);
+qed "suminf_NSsuminf_iff";
+
+Goalw [NSsums_def,NSsummable_def] 
+      "f NSsums l ==> NSsummable f";
+by (Blast_tac 1);
+qed "NSsums_NSsummable";
+
+Goalw [NSsummable_def,NSsuminf_def]
+     "NSsummable f ==> f NSsums (NSsuminf f)";
+by (blast_tac (claset() addIs [someI2]) 1);
+qed "NSsummable_NSsums";
+
+Goal "f NSsums s ==> (s = NSsuminf f)";
+by (asm_full_simp_tac 
+    (simpset() addsimps [suminf_NSsuminf_iff RS sym,sums_NSsums_iff,
+                         sums_unique]) 1);
+qed "NSsums_unique";
+
+Goal "ALL m. n <= Suc m --> f(m) = #0 ==> f NSsums (sumr 0 n f)";
+by (asm_simp_tac (simpset() addsimps [sums_NSsums_iff RS sym, series_zero]) 1);
+qed "NSseries_zero";
+
+Goal "NSsummable f = \
+\     (ALL M: HNatInfinite. ALL N: HNatInfinite. abs (sumhr(M,N,f)) @= #0)";
+by (auto_tac (claset(),
+              simpset() addsimps [summable_NSsummable_iff RS sym,
+                 summable_convergent_sumr_iff, convergent_NSconvergent_iff,
+                 NSCauchy_NSconvergent_iff RS sym, NSCauchy_def,
+                 starfunNat_sumr]));
+by (cut_inst_tac [("x","M"),("y","N")] hypnat_linear 1);
+by (auto_tac (claset(), simpset() addsimps [inf_close_refl]));
+by (rtac ((inf_close_minus_iff RS iffD2) RS inf_close_sym) 1);
+by (rtac (inf_close_minus_iff RS iffD2) 2);
+by (auto_tac (claset() addDs [inf_close_hrabs_zero_cancel],
+              simpset() addsimps [sumhr_split_add_minus]));
+qed "NSsummable_NSCauchy";
+
+(*-------------------------------------------------------------------
+         Terms of a convergent series tend to zero
+ -------------------------------------------------------------------*)
+Goalw [NSLIMSEQ_def] "NSsummable f ==> f ----NS> #0";
+by (auto_tac (claset(), simpset() addsimps [NSsummable_NSCauchy]));
+by (dtac bspec 1 THEN Auto_tac);
+by (dres_inst_tac [("x","N + 1hn")] bspec 1);
+by (auto_tac (claset() addIs [HNatInfinite_add_one,
+                              inf_close_hrabs_zero_cancel],
+              simpset() addsimps [rename_numerals hypreal_of_real_zero]));
+qed "NSsummable_NSLIMSEQ_zero";
+
+(* Easy to prove stsandard case now *)
+Goal "summable f ==> f ----> #0";
+by (auto_tac (claset(),
+        simpset() addsimps [summable_NSsummable_iff,
+                            LIMSEQ_NSLIMSEQ_iff, NSsummable_NSLIMSEQ_zero]));
+qed "summable_LIMSEQ_zero";
+
+(*-------------------------------------------------------------------
+                  NS Comparison test
+ -------------------------------------------------------------------*)
+
+Goal "[| EX N. ALL n. N <= n --> abs(f n) <= g n; \
+\        NSsummable g \
+\     |] ==> NSsummable f";
+by (auto_tac (claset() addIs [summable_comparison_test],
+              simpset() addsimps [summable_NSsummable_iff RS sym]));
+qed "NSsummable_comparison_test";
+
+Goal "[| EX N. ALL n. N <= n --> abs(f n) <= g n; \
+\        NSsummable g \
+\     |] ==> NSsummable (%k. abs (f k))";
+by (rtac NSsummable_comparison_test 1);
+by (auto_tac (claset(), simpset() addsimps [abs_idempotent]));
+qed "NSsummable_rabs_comparison_test";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HSeries.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,30 @@
+(*  Title       : HSeries.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Finite summation and infinite series
+                  for hyperreals
+*) 
+
+HSeries = Series +
+
+consts 
+   sumhr :: "(hypnat * hypnat * (nat=>real)) => hypreal"
+
+defs
+   sumhr_def
+   "sumhr p
+       == Abs_hypreal(UN X:Rep_hypnat(fst p). 
+              UN Y: Rep_hypnat(fst(snd p)).
+              hyprel ^^{%n::nat. sumr (X n) (Y n) (snd(snd p))})"
+
+constdefs
+   NSsums  :: [nat=>real,real] => bool     (infixr 80)
+   "f NSsums s  == (%n. sumr 0 n f) ----NS> s"
+
+   NSsummable :: (nat=>real) => bool
+   "NSsummable f == (EX s. f NSsums s)"
+
+   NSsuminf   :: (nat=>real) => real
+   "NSsuminf f == (@s. f NSsums s)"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HyperArith.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,6 @@
+theory HyperArith = HyperArith0
+files "hypreal_arith.ML":
+
+setup hypreal_arith_setup
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HyperArith0.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,679 @@
+(*  Title:      HOL/Hyperreal/HyperRealArith0.ML
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1999  University of Cambridge
+
+Assorted facts that need binary literals and the arithmetic decision procedure
+
+Also, common factor cancellation
+*)
+
+Goal "((x * y = #0) = (x = #0 | y = (#0::hypreal)))";
+by Auto_tac;  
+by (cut_inst_tac [("x","x"),("y","y")] hypreal_mult_zero_disj 1);
+by Auto_tac;  
+qed "hypreal_mult_is_0";
+AddIffs [hypreal_mult_is_0];
+
+(** Division and inverse **)
+
+Goal "#0/x = (#0::hypreal)";
+by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); 
+qed "hypreal_0_divide";
+Addsimps [hypreal_0_divide];
+
+Goal "((#0::hypreal) < inverse x) = (#0 < x)";
+by (case_tac "x=#0" 1);
+by (asm_simp_tac (HOL_ss addsimps [rename_numerals HYPREAL_INVERSE_ZERO]) 1); 
+by (auto_tac (claset() addDs [hypreal_inverse_less_0], 
+              simpset() addsimps [linorder_neq_iff, 
+                                  hypreal_inverse_gt_0]));  
+qed "hypreal_0_less_inverse_iff";
+Addsimps [hypreal_0_less_inverse_iff];
+
+Goal "(inverse x < (#0::hypreal)) = (x < #0)";
+by (case_tac "x=#0" 1);
+by (asm_simp_tac (HOL_ss addsimps [rename_numerals HYPREAL_INVERSE_ZERO]) 1); 
+by (auto_tac (claset() addDs [hypreal_inverse_less_0], 
+              simpset() addsimps [linorder_neq_iff, 
+                                  hypreal_inverse_gt_0]));  
+qed "hypreal_inverse_less_0_iff";
+Addsimps [hypreal_inverse_less_0_iff];
+
+Goal "((#0::hypreal) <= inverse x) = (#0 <= x)";
+by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); 
+qed "hypreal_0_le_inverse_iff";
+Addsimps [hypreal_0_le_inverse_iff];
+
+Goal "(inverse x <= (#0::hypreal)) = (x <= #0)";
+by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); 
+qed "hypreal_inverse_le_0_iff";
+Addsimps [hypreal_inverse_le_0_iff];
+
+Goalw [hypreal_divide_def] "x/(#0::hypreal) = #0";
+by (stac (rename_numerals HYPREAL_INVERSE_ZERO) 1); 
+by (Simp_tac 1); 
+qed "HYPREAL_DIVIDE_ZERO";
+
+Goal "inverse (x::hypreal) = #1/x";
+by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); 
+qed "hypreal_inverse_eq_divide";
+
+Goal "((#0::hypreal) < x/y) = (#0 < x & #0 < y | x < #0 & y < #0)";
+by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_0_less_mult_iff]) 1);
+qed "hypreal_0_less_divide_iff";
+Addsimps [inst "x" "number_of ?w" hypreal_0_less_divide_iff];
+
+Goal "(x/y < (#0::hypreal)) = (#0 < x & y < #0 | x < #0 & #0 < y)";
+by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_less_0_iff]) 1);
+qed "hypreal_divide_less_0_iff";
+Addsimps [inst "x" "number_of ?w" hypreal_divide_less_0_iff];
+
+Goal "((#0::hypreal) <= x/y) = ((x <= #0 | #0 <= y) & (#0 <= x | y <= #0))";
+by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_0_le_mult_iff]) 1);
+by Auto_tac;  
+qed "hypreal_0_le_divide_iff";
+Addsimps [inst "x" "number_of ?w" hypreal_0_le_divide_iff];
+
+Goal "(x/y <= (#0::hypreal)) = ((x <= #0 | y <= #0) & (#0 <= x | #0 <= y))";
+by (simp_tac (simpset() addsimps [hypreal_divide_def, 
+                                  hypreal_mult_le_0_iff]) 1);
+by Auto_tac;  
+qed "hypreal_divide_le_0_iff";
+Addsimps [inst "x" "number_of ?w" hypreal_divide_le_0_iff];
+
+Goal "(inverse(x::hypreal) = #0) = (x = #0)";
+by (auto_tac (claset(), 
+              simpset() addsimps [rename_numerals HYPREAL_INVERSE_ZERO]));  
+by (rtac ccontr 1); 
+by (blast_tac (claset() addDs [rename_numerals hypreal_inverse_not_zero]) 1); 
+qed "hypreal_inverse_zero_iff";
+Addsimps [hypreal_inverse_zero_iff];
+
+Goal "(x/y = #0) = (x=#0 | y=(#0::hypreal))";
+by (auto_tac (claset(), simpset() addsimps [hypreal_divide_def]));  
+qed "hypreal_divide_eq_0_iff";
+Addsimps [hypreal_divide_eq_0_iff];
+
+Goal "h ~= (#0::hypreal) ==> h/h = #1";
+by (asm_simp_tac 
+    (simpset() addsimps [hypreal_divide_def, hypreal_mult_inverse_left]) 1);
+qed "hypreal_divide_self_eq"; 
+Addsimps [hypreal_divide_self_eq];
+
+
+(**** Factor cancellation theorems for "hypreal" ****)
+
+(** Cancellation laws for k*m < k*n and m*k < n*k, also for <= and =,
+    but not (yet?) for k*m < n*k. **)
+
+bind_thm ("hypreal_mult_minus_right", hypreal_minus_mult_eq2 RS sym);
+
+Goal "(-y < -x) = ((x::hypreal) < y)";
+by (arith_tac 1);
+qed "hypreal_minus_less_minus";
+Addsimps [hypreal_minus_less_minus];
+
+Goal "[| i<j;  k < (0::hypreal) |] ==> j*k < i*k";
+by (rtac (hypreal_minus_less_minus RS iffD1) 1);
+by (auto_tac (claset(), 
+              simpset() delsimps [hypreal_minus_mult_eq2 RS sym]
+                        addsimps [hypreal_minus_mult_eq2,
+                                  hypreal_mult_less_mono1])); 
+qed "hypreal_mult_less_mono1_neg";
+
+Goal "[| i<j;  k < (0::hypreal) |] ==> k*j < k*i";
+by (rtac (hypreal_minus_less_minus RS iffD1) 1);
+by (auto_tac (claset(), 
+              simpset() delsimps [hypreal_minus_mult_eq1 RS sym]
+                        addsimps [hypreal_minus_mult_eq1,
+                                  hypreal_mult_less_mono2]));
+qed "hypreal_mult_less_mono2_neg";
+
+Goal "[| i <= j;  (0::hypreal) <= k |] ==> i*k <= j*k";
+by (auto_tac (claset(), 
+              simpset() addsimps [order_le_less, hypreal_mult_less_mono1]));  
+qed "hypreal_mult_le_mono1";
+
+Goal "[| i <= j;  k <= (0::hypreal) |] ==> j*k <= i*k";
+by (auto_tac (claset(), 
+          simpset() addsimps [order_le_less, hypreal_mult_less_mono1_neg]));  
+qed "hypreal_mult_le_mono1_neg";
+
+Goal "[| i <= j;  (0::hypreal) <= k |] ==> k*i <= k*j";
+by (dtac hypreal_mult_le_mono1 1);
+by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [hypreal_mult_commute])));
+qed "hypreal_mult_le_mono2";
+
+Goal "[| i <= j;  k <= (0::hypreal) |] ==> k*j <= k*i";
+by (dtac hypreal_mult_le_mono1_neg 1);
+by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [hypreal_mult_commute])));
+qed "hypreal_mult_le_mono2_neg";
+
+Goal "[| u <= v;  x <= y;  0 <= v;  (0::hypreal) <= x |] ==> u * x <= v * y";
+by (etac (hypreal_mult_le_mono1 RS order_trans) 1);
+by (assume_tac 1);
+by (etac hypreal_mult_le_mono2 1);
+by (assume_tac 1);
+qed "hypreal_mult_le_mono";
+
+Goal "(m*k < n*k) = (((#0::hypreal) < k & m<n) | (k < #0 & n<m))";
+by (case_tac "k = (0::hypreal)" 1);
+by (auto_tac (claset(), 
+          simpset() addsimps [linorder_neq_iff, 
+                      hypreal_mult_less_mono1, hypreal_mult_less_mono1_neg]));  
+by (auto_tac (claset(), 
+              simpset() addsimps [linorder_not_less,
+				  inst "y1" "m*k" (linorder_not_le RS sym),
+                                  inst "y1" "m" (linorder_not_le RS sym)]));
+by (TRYALL (etac notE));
+by (auto_tac (claset(), 
+              simpset() addsimps [order_less_imp_le, hypreal_mult_le_mono1,
+                                  hypreal_mult_le_mono1_neg]));  
+qed "hypreal_mult_less_cancel2";
+
+Goal "(m*k <= n*k) = (((#0::hypreal) < k --> m<=n) & (k < #0 --> n<=m))";
+by (simp_tac (simpset() addsimps [linorder_not_less RS sym, 
+                                  hypreal_mult_less_cancel2]) 1);
+qed "hypreal_mult_le_cancel2";
+
+Goal "(k*m < k*n) = (((#0::hypreal) < k & m<n) | (k < #0 & n<m))";
+by (simp_tac (simpset() addsimps [inst "z" "k" hypreal_mult_commute, 
+                                  hypreal_mult_less_cancel2]) 1);
+qed "hypreal_mult_less_cancel1";
+
+Goal "!!k::hypreal. (k*m <= k*n) = ((#0 < k --> m<=n) & (k < #0 --> n<=m))";
+by (simp_tac (simpset() addsimps [linorder_not_less RS sym, 
+                                  hypreal_mult_less_cancel1]) 1);
+qed "hypreal_mult_le_cancel1";
+
+Goal "!!k::hypreal. (k*m = k*n) = (k = #0 | m=n)";
+by (case_tac "k=0" 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_mult_left_cancel]));  
+qed "hypreal_mult_eq_cancel1";
+
+Goal "!!k::hypreal. (m*k = n*k) = (k = #0 | m=n)";
+by (case_tac "k=0" 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_mult_right_cancel]));  
+qed "hypreal_mult_eq_cancel2";
+
+Goal "!!k::hypreal. k~=#0 ==> (k*m) / (k*n) = (m/n)";
+by (asm_simp_tac
+    (simpset() addsimps [hypreal_divide_def, hypreal_inverse_distrib]) 1); 
+by (subgoal_tac "k * m * (inverse k * inverse n) = \
+\                (k * inverse k) * (m * inverse n)" 1);
+by (asm_full_simp_tac (simpset() addsimps []) 1); 
+by (asm_full_simp_tac (HOL_ss addsimps hypreal_mult_ac) 1); 
+qed "hypreal_mult_div_cancel1";
+
+(*For ExtractCommonTerm*)
+Goal "(k*m) / (k*n) = (if k = (#0::hypreal) then #0 else m/n)";
+by (simp_tac (simpset() addsimps [hypreal_mult_div_cancel1]) 1); 
+qed "hypreal_mult_div_cancel_disj";
+
+
+local
+  open Hyperreal_Numeral_Simprocs
+in
+
+val rel_hypreal_number_of = [eq_hypreal_number_of, less_hypreal_number_of, 
+                          le_hypreal_number_of_eq_not_less];
+
+structure CancelNumeralFactorCommon =
+  struct
+  val mk_coeff		= mk_coeff
+  val dest_coeff	= dest_coeff 1
+  val trans_tac         = trans_tac
+  val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_plus_1s))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@hypreal_mult_minus_simps))
+     THEN ALLGOALS
+	  (simp_tac 
+	   (HOL_ss addsimps [eq_hypreal_number_of, mult_hypreal_number_of, 
+                             hypreal_mult_number_of_left]@
+           hypreal_minus_from_mult_simps @ hypreal_mult_ac))
+  val numeral_simp_tac	= 
+         ALLGOALS (simp_tac (HOL_ss addsimps rel_hypreal_number_of@bin_simps))
+  val simplify_meta_eq  = simplify_meta_eq
+  end
+
+structure DivCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = prove_conv "hyprealdiv_cancel_numeral_factor"
+  val mk_bal   = HOLogic.mk_binop "HOL.divide"
+  val dest_bal = HOLogic.dest_bin "HOL.divide" hyprealT
+  val cancel = hypreal_mult_div_cancel1 RS trans
+  val neg_exchanges = false
+)
+
+structure EqCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = prove_conv "hyprealeq_cancel_numeral_factor"
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" hyprealT
+  val cancel = hypreal_mult_eq_cancel1 RS trans
+  val neg_exchanges = false
+)
+
+structure LessCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = prove_conv "hyprealless_cancel_numeral_factor"
+  val mk_bal   = HOLogic.mk_binrel "op <"
+  val dest_bal = HOLogic.dest_bin "op <" hyprealT
+  val cancel = hypreal_mult_less_cancel1 RS trans
+  val neg_exchanges = true
+)
+
+structure LeCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = prove_conv "hyprealle_cancel_numeral_factor"
+  val mk_bal   = HOLogic.mk_binrel "op <="
+  val dest_bal = HOLogic.dest_bin "op <=" hyprealT
+  val cancel = hypreal_mult_le_cancel1 RS trans
+  val neg_exchanges = true
+)
+
+val hypreal_cancel_numeral_factors_relations = 
+  map prep_simproc
+   [("hyprealeq_cancel_numeral_factor",
+     prep_pats ["(l::hypreal) * m = n", "(l::hypreal) = m * n"], 
+     EqCancelNumeralFactor.proc),
+    ("hyprealless_cancel_numeral_factor", 
+     prep_pats ["(l::hypreal) * m < n", "(l::hypreal) < m * n"], 
+     LessCancelNumeralFactor.proc),
+    ("hyprealle_cancel_numeral_factor", 
+     prep_pats ["(l::hypreal) * m <= n", "(l::hypreal) <= m * n"], 
+     LeCancelNumeralFactor.proc)];
+
+val hypreal_cancel_numeral_factors_divide = prep_simproc
+	("hyprealdiv_cancel_numeral_factor", 
+	 prep_pats ["((l::hypreal) * m) / n", "(l::hypreal) / (m * n)"], 
+	 DivCancelNumeralFactor.proc);
+
+val hypreal_cancel_numeral_factors = 
+    hypreal_cancel_numeral_factors_relations @ 
+    [hypreal_cancel_numeral_factors_divide];
+
+end;
+
+Addsimprocs hypreal_cancel_numeral_factors;
+
+
+(*examples:
+print_depth 22;
+set timing;
+set trace_simp;
+fun test s = (Goal s; by (Simp_tac 1)); 
+
+test "#0 <= (y::hypreal) * #-2";
+test "#9*x = #12 * (y::hypreal)";
+test "(#9*x) / (#12 * (y::hypreal)) = z";
+test "#9*x < #12 * (y::hypreal)";
+test "#9*x <= #12 * (y::hypreal)";
+
+test "#-99*x = #132 * (y::hypreal)";
+test "(#-99*x) / (#132 * (y::hypreal)) = z";
+test "#-99*x < #132 * (y::hypreal)";
+test "#-99*x <= #132 * (y::hypreal)";
+
+test "#999*x = #-396 * (y::hypreal)";
+test "(#999*x) / (#-396 * (y::hypreal)) = z";
+test "#999*x < #-396 * (y::hypreal)";
+test "#999*x <= #-396 * (y::hypreal)";
+
+test "#-99*x = #-81 * (y::hypreal)";
+test "(#-99*x) / (#-81 * (y::hypreal)) = z";
+test "#-99*x <= #-81 * (y::hypreal)";
+test "#-99*x < #-81 * (y::hypreal)";
+
+test "#-2 * x = #-1 * (y::hypreal)";
+test "#-2 * x = -(y::hypreal)";
+test "(#-2 * x) / (#-1 * (y::hypreal)) = z";
+test "#-2 * x < -(y::hypreal)";
+test "#-2 * x <= #-1 * (y::hypreal)";
+test "-x < #-23 * (y::hypreal)";
+test "-x <= #-23 * (y::hypreal)";
+*)
+
+
+(** Declarations for ExtractCommonTerm **)
+
+local
+  open Hyperreal_Numeral_Simprocs
+in
+
+structure CancelFactorCommon =
+  struct
+  val mk_sum    	= long_mk_prod
+  val dest_sum		= dest_prod
+  val mk_coeff		= mk_coeff
+  val dest_coeff	= dest_coeff
+  val find_first	= find_first []
+  val trans_tac         = trans_tac
+  val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@hypreal_mult_ac))
+  end;
+
+structure EqCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = prove_conv "hypreal_eq_cancel_factor"
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" hyprealT
+  val simplify_meta_eq  = cancel_simplify_meta_eq hypreal_mult_eq_cancel1
+);
+
+
+structure DivideCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = prove_conv "hypreal_divide_cancel_factor"
+  val mk_bal   = HOLogic.mk_binop "HOL.divide"
+  val dest_bal = HOLogic.dest_bin "HOL.divide" hyprealT
+  val simplify_meta_eq  = cancel_simplify_meta_eq hypreal_mult_div_cancel_disj
+);
+
+val hypreal_cancel_factor = 
+  map prep_simproc
+   [("hypreal_eq_cancel_factor",
+     prep_pats ["(l::hypreal) * m = n", "(l::hypreal) = m * n"], 
+     EqCancelFactor.proc),
+    ("hypreal_divide_cancel_factor", 
+     prep_pats ["((l::hypreal) * m) / n", "(l::hypreal) / (m * n)"], 
+     DivideCancelFactor.proc)];
+
+end;
+
+Addsimprocs hypreal_cancel_factor;
+
+
+(*examples:
+print_depth 22;
+set timing;
+set trace_simp;
+fun test s = (Goal s; by (Asm_simp_tac 1)); 
+
+test "x*k = k*(y::hypreal)";
+test "k = k*(y::hypreal)"; 
+test "a*(b*c) = (b::hypreal)";
+test "a*(b*c) = d*(b::hypreal)*(x*a)";
+
+
+test "(x*k) / (k*(y::hypreal)) = (uu::hypreal)";
+test "(k) / (k*(y::hypreal)) = (uu::hypreal)"; 
+test "(a*(b*c)) / ((b::hypreal)) = (uu::hypreal)";
+test "(a*(b*c)) / (d*(b::hypreal)*(x*a)) = (uu::hypreal)";
+
+(*FIXME: what do we do about this?*)
+test "a*(b*c)/(y*z) = d*(b::hypreal)*(x*a)/z";
+*)
+
+
+(*** Simplification of inequalities involving literal divisors ***)
+
+Goal "#0<z ==> ((x::hypreal) <= y/z) = (x*z <= y)";
+by (subgoal_tac "(x*z <= y) = (x*z <= (y/z)*z)" 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac hypreal_mult_le_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "pos_hypreal_le_divide_eq";
+Addsimps [inst "z" "number_of ?w" pos_hypreal_le_divide_eq];
+
+Goal "z<#0 ==> ((x::hypreal) <= y/z) = (y <= x*z)";
+by (subgoal_tac "(y <= x*z) = ((y/z)*z <= x*z)" 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac hypreal_mult_le_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "neg_hypreal_le_divide_eq";
+Addsimps [inst "z" "number_of ?w" neg_hypreal_le_divide_eq];
+
+Goal "#0<z ==> (y/z <= (x::hypreal)) = (y <= x*z)";
+by (subgoal_tac "(y <= x*z) = ((y/z)*z <= x*z)" 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac hypreal_mult_le_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "pos_hypreal_divide_le_eq";
+Addsimps [inst "z" "number_of ?w" pos_hypreal_divide_le_eq];
+
+Goal "z<#0 ==> (y/z <= (x::hypreal)) = (x*z <= y)";
+by (subgoal_tac "(x*z <= y) = (x*z <= (y/z)*z)" 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac hypreal_mult_le_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "neg_hypreal_divide_le_eq";
+Addsimps [inst "z" "number_of ?w" neg_hypreal_divide_le_eq];
+
+Goal "#0<z ==> ((x::hypreal) < y/z) = (x*z < y)";
+by (subgoal_tac "(x*z < y) = (x*z < (y/z)*z)" 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac hypreal_mult_less_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "pos_hypreal_less_divide_eq";
+Addsimps [inst "z" "number_of ?w" pos_hypreal_less_divide_eq];
+
+Goal "z<#0 ==> ((x::hypreal) < y/z) = (y < x*z)";
+by (subgoal_tac "(y < x*z) = ((y/z)*z < x*z)" 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac hypreal_mult_less_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "neg_hypreal_less_divide_eq";
+Addsimps [inst "z" "number_of ?w" neg_hypreal_less_divide_eq];
+
+Goal "#0<z ==> (y/z < (x::hypreal)) = (y < x*z)";
+by (subgoal_tac "(y < x*z) = ((y/z)*z < x*z)" 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac hypreal_mult_less_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "pos_hypreal_divide_less_eq";
+Addsimps [inst "z" "number_of ?w" pos_hypreal_divide_less_eq];
+
+Goal "z<#0 ==> (y/z < (x::hypreal)) = (x*z < y)";
+by (subgoal_tac "(x*z < y) = (x*z < (y/z)*z)" 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac hypreal_mult_less_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "neg_hypreal_divide_less_eq";
+Addsimps [inst "z" "number_of ?w" neg_hypreal_divide_less_eq];
+
+Goal "z~=#0 ==> ((x::hypreal) = y/z) = (x*z = y)";
+by (subgoal_tac "(x*z = y) = (x*z = (y/z)*z)" 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac hypreal_mult_eq_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "hypreal_eq_divide_eq";
+Addsimps [inst "z" "number_of ?w" hypreal_eq_divide_eq];
+
+Goal "z~=#0 ==> (y/z = (x::hypreal)) = (y = x*z)";
+by (subgoal_tac "(y = x*z) = ((y/z)*z = x*z)" 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac hypreal_mult_eq_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "hypreal_divide_eq_eq";
+Addsimps [inst "z" "number_of ?w" hypreal_divide_eq_eq];
+
+Goal "(m/k = n/k) = (k = #0 | m = (n::hypreal))";
+by (case_tac "k=#0" 1);
+by (asm_simp_tac (simpset() addsimps [HYPREAL_DIVIDE_ZERO]) 1); 
+by (asm_simp_tac (simpset() addsimps [hypreal_divide_eq_eq, hypreal_eq_divide_eq, 
+                                      hypreal_mult_eq_cancel2]) 1); 
+qed "hypreal_divide_eq_cancel2";
+
+Goal "(k/m = k/n) = (k = #0 | m = (n::hypreal))";
+by (case_tac "m=#0 | n = #0" 1);
+by (auto_tac (claset(), 
+              simpset() addsimps [HYPREAL_DIVIDE_ZERO, hypreal_divide_eq_eq, 
+                                  hypreal_eq_divide_eq, hypreal_mult_eq_cancel1]));  
+qed "hypreal_divide_eq_cancel1";
+
+Goal "[| #0 < r; #0 < x|] ==> (inverse x < inverse (r::hypreal)) = (r < x)";
+by (auto_tac (claset() addIs [hypreal_inverse_less_swap], simpset()));
+by (res_inst_tac [("t","r")] (hypreal_inverse_inverse RS subst) 1);
+by (res_inst_tac [("t","x")] (hypreal_inverse_inverse RS subst) 1);
+by (auto_tac (claset() addIs [hypreal_inverse_less_swap],
+	      simpset() delsimps [hypreal_inverse_inverse]
+			addsimps [hypreal_inverse_gt_zero]));
+qed "hypreal_inverse_less_iff";
+
+Goal "[| #0 < r; #0 < x|] ==> (inverse x <= inverse r) = (r <= (x::hypreal))";
+by (asm_simp_tac (simpset() addsimps [linorder_not_less RS sym, 
+                                      hypreal_inverse_less_iff]) 1); 
+qed "hypreal_inverse_le_iff";
+
+(** Division by 1, -1 **)
+
+Goal "(x::hypreal)/#1 = x";
+by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); 
+qed "hypreal_divide_1";
+Addsimps [hypreal_divide_1];
+
+Goal "x/#-1 = -(x::hypreal)";
+by (Simp_tac 1); 
+qed "hypreal_divide_minus1";
+Addsimps [hypreal_divide_minus1];
+
+Goal "#-1/(x::hypreal) = - (#1/x)";
+by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_minus_inverse]) 1); 
+qed "hypreal_minus1_divide";
+Addsimps [hypreal_minus1_divide];
+
+Goal "[| (#0::hypreal) < d1; #0 < d2 |] ==> EX e. #0 < e & e < d1 & e < d2";
+by (res_inst_tac [("x","(min d1 d2)/#2")] exI 1); 
+by (asm_simp_tac (simpset() addsimps [min_def]) 1); 
+qed "hypreal_lbound_gt_zero";
+
+
+(*** General rewrites to improve automation, like those for type "int" ***)
+
+(** The next several equations can make the simplifier loop! **)
+
+Goal "(x < - y) = (y < - (x::hypreal))";
+by Auto_tac;  
+qed "hypreal_less_minus"; 
+
+Goal "(- x < y) = (- y < (x::hypreal))";
+by Auto_tac;  
+qed "hypreal_minus_less"; 
+
+Goal "(x <= - y) = (y <= - (x::hypreal))";
+by Auto_tac;  
+qed "hypreal_le_minus"; 
+
+Goal "(- x <= y) = (- y <= (x::hypreal))";
+by Auto_tac;  
+qed "hypreal_minus_le"; 
+
+Goal "(x = - y) = (y = - (x::hypreal))";
+by Auto_tac;
+qed "hypreal_equation_minus";
+
+Goal "(- x = y) = (- (y::hypreal) = x)";
+by Auto_tac;
+qed "hypreal_minus_equation";
+
+Goal "(x + - a = (#0::hypreal)) = (x=a)";
+by (arith_tac 1);
+qed "hypreal_add_minus_iff";
+Addsimps [hypreal_add_minus_iff];
+
+Goal "(-b = -a) = (b = (a::hypreal))";
+by (arith_tac 1);
+qed "hypreal_minus_eq_cancel";
+Addsimps [hypreal_minus_eq_cancel];
+
+Goal "(-s <= -r) = ((r::hypreal) <= s)";
+by (stac hypreal_minus_le 1); 
+by (Simp_tac 1); 
+qed "hypreal_le_minus_iff";
+Addsimps [hypreal_le_minus_iff];          
+
+
+(*Distributive laws for literals*)
+Addsimps (map (inst "w" "number_of ?v")
+	  [hypreal_add_mult_distrib, hypreal_add_mult_distrib2,
+	   hypreal_diff_mult_distrib, hypreal_diff_mult_distrib2]);
+
+Addsimps (map (inst "x" "number_of ?v") 
+	  [hypreal_less_minus, hypreal_le_minus, hypreal_equation_minus]);
+Addsimps (map (inst "y" "number_of ?v") 
+	  [hypreal_minus_less, hypreal_minus_le, hypreal_minus_equation]);
+
+
+(*** Simprules combining x+y and #0 ***)
+
+Goal "(x+y = (#0::hypreal)) = (y = -x)";
+by Auto_tac;  
+qed "hypreal_add_eq_0_iff";
+AddIffs [hypreal_add_eq_0_iff];
+
+Goal "(x+y < (#0::hypreal)) = (y < -x)";
+by Auto_tac;  
+qed "hypreal_add_less_0_iff";
+AddIffs [hypreal_add_less_0_iff];
+
+Goal "((#0::hypreal) < x+y) = (-x < y)";
+by Auto_tac;  
+qed "hypreal_0_less_add_iff";
+AddIffs [hypreal_0_less_add_iff];
+
+Goal "(x+y <= (#0::hypreal)) = (y <= -x)";
+by Auto_tac;  
+qed "hypreal_add_le_0_iff";
+AddIffs [hypreal_add_le_0_iff];
+
+Goal "((#0::hypreal) <= x+y) = (-x <= y)";
+by Auto_tac;  
+qed "hypreal_0_le_add_iff";
+AddIffs [hypreal_0_le_add_iff];
+
+
+(** Simprules combining x-y and #0; see also hypreal_less_iff_diff_less_0 etc
+    in HyperBin
+**)
+
+Goal "((#0::hypreal) < x-y) = (y < x)";
+by Auto_tac;  
+qed "hypreal_0_less_diff_iff";
+AddIffs [hypreal_0_less_diff_iff];
+
+Goal "((#0::hypreal) <= x-y) = (y <= x)";
+by Auto_tac;  
+qed "hypreal_0_le_diff_iff";
+AddIffs [hypreal_0_le_diff_iff];
+
+(*
+FIXME: we should have this, as for type int, but many proofs would break.
+It replaces x+-y by x-y.
+Addsimps [symmetric hypreal_diff_def];
+*)
+
+Goal "-(x-y) = y - (x::hypreal)";
+by (arith_tac 1);
+qed "hypreal_minus_diff_eq";
+Addsimps [hypreal_minus_diff_eq];
+
+
+(*** Density of the Hyperreals ***)
+
+Goal "x < y ==> x < (x+y) / (#2::hypreal)";
+by Auto_tac;
+qed "hypreal_less_half_sum";
+
+Goal "x < y ==> (x+y)/(#2::hypreal) < y";
+by Auto_tac;
+qed "hypreal_gt_half_sum";
+
+Goal "x < y ==> EX r::hypreal. x < r & r < y";
+by (blast_tac (claset() addSIs [hypreal_less_half_sum, hypreal_gt_half_sum]) 1);
+qed "hypreal_dense";
+
+
+(*Replaces "inverse #nn" by #1/#nn *)
+Addsimps [inst "x" "number_of ?w" hypreal_inverse_eq_divide];
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HyperArith0.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,6 @@
+theory HyperArith0 = HyperBin
+files "hypreal_arith0.ML":
+
+setup hypreal_arith_setup
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HyperBin.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,626 @@
+(*  Title:      HOL/Hyperreal/HyperBin.ML
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   2000  University of Cambridge
+
+Binary arithmetic for the hypreals (integer literals only).
+*)
+
+(** hypreal_of_real (coercion from int to real) **)
+
+Goal "hypreal_of_real (number_of w) = number_of w";
+by (simp_tac (simpset() addsimps [hypreal_number_of_def]) 1);
+qed "hypreal_number_of";
+Addsimps [hypreal_number_of];
+
+Goalw [hypreal_number_of_def] "(0::hypreal) = #0";
+by (simp_tac (simpset() addsimps [hypreal_of_real_zero RS sym]) 1);
+qed "zero_eq_numeral_0";
+
+Goalw [hypreal_number_of_def] "1hr = #1";
+by (simp_tac (simpset() addsimps [hypreal_of_real_one RS sym]) 1);
+qed "one_eq_numeral_1";
+
+(** Addition **)
+
+Goal "(number_of v :: hypreal) + number_of v' = number_of (bin_add v v')";
+by (simp_tac
+    (HOL_ss addsimps [hypreal_number_of_def, 
+                      hypreal_of_real_add RS sym, add_real_number_of]) 1);
+qed "add_hypreal_number_of";
+Addsimps [add_hypreal_number_of];
+
+
+(** Subtraction **)
+
+Goalw [hypreal_number_of_def]
+     "- (number_of w :: hypreal) = number_of (bin_minus w)";
+by (simp_tac
+    (HOL_ss addsimps [minus_real_number_of, hypreal_of_real_minus RS sym]) 1);
+qed "minus_hypreal_number_of";
+Addsimps [minus_hypreal_number_of];
+
+Goalw [hypreal_number_of_def, hypreal_diff_def]
+     "(number_of v :: hypreal) - number_of w = \
+\     number_of (bin_add v (bin_minus w))";
+by (Simp_tac 1); 
+qed "diff_hypreal_number_of";
+Addsimps [diff_hypreal_number_of];
+
+
+(** Multiplication **)
+
+Goal "(number_of v :: hypreal) * number_of v' = number_of (bin_mult v v')";
+by (simp_tac
+    (HOL_ss addsimps [hypreal_number_of_def, 
+	              hypreal_of_real_mult RS sym, mult_real_number_of]) 1);
+qed "mult_hypreal_number_of";
+Addsimps [mult_hypreal_number_of];
+
+Goal "(#2::hypreal) = #1 + #1";
+by (Simp_tac 1);
+val lemma = result();
+
+(*For specialist use: NOT as default simprules*)
+Goal "#2 * z = (z+z::hypreal)";
+by (simp_tac (simpset ()
+	      addsimps [lemma, hypreal_add_mult_distrib,
+			one_eq_numeral_1 RS sym]) 1);
+qed "hypreal_mult_2";
+
+Goal "z * #2 = (z+z::hypreal)";
+by (stac hypreal_mult_commute 1 THEN rtac hypreal_mult_2 1);
+qed "hypreal_mult_2_right";
+
+
+(*** Comparisons ***)
+
+(** Equals (=) **)
+
+Goal "((number_of v :: hypreal) = number_of v') = \
+\     iszero (number_of (bin_add v (bin_minus v')))";
+by (simp_tac
+    (HOL_ss addsimps [hypreal_number_of_def, 
+	              hypreal_of_real_eq_iff, eq_real_number_of]) 1);
+qed "eq_hypreal_number_of";
+Addsimps [eq_hypreal_number_of];
+
+(** Less-than (<) **)
+
+(*"neg" is used in rewrite rules for binary comparisons*)
+Goal "((number_of v :: hypreal) < number_of v') = \
+\     neg (number_of (bin_add v (bin_minus v')))";
+by (simp_tac
+    (HOL_ss addsimps [hypreal_number_of_def, hypreal_of_real_less_iff, 
+		      less_real_number_of]) 1);
+qed "less_hypreal_number_of";
+Addsimps [less_hypreal_number_of];
+
+
+(** Less-than-or-equals (<=) **)
+
+Goal "(number_of x <= (number_of y::hypreal)) = \
+\     (~ number_of y < (number_of x::hypreal))";
+by (rtac (linorder_not_less RS sym) 1);
+qed "le_hypreal_number_of_eq_not_less"; 
+Addsimps [le_hypreal_number_of_eq_not_less];
+
+(*** New versions of existing theorems involving 0, 1hr ***)
+
+Goal "- #1 = (#-1::hypreal)";
+by (Simp_tac 1);
+qed "minus_numeral_one";
+
+(*Maps 0 to #0 and 1hr to #1 and -1hr to #-1*)
+val hypreal_numeral_ss = 
+    real_numeral_ss addsimps [zero_eq_numeral_0, one_eq_numeral_1, 
+		              minus_numeral_one];
+
+fun rename_numerals th = 
+    asm_full_simplify hypreal_numeral_ss (Thm.transfer (the_context ()) th);
+
+
+(*Now insert some identities previously stated for 0 and 1hr*)
+
+(** HyperDef **)
+
+Addsimps (map rename_numerals
+	  [hypreal_minus_zero, hypreal_minus_zero_iff,
+	   hypreal_add_zero_left, hypreal_add_zero_right, 
+	   hypreal_diff_zero, hypreal_diff_zero_right,
+	   hypreal_mult_0_right, hypreal_mult_0, 
+           hypreal_mult_1_right, hypreal_mult_1,
+	   hypreal_inverse_1, hypreal_minus_zero_less_iff]);
+
+bind_thm ("hypreal_0_less_mult_iff", 
+	  rename_numerals hypreal_zero_less_mult_iff);
+bind_thm ("hypreal_0_le_mult_iff", 
+	  rename_numerals hypreal_zero_le_mult_iff);
+bind_thm ("hypreal_mult_less_0_iff", 
+	  rename_numerals hypreal_mult_less_zero_iff);
+bind_thm ("hypreal_mult_le_0_iff", 
+	  rename_numerals hypreal_mult_le_zero_iff);
+
+bind_thm ("hypreal_inverse_less_0", rename_numerals hypreal_inverse_less_zero);
+bind_thm ("hypreal_inverse_gt_0", rename_numerals hypreal_inverse_gt_zero);
+
+Addsimps [zero_eq_numeral_0,one_eq_numeral_1];
+
+
+(** Simplification of arithmetic when nested to the right **)
+
+Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::hypreal)";
+by Auto_tac; 
+qed "hypreal_add_number_of_left";
+
+Goal "number_of v *(number_of w * z) = (number_of(bin_mult v w) * z::hypreal)";
+by (simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym]) 1);
+qed "hypreal_mult_number_of_left";
+
+Goalw [hypreal_diff_def]
+    "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::hypreal)";
+by (rtac hypreal_add_number_of_left 1);
+qed "hypreal_add_number_of_diff1";
+
+Goal "number_of v + (c - number_of w) = \
+\     number_of (bin_add v (bin_minus w)) + (c::hypreal)";
+by (stac (diff_hypreal_number_of RS sym) 1);
+by Auto_tac;
+qed "hypreal_add_number_of_diff2";
+
+Addsimps [hypreal_add_number_of_left, hypreal_mult_number_of_left,
+	  hypreal_add_number_of_diff1, hypreal_add_number_of_diff2]; 
+
+
+(*"neg" is used in rewrite rules for binary comparisons*)
+Goal "hypreal_of_nat (number_of v :: nat) = \
+\        (if neg (number_of v) then #0 \
+\         else (number_of v :: hypreal))";
+by (simp_tac (simpset() addsimps [hypreal_of_nat_real_of_nat]) 1);
+qed "hypreal_of_nat_number_of";
+Addsimps [hypreal_of_nat_number_of];
+
+
+(**** Simprocs for numeric literals ****)
+
+(** Combining of literal coefficients in sums of products **)
+
+Goal "(x < y) = (x-y < (#0::hypreal))";
+by (simp_tac (simpset() addsimps [hypreal_diff_less_eq]) 1);   
+qed "hypreal_less_iff_diff_less_0";
+
+Goal "(x = y) = (x-y = (#0::hypreal))";
+by (simp_tac (simpset() addsimps [hypreal_diff_eq_eq]) 1);   
+qed "hypreal_eq_iff_diff_eq_0";
+
+Goal "(x <= y) = (x-y <= (#0::hypreal))";
+by (simp_tac (simpset() addsimps [hypreal_diff_le_eq]) 1);   
+qed "hypreal_le_iff_diff_le_0";
+
+
+(** For combine_numerals **)
+
+Goal "i*u + (j*u + k) = (i+j)*u + (k::hypreal)";
+by (asm_simp_tac (simpset() addsimps [hypreal_add_mult_distrib]) 1);
+qed "left_hypreal_add_mult_distrib";
+
+
+(** For cancel_numerals **)
+
+val rel_iff_rel_0_rls = 
+    map (inst "y" "?u+?v")
+      [hypreal_less_iff_diff_less_0, hypreal_eq_iff_diff_eq_0, 
+       hypreal_le_iff_diff_le_0] @
+    map (inst "y" "n")
+      [hypreal_less_iff_diff_less_0, hypreal_eq_iff_diff_eq_0, 
+       hypreal_le_iff_diff_le_0];
+
+Goal "!!i::hypreal. (i*u + m = j*u + n) = ((i-j)*u + m = n)";
+by (asm_simp_tac
+    (simpset() addsimps [hypreal_diff_def, hypreal_add_mult_distrib]@
+	                 hypreal_add_ac@rel_iff_rel_0_rls) 1);
+qed "hypreal_eq_add_iff1";
+
+Goal "!!i::hypreal. (i*u + m = j*u + n) = (m = (j-i)*u + n)";
+by (asm_simp_tac
+    (simpset() addsimps [hypreal_diff_def, hypreal_add_mult_distrib]@
+                         hypreal_add_ac@rel_iff_rel_0_rls) 1);
+qed "hypreal_eq_add_iff2";
+
+Goal "!!i::hypreal. (i*u + m < j*u + n) = ((i-j)*u + m < n)";
+by (asm_simp_tac
+    (simpset() addsimps [hypreal_diff_def, hypreal_add_mult_distrib]@
+                         hypreal_add_ac@rel_iff_rel_0_rls) 1);
+qed "hypreal_less_add_iff1";
+
+Goal "!!i::hypreal. (i*u + m < j*u + n) = (m < (j-i)*u + n)";
+by (asm_simp_tac
+    (simpset() addsimps [hypreal_diff_def, hypreal_add_mult_distrib]@
+                         hypreal_add_ac@rel_iff_rel_0_rls) 1);
+qed "hypreal_less_add_iff2";
+
+Goal "!!i::hypreal. (i*u + m <= j*u + n) = ((i-j)*u + m <= n)";
+by (asm_simp_tac
+    (simpset() addsimps [hypreal_diff_def, hypreal_add_mult_distrib]@
+                         hypreal_add_ac@rel_iff_rel_0_rls) 1);
+qed "hypreal_le_add_iff1";
+
+Goal "!!i::hypreal. (i*u + m <= j*u + n) = (m <= (j-i)*u + n)";
+by (asm_simp_tac
+    (simpset() addsimps [hypreal_diff_def, hypreal_add_mult_distrib]@
+                        hypreal_add_ac@rel_iff_rel_0_rls) 1);
+qed "hypreal_le_add_iff2";
+
+Goal "(z::hypreal) * #-1 = -z";
+by (stac (minus_numeral_one RS sym) 1);
+by (stac (hypreal_minus_mult_eq2 RS sym) 1); 
+by Auto_tac;  
+qed "hypreal_mult_minus_1_right";
+Addsimps [hypreal_mult_minus_1_right];
+
+Goal "#-1 * (z::hypreal) = -z";
+by (simp_tac (simpset() addsimps [hypreal_mult_commute]) 1); 
+qed "hypreal_mult_minus_1";
+Addsimps [hypreal_mult_minus_1];
+
+
+
+structure Hyperreal_Numeral_Simprocs =
+struct
+
+(*Utilities*)
+
+
+val hyprealT = Type("HyperDef.hypreal",[]);
+
+fun mk_numeral n = HOLogic.number_of_const hyprealT $ HOLogic.mk_bin n;
+
+val dest_numeral = Real_Numeral_Simprocs.dest_numeral;
+
+val find_first_numeral = Real_Numeral_Simprocs.find_first_numeral;
+
+val zero = mk_numeral 0;
+val mk_plus = HOLogic.mk_binop "op +";
+
+val uminus_const = Const ("uminus", hyprealT --> hyprealT);
+
+(*Thus mk_sum[t] yields t+#0; longer sums don't have a trailing zero*)
+fun mk_sum []        = zero
+  | mk_sum [t,u]     = mk_plus (t, u)
+  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
+
+(*this version ALWAYS includes a trailing zero*)
+fun long_mk_sum []        = zero
+  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
+
+val dest_plus = HOLogic.dest_bin "op +" hyprealT;
+
+(*decompose additions AND subtractions as a sum*)
+fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
+        dest_summing (pos, t, dest_summing (pos, u, ts))
+  | dest_summing (pos, Const ("op -", _) $ t $ u, ts) =
+        dest_summing (pos, t, dest_summing (not pos, u, ts))
+  | dest_summing (pos, t, ts) =
+	if pos then t::ts else uminus_const$t :: ts;
+
+fun dest_sum t = dest_summing (true, t, []);
+
+val mk_diff = HOLogic.mk_binop "op -";
+val dest_diff = HOLogic.dest_bin "op -" hyprealT;
+
+val one = mk_numeral 1;
+val mk_times = HOLogic.mk_binop "op *";
+
+fun mk_prod [] = one
+  | mk_prod [t] = t
+  | mk_prod (t :: ts) = if t = one then mk_prod ts
+                        else mk_times (t, mk_prod ts);
+
+val dest_times = HOLogic.dest_bin "op *" hyprealT;
+
+fun dest_prod t =
+      let val (t,u) = dest_times t 
+      in  dest_prod t @ dest_prod u  end
+      handle TERM _ => [t];
+
+(*DON'T do the obvious simplifications; that would create special cases*) 
+fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts);
+
+(*Express t as a product of (possibly) a numeral with other sorted terms*)
+fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t
+  | dest_coeff sign t =
+    let val ts = sort Term.term_ord (dest_prod t)
+	val (n, ts') = find_first_numeral [] ts
+                          handle TERM _ => (1, ts)
+    in (sign*n, mk_prod ts') end;
+
+(*Find first coefficient-term THAT MATCHES u*)
+fun find_first_coeff past u [] = raise TERM("find_first_coeff", []) 
+  | find_first_coeff past u (t::terms) =
+	let val (n,u') = dest_coeff 1 t
+	in  if u aconv u' then (n, rev past @ terms)
+			  else find_first_coeff (t::past) u terms
+	end
+	handle TERM _ => find_first_coeff (t::past) u terms;
+
+
+(*Simplify #1*n and n*#1 to n*)
+val add_0s = map rename_numerals
+                 [hypreal_add_zero_left, hypreal_add_zero_right];
+val mult_plus_1s = map rename_numerals
+                      [hypreal_mult_1, hypreal_mult_1_right];
+val mult_minus_1s = map rename_numerals
+                      [hypreal_mult_minus_1, hypreal_mult_minus_1_right];
+val mult_1s = mult_plus_1s @ mult_minus_1s;
+
+(*To perform binary arithmetic*)
+val bin_simps =
+    [add_hypreal_number_of, hypreal_add_number_of_left, 
+     minus_hypreal_number_of, diff_hypreal_number_of, mult_hypreal_number_of, 
+     hypreal_mult_number_of_left] @ bin_arith_simps @ bin_rel_simps;
+
+(*To evaluate binary negations of coefficients*)
+val hypreal_minus_simps = NCons_simps @
+                   [minus_hypreal_number_of, 
+		    bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min,
+		    bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min];
+
+(*To let us treat subtraction as addition*)
+val diff_simps = [hypreal_diff_def, hypreal_minus_add_distrib, 
+                  hypreal_minus_minus];
+
+(*push the unary minus down: - x * y = x * - y *)
+val hypreal_minus_mult_eq_1_to_2 = 
+    [hypreal_minus_mult_eq1 RS sym, hypreal_minus_mult_eq2] MRS trans 
+    |> standard;
+
+(*to extract again any uncancelled minuses*)
+val hypreal_minus_from_mult_simps = 
+    [hypreal_minus_minus, hypreal_minus_mult_eq1 RS sym, 
+     hypreal_minus_mult_eq2 RS sym];
+
+(*combine unary minus with numeric literals, however nested within a product*)
+val hypreal_mult_minus_simps =
+    [hypreal_mult_assoc, hypreal_minus_mult_eq1, hypreal_minus_mult_eq_1_to_2];
+
+(*Apply the given rewrite (if present) just once*)
+fun trans_tac None      = all_tac
+  | trans_tac (Some th) = ALLGOALS (rtac (th RS trans));
+
+fun prove_conv name tacs sg (hyps: thm list) (t,u) =
+  if t aconv u then None
+  else
+  let val ct = cterm_of sg (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u)))
+  in Some
+     (prove_goalw_cterm [] ct (K tacs)
+      handle ERROR => error 
+	  ("The error(s) above occurred while trying to prove " ^
+	   string_of_cterm ct ^ "\nInternal failure of simproc " ^ name))
+  end;
+
+(*version without the hyps argument*)
+fun prove_conv_nohyps name tacs sg = prove_conv name tacs sg [];
+
+(*Final simplification: cancel + and *  *)
+val simplify_meta_eq = 
+    Int_Numeral_Simprocs.simplify_meta_eq
+         [hypreal_add_zero_left, hypreal_add_zero_right,
+ 	  hypreal_mult_0, hypreal_mult_0_right, hypreal_mult_1, 
+          hypreal_mult_1_right];
+
+val prep_simproc = Real_Numeral_Simprocs.prep_simproc;
+val prep_pats = map Real_Numeral_Simprocs.prep_pat;
+
+structure CancelNumeralsCommon =
+  struct
+  val mk_sum    	= mk_sum
+  val dest_sum		= dest_sum
+  val mk_coeff		= mk_coeff
+  val dest_coeff	= dest_coeff 1
+  val find_first_coeff	= find_first_coeff []
+  val trans_tac         = trans_tac
+  val norm_tac = 
+     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
+                                         hypreal_minus_simps@hypreal_add_ac))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@hypreal_mult_minus_simps))
+     THEN ALLGOALS
+              (simp_tac (HOL_ss addsimps hypreal_minus_from_mult_simps@
+                                         hypreal_add_ac@hypreal_mult_ac))
+  val numeral_simp_tac	= ALLGOALS (simp_tac (HOL_ss addsimps add_0s@bin_simps))
+  val simplify_meta_eq  = simplify_meta_eq
+  end;
+
+
+structure EqCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = prove_conv "hyprealeq_cancel_numerals"
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" hyprealT
+  val bal_add1 = hypreal_eq_add_iff1 RS trans
+  val bal_add2 = hypreal_eq_add_iff2 RS trans
+);
+
+structure LessCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = prove_conv "hyprealless_cancel_numerals"
+  val mk_bal   = HOLogic.mk_binrel "op <"
+  val dest_bal = HOLogic.dest_bin "op <" hyprealT
+  val bal_add1 = hypreal_less_add_iff1 RS trans
+  val bal_add2 = hypreal_less_add_iff2 RS trans
+);
+
+structure LeCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = prove_conv "hyprealle_cancel_numerals"
+  val mk_bal   = HOLogic.mk_binrel "op <="
+  val dest_bal = HOLogic.dest_bin "op <=" hyprealT
+  val bal_add1 = hypreal_le_add_iff1 RS trans
+  val bal_add2 = hypreal_le_add_iff2 RS trans
+);
+
+val cancel_numerals = 
+  map prep_simproc
+   [("hyprealeq_cancel_numerals",
+     prep_pats ["(l::hypreal) + m = n", "(l::hypreal) = m + n", 
+		"(l::hypreal) - m = n", "(l::hypreal) = m - n", 
+		"(l::hypreal) * m = n", "(l::hypreal) = m * n"], 
+     EqCancelNumerals.proc),
+    ("hyprealless_cancel_numerals", 
+     prep_pats ["(l::hypreal) + m < n", "(l::hypreal) < m + n", 
+		"(l::hypreal) - m < n", "(l::hypreal) < m - n", 
+		"(l::hypreal) * m < n", "(l::hypreal) < m * n"], 
+     LessCancelNumerals.proc),
+    ("hyprealle_cancel_numerals", 
+     prep_pats ["(l::hypreal) + m <= n", "(l::hypreal) <= m + n", 
+		"(l::hypreal) - m <= n", "(l::hypreal) <= m - n", 
+		"(l::hypreal) * m <= n", "(l::hypreal) <= m * n"], 
+     LeCancelNumerals.proc)];
+
+
+structure CombineNumeralsData =
+  struct
+  val add		= op + : int*int -> int 
+  val mk_sum    	= long_mk_sum    (*to work for e.g. #2*x + #3*x *)
+  val dest_sum		= dest_sum
+  val mk_coeff		= mk_coeff
+  val dest_coeff	= dest_coeff 1
+  val left_distrib	= left_hypreal_add_mult_distrib RS trans
+  val prove_conv	= prove_conv_nohyps "hypreal_combine_numerals"
+  val trans_tac          = trans_tac
+  val norm_tac = 
+     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
+                                         hypreal_minus_simps@hypreal_add_ac))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@hypreal_mult_minus_simps))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps hypreal_minus_from_mult_simps@
+                                              hypreal_add_ac@hypreal_mult_ac))
+  val numeral_simp_tac	= ALLGOALS 
+                    (simp_tac (HOL_ss addsimps add_0s@bin_simps))
+  val simplify_meta_eq  = simplify_meta_eq
+  end;
+
+structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
+  
+val combine_numerals = 
+    prep_simproc ("hypreal_combine_numerals",
+		  prep_pats ["(i::hypreal) + j", "(i::hypreal) - j"],
+		  CombineNumerals.proc);
+
+
+(** Declarations for ExtractCommonTerm **)
+
+(*this version ALWAYS includes a trailing one*)
+fun long_mk_prod []        = one
+  | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts);
+
+(*Find first term that matches u*)
+fun find_first past u []         = raise TERM("find_first", []) 
+  | find_first past u (t::terms) =
+	if u aconv t then (rev past @ terms)
+        else find_first (t::past) u terms
+	handle TERM _ => find_first (t::past) u terms;
+
+(*Final simplification: cancel + and *  *)
+fun cancel_simplify_meta_eq cancel_th th = 
+    Int_Numeral_Simprocs.simplify_meta_eq 
+        [hypreal_mult_1, hypreal_mult_1_right] 
+        (([th, cancel_th]) MRS trans);
+
+end;
+
+Addsimprocs Hyperreal_Numeral_Simprocs.cancel_numerals;
+Addsimprocs [Hyperreal_Numeral_Simprocs.combine_numerals];
+
+(*The Abel_Cancel simprocs are now obsolete*)
+Delsimprocs [Hyperreal_Cancel.sum_conv, Hyperreal_Cancel.rel_conv];
+
+(*examples:
+print_depth 22;
+set timing;
+set trace_simp;
+fun test s = (Goal s, by (Simp_tac 1)); 
+
+test "l + #2 + #2 + #2 + (l + #2) + (oo + #2) = (uu::hypreal)";
+test "#2*u = (u::hypreal)";
+test "(i + j + #12 + (k::hypreal)) - #15 = y";
+test "(i + j + #12 + (k::hypreal)) - #5 = y";
+
+test "y - b < (b::hypreal)";
+test "y - (#3*b + c) < (b::hypreal) - #2*c";
+
+test "(#2*x - (u*v) + y) - v*#3*u = (w::hypreal)";
+test "(#2*x*u*v + (u*v)*#4 + y) - v*u*#4 = (w::hypreal)";
+test "(#2*x*u*v + (u*v)*#4 + y) - v*u = (w::hypreal)";
+test "u*v - (x*u*v + (u*v)*#4 + y) = (w::hypreal)";
+
+test "(i + j + #12 + (k::hypreal)) = u + #15 + y";
+test "(i + j*#2 + #12 + (k::hypreal)) = j + #5 + y";
+
+test "#2*y + #3*z + #6*w + #2*y + #3*z + #2*u = #2*y' + #3*z' + #6*w' + #2*y' + #3*z' + u + (vv::hypreal)";
+
+test "a + -(b+c) + b = (d::hypreal)";
+test "a + -(b+c) - b = (d::hypreal)";
+
+(*negative numerals*)
+test "(i + j + #-2 + (k::hypreal)) - (u + #5 + y) = zz";
+test "(i + j + #-3 + (k::hypreal)) < u + #5 + y";
+test "(i + j + #3 + (k::hypreal)) < u + #-6 + y";
+test "(i + j + #-12 + (k::hypreal)) - #15 = y";
+test "(i + j + #12 + (k::hypreal)) - #-15 = y";
+test "(i + j + #-12 + (k::hypreal)) - #-15 = y";
+*)
+
+
+(** Constant folding for hypreal plus and times **)
+
+(*We do not need
+    structure Hyperreal_Plus_Assoc = Assoc_Fold (Hyperreal_Plus_Assoc_Data);
+  because combine_numerals does the same thing*)
+
+structure Hyperreal_Times_Assoc_Data : ASSOC_FOLD_DATA =
+struct
+  val ss		= HOL_ss
+  val eq_reflection	= eq_reflection
+  val sg_ref    = Sign.self_ref (Theory.sign_of (the_context ()))
+  val T	     = Hyperreal_Numeral_Simprocs.hyprealT
+  val plus   = Const ("op *", [T,T] ---> T)
+  val add_ac = hypreal_mult_ac
+end;
+
+structure Hyperreal_Times_Assoc = Assoc_Fold (Hyperreal_Times_Assoc_Data);
+
+Addsimprocs [Hyperreal_Times_Assoc.conv];
+
+Addsimps [rename_numerals hypreal_of_real_zero_iff];
+
+(*Simplification of  x-y < 0, etc.*)
+AddIffs [hypreal_less_iff_diff_less_0 RS sym];
+AddIffs [hypreal_eq_iff_diff_eq_0 RS sym];
+AddIffs [hypreal_le_iff_diff_le_0 RS sym];
+
+
+(** number_of related to hypreal_of_real **)
+
+Goal "(number_of w < hypreal_of_real z) = (number_of w < z)";
+by (stac (hypreal_of_real_less_iff RS sym) 1); 
+by (Simp_tac 1); 
+qed "number_of_less_hypreal_of_real_iff";
+Addsimps [number_of_less_hypreal_of_real_iff];
+
+Goal "(number_of w <= hypreal_of_real z) = (number_of w <= z)";
+by (stac (hypreal_of_real_le_iff RS sym) 1); 
+by (Simp_tac 1); 
+qed "number_of_le_hypreal_of_real_iff";
+Addsimps [number_of_le_hypreal_of_real_iff];
+
+Goal "(hypreal_of_real z < number_of w) = (z < number_of w)";
+by (stac (hypreal_of_real_less_iff RS sym) 1); 
+by (Simp_tac 1); 
+qed "hypreal_of_real_less_number_of_iff";
+Addsimps [hypreal_of_real_less_number_of_iff];
+
+Goal "(hypreal_of_real z <= number_of w) = (z <= number_of w)";
+by (stac (hypreal_of_real_le_iff RS sym) 1); 
+by (Simp_tac 1); 
+qed "hypreal_of_real_le_number_of_iff";
+Addsimps [hypreal_of_real_le_number_of_iff];
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HyperBin.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,21 @@
+(*  Title:      HOL/HyperBin.thy
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1999  University of Cambridge
+
+Binary arithmetic for the hyperreals
+
+This case is reduced to that for the reals.
+*)
+
+HyperBin = HyperOrd +
+
+instance
+  hypreal :: number 
+
+defs
+  hypreal_number_of_def
+    "number_of v == hypreal_of_real (number_of v)"
+     (*::bin=>hypreal               ::bin=>real*)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HyperDef.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,1354 @@
+(*  Title       : HOL/Real/Hyperreal/Hyper.ML
+    ID          : $Id$
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Ultrapower construction of hyperreals
+*) 
+
+(*------------------------------------------------------------------------
+             Proof that the set of naturals is not finite
+ ------------------------------------------------------------------------*)
+
+(*** based on James' proof that the set of naturals is not finite ***)
+Goal "finite (A::nat set) --> (EX n. ALL m. Suc (n + m) ~: A)";
+by (rtac impI 1);
+by (eres_inst_tac [("F","A")] finite_induct 1);
+by (Blast_tac 1 THEN etac exE 1);
+by (res_inst_tac [("x","n + x")] exI 1);
+by (rtac allI 1 THEN eres_inst_tac [("x","x + m")] allE 1);
+by (auto_tac (claset(), simpset() addsimps add_ac));
+by (auto_tac (claset(),
+	      simpset() addsimps [add_assoc RS sym,
+				  less_add_Suc2 RS less_not_refl2]));
+qed_spec_mp "finite_exhausts";
+
+Goal "finite (A :: nat set) --> (EX n. n ~:A)";
+by (rtac impI 1 THEN dtac finite_exhausts 1);
+by (Blast_tac 1);
+qed_spec_mp "finite_not_covers";
+
+Goal "~ finite(UNIV:: nat set)";
+by (fast_tac (claset() addSDs [finite_exhausts]) 1);
+qed "not_finite_nat";
+
+(*------------------------------------------------------------------------
+   Existence of free ultrafilter over the naturals and proof of various 
+   properties of the FreeUltrafilterNat- an arbitrary free ultrafilter
+ ------------------------------------------------------------------------*)
+
+Goal "EX U. U: FreeUltrafilter (UNIV::nat set)";
+by (rtac (not_finite_nat RS FreeUltrafilter_Ex) 1);
+qed "FreeUltrafilterNat_Ex";
+
+Goalw [FreeUltrafilterNat_def] 
+     "FreeUltrafilterNat: FreeUltrafilter(UNIV:: nat set)";
+by (rtac (FreeUltrafilterNat_Ex RS exE) 1);
+by (rtac someI2 1 THEN ALLGOALS(assume_tac));
+qed "FreeUltrafilterNat_mem";
+Addsimps [FreeUltrafilterNat_mem];
+
+Goalw [FreeUltrafilterNat_def] "finite x ==> x ~: FreeUltrafilterNat";
+by (rtac (FreeUltrafilterNat_Ex RS exE) 1);
+by (rtac someI2 1 THEN assume_tac 1);
+by (blast_tac (claset() addDs [mem_FreeUltrafiltersetD1]) 1);
+qed "FreeUltrafilterNat_finite";
+
+Goal "x: FreeUltrafilterNat ==> ~ finite x";
+by (blast_tac (claset() addDs [FreeUltrafilterNat_finite]) 1);
+qed "FreeUltrafilterNat_not_finite";
+
+Goalw [FreeUltrafilterNat_def] "{} ~: FreeUltrafilterNat";
+by (rtac (FreeUltrafilterNat_Ex RS exE) 1);
+by (rtac someI2 1 THEN assume_tac 1);
+by (blast_tac (claset() addDs [FreeUltrafilter_Ultrafilter,
+			       Ultrafilter_Filter,Filter_empty_not_mem]) 1);
+qed "FreeUltrafilterNat_empty";
+Addsimps [FreeUltrafilterNat_empty];
+
+Goal "[| X: FreeUltrafilterNat;  Y: FreeUltrafilterNat |]  \
+\     ==> X Int Y : FreeUltrafilterNat";
+by (cut_facts_tac [FreeUltrafilterNat_mem] 1);
+by (blast_tac (claset() addDs [FreeUltrafilter_Ultrafilter,
+			       Ultrafilter_Filter,mem_FiltersetD1]) 1);
+qed "FreeUltrafilterNat_Int";
+
+Goal "[| X: FreeUltrafilterNat;  X <= Y |] \
+\     ==> Y : FreeUltrafilterNat";
+by (cut_facts_tac [FreeUltrafilterNat_mem] 1);
+by (blast_tac (claset() addDs [FreeUltrafilter_Ultrafilter,
+			       Ultrafilter_Filter,mem_FiltersetD2]) 1);
+qed "FreeUltrafilterNat_subset";
+
+Goal "X: FreeUltrafilterNat ==> -X ~: FreeUltrafilterNat";
+by (Step_tac 1);
+by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
+by Auto_tac;
+qed "FreeUltrafilterNat_Compl";
+
+Goal "X~: FreeUltrafilterNat ==> -X : FreeUltrafilterNat";
+by (cut_facts_tac [FreeUltrafilterNat_mem RS (FreeUltrafilter_iff RS iffD1)] 1);
+by (Step_tac 1 THEN dres_inst_tac [("x","X")] bspec 1);
+by (auto_tac (claset(), simpset() addsimps [UNIV_diff_Compl]));
+qed "FreeUltrafilterNat_Compl_mem";
+
+Goal "(X ~: FreeUltrafilterNat) = (-X: FreeUltrafilterNat)";
+by (blast_tac (claset() addDs [FreeUltrafilterNat_Compl,
+			       FreeUltrafilterNat_Compl_mem]) 1);
+qed "FreeUltrafilterNat_Compl_iff1";
+
+Goal "(X: FreeUltrafilterNat) = (-X ~: FreeUltrafilterNat)";
+by (auto_tac (claset(),
+	      simpset() addsimps [FreeUltrafilterNat_Compl_iff1 RS sym]));
+qed "FreeUltrafilterNat_Compl_iff2";
+
+Goal "(UNIV::nat set) : FreeUltrafilterNat";
+by (rtac (FreeUltrafilterNat_mem RS FreeUltrafilter_Ultrafilter RS 
+          Ultrafilter_Filter RS mem_FiltersetD4) 1);
+qed "FreeUltrafilterNat_UNIV";
+Addsimps [FreeUltrafilterNat_UNIV];
+
+Goal "UNIV : FreeUltrafilterNat";
+by Auto_tac;
+qed "FreeUltrafilterNat_Nat_set";
+Addsimps [FreeUltrafilterNat_Nat_set];
+
+Goal "{n. P(n) = P(n)} : FreeUltrafilterNat";
+by (Simp_tac 1);
+qed "FreeUltrafilterNat_Nat_set_refl";
+AddIs [FreeUltrafilterNat_Nat_set_refl];
+
+Goal "{n::nat. P} : FreeUltrafilterNat ==> P";
+by (rtac ccontr 1);
+by (rotate_tac 1 1);
+by (Asm_full_simp_tac 1);
+qed "FreeUltrafilterNat_P";
+
+Goal "{n. P(n)} : FreeUltrafilterNat ==> EX n. P(n)";
+by (rtac ccontr 1 THEN rotate_tac 1 1);
+by (Asm_full_simp_tac 1);
+qed "FreeUltrafilterNat_Ex_P";
+
+Goal "ALL n. P(n) ==> {n. P(n)} : FreeUltrafilterNat";
+by (auto_tac (claset() addIs [FreeUltrafilterNat_Nat_set], simpset()));
+qed "FreeUltrafilterNat_all";
+
+(*-------------------------------------------------------
+     Define and use Ultrafilter tactics
+ -------------------------------------------------------*)
+use "fuf.ML";
+
+(*-------------------------------------------------------
+  Now prove one further property of our free ultrafilter
+ -------------------------------------------------------*)
+Goal "X Un Y: FreeUltrafilterNat \
+\     ==> X: FreeUltrafilterNat | Y: FreeUltrafilterNat";
+by Auto_tac;
+by (Ultra_tac 1);
+qed "FreeUltrafilterNat_Un";
+
+(*-------------------------------------------------------
+   Properties of hyprel
+ -------------------------------------------------------*)
+
+(** Proving that hyprel is an equivalence relation **)
+(** Natural deduction for hyprel **)
+
+Goalw [hyprel_def]
+   "((X,Y): hyprel) = ({n. X n = Y n}: FreeUltrafilterNat)";
+by (Fast_tac 1);
+qed "hyprel_iff";
+
+Goalw [hyprel_def] 
+     "{n. X n = Y n}: FreeUltrafilterNat  ==> (X,Y): hyprel";
+by (Fast_tac 1);
+qed "hyprelI";
+
+Goalw [hyprel_def]
+  "p: hyprel --> (EX X Y. \
+\                 p = (X,Y) & {n. X n = Y n} : FreeUltrafilterNat)";
+by (Fast_tac 1);
+qed "hyprelE_lemma";
+
+val [major,minor] = goal (the_context ())
+  "[| p: hyprel;  \
+\     !!X Y. [| p = (X,Y); {n. X n = Y n}: FreeUltrafilterNat\
+\                    |] ==> Q |] ==> Q";
+by (cut_facts_tac [major RS (hyprelE_lemma RS mp)] 1);
+by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
+qed "hyprelE";
+
+AddSIs [hyprelI];
+AddSEs [hyprelE];
+
+Goalw [hyprel_def] "(x,x): hyprel";
+by (auto_tac (claset(),
+              simpset() addsimps [FreeUltrafilterNat_Nat_set]));
+qed "hyprel_refl";
+
+Goal "{n. X n = Y n} = {n. Y n = X n}";
+by Auto_tac;
+qed "lemma_perm";
+
+Goalw [hyprel_def] "(x,y): hyprel --> (y,x):hyprel";
+by (auto_tac (claset() addIs [lemma_perm RS subst], simpset()));
+qed_spec_mp "hyprel_sym";
+
+Goalw [hyprel_def]
+      "(x,y): hyprel --> (y,z):hyprel --> (x,z):hyprel";
+by Auto_tac;
+by (Ultra_tac 1);
+qed_spec_mp "hyprel_trans";
+
+Goalw [equiv_def, refl_def, sym_def, trans_def] "equiv UNIV hyprel";
+by (auto_tac (claset() addSIs [hyprel_refl] 
+                       addSEs [hyprel_sym,hyprel_trans] 
+                       delrules [hyprelI,hyprelE],
+	      simpset() addsimps [FreeUltrafilterNat_Nat_set]));
+qed "equiv_hyprel";
+
+(* (hyprel ^^ {x} = hyprel ^^ {y}) = ((x,y) : hyprel) *)
+bind_thm ("equiv_hyprel_iff",
+    	  [equiv_hyprel, UNIV_I, UNIV_I] MRS eq_equiv_class_iff);
+
+Goalw  [hypreal_def,hyprel_def,quotient_def] "hyprel^^{x}:hypreal";
+by (Blast_tac 1);
+qed "hyprel_in_hypreal";
+
+Goal "inj_on Abs_hypreal hypreal";
+by (rtac inj_on_inverseI 1);
+by (etac Abs_hypreal_inverse 1);
+qed "inj_on_Abs_hypreal";
+
+Addsimps [equiv_hyprel_iff,inj_on_Abs_hypreal RS inj_on_iff,
+          hyprel_iff, hyprel_in_hypreal, Abs_hypreal_inverse];
+
+Addsimps [equiv_hyprel RS eq_equiv_class_iff];
+bind_thm ("eq_hyprelD", equiv_hyprel RSN (2,eq_equiv_class));
+
+Goal "inj(Rep_hypreal)";
+by (rtac inj_inverseI 1);
+by (rtac Rep_hypreal_inverse 1);
+qed "inj_Rep_hypreal";
+
+Goalw [hyprel_def] "x: hyprel ^^ {x}";
+by (Step_tac 1);
+by (auto_tac (claset() addSIs [FreeUltrafilterNat_Nat_set], simpset()));
+qed "lemma_hyprel_refl";
+
+Addsimps [lemma_hyprel_refl];
+
+Goalw [hypreal_def] "{} ~: hypreal";
+by (auto_tac (claset() addSEs [quotientE], simpset()));
+qed "hypreal_empty_not_mem";
+
+Addsimps [hypreal_empty_not_mem];
+
+Goal "Rep_hypreal x ~= {}";
+by (cut_inst_tac [("x","x")] Rep_hypreal 1);
+by Auto_tac;
+qed "Rep_hypreal_nonempty";
+
+Addsimps [Rep_hypreal_nonempty];
+
+(*------------------------------------------------------------------------
+   hypreal_of_real: the injection from real to hypreal
+ ------------------------------------------------------------------------*)
+
+Goal "inj(hypreal_of_real)";
+by (rtac injI 1);
+by (rewtac hypreal_of_real_def);
+by (dtac (inj_on_Abs_hypreal RS inj_onD) 1);
+by (REPEAT (rtac hyprel_in_hypreal 1));
+by (dtac eq_equiv_class 1);
+by (rtac equiv_hyprel 1);
+by (Fast_tac 1);
+by (rtac ccontr 1 THEN rotate_tac 1 1);
+by Auto_tac;
+qed "inj_hypreal_of_real";
+
+val [prem] = goal (the_context ())
+    "(!!x y. z = Abs_hypreal(hyprel^^{x}) ==> P) ==> P";
+by (res_inst_tac [("x1","z")] 
+    (rewrite_rule [hypreal_def] Rep_hypreal RS quotientE) 1);
+by (dres_inst_tac [("f","Abs_hypreal")] arg_cong 1);
+by (res_inst_tac [("x","x")] prem 1);
+by (asm_full_simp_tac (simpset() addsimps [Rep_hypreal_inverse]) 1);
+qed "eq_Abs_hypreal";
+
+(**** hypreal_minus: additive inverse on hypreal ****)
+
+Goalw [congruent_def]
+  "congruent hyprel (%X. hyprel^^{%n. - (X n)})";
+by Safe_tac;
+by (ALLGOALS Ultra_tac);
+qed "hypreal_minus_congruent";
+
+Goalw [hypreal_minus_def]
+   "- (Abs_hypreal(hyprel^^{%n. X n})) = Abs_hypreal(hyprel ^^ {%n. -(X n)})";
+by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
+by (simp_tac (simpset() addsimps 
+      [hyprel_in_hypreal RS Abs_hypreal_inverse,
+       [equiv_hyprel, hypreal_minus_congruent] MRS UN_equiv_class]) 1);
+qed "hypreal_minus";
+
+Goal "- (- z) = (z::hypreal)";
+by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_minus]) 1);
+qed "hypreal_minus_minus";
+
+Addsimps [hypreal_minus_minus];
+
+Goal "inj(%r::hypreal. -r)";
+by (rtac injI 1);
+by (dres_inst_tac [("f","uminus")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_minus_minus]) 1);
+qed "inj_hypreal_minus";
+
+Goalw [hypreal_zero_def] "-0 = (0::hypreal)";
+by (simp_tac (simpset() addsimps [hypreal_minus]) 1);
+qed "hypreal_minus_zero";
+Addsimps [hypreal_minus_zero];
+
+Goal "(-x = 0) = (x = (0::hypreal))"; 
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+       simpset() addsimps [hypreal_zero_def, hypreal_minus, eq_commute] @ 
+                          real_add_ac));
+qed "hypreal_minus_zero_iff";
+
+Addsimps [hypreal_minus_zero_iff];
+
+
+(**** hyperreal addition: hypreal_add  ****)
+
+Goalw [congruent2_def]
+    "congruent2 hyprel (%X Y. hyprel^^{%n. X n + Y n})";
+by Safe_tac;
+by (ALLGOALS(Ultra_tac));
+qed "hypreal_add_congruent2";
+
+Goalw [hypreal_add_def]
+  "Abs_hypreal(hyprel^^{%n. X n}) + Abs_hypreal(hyprel^^{%n. Y n}) = \
+\  Abs_hypreal(hyprel^^{%n. X n + Y n})";
+by (simp_tac (simpset() addsimps 
+         [[equiv_hyprel, hypreal_add_congruent2] MRS UN_equiv_class2]) 1);
+qed "hypreal_add";
+
+Goal "Abs_hypreal(hyprel^^{%n. X n}) - Abs_hypreal(hyprel^^{%n. Y n}) = \
+\     Abs_hypreal(hyprel^^{%n. X n - Y n})";
+by (simp_tac (simpset() addsimps 
+         [hypreal_diff_def, hypreal_minus,hypreal_add]) 1);
+qed "hypreal_diff";
+
+Goal "(z::hypreal) + w = w + z";
+by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","w")] eq_Abs_hypreal 1);
+by (asm_simp_tac (simpset() addsimps (real_add_ac @ [hypreal_add])) 1);
+qed "hypreal_add_commute";
+
+Goal "((z1::hypreal) + z2) + z3 = z1 + (z2 + z3)";
+by (res_inst_tac [("z","z1")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","z2")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","z3")] eq_Abs_hypreal 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_add, real_add_assoc]) 1);
+qed "hypreal_add_assoc";
+
+(*For AC rewriting*)
+Goal "(x::hypreal)+(y+z)=y+(x+z)";
+by (rtac (hypreal_add_commute RS trans) 1);
+by (rtac (hypreal_add_assoc RS trans) 1);
+by (rtac (hypreal_add_commute RS arg_cong) 1);
+qed "hypreal_add_left_commute";
+
+(* hypreal addition is an AC operator *)
+bind_thms ("hypreal_add_ac", [hypreal_add_assoc,hypreal_add_commute,
+                      hypreal_add_left_commute]);
+
+Goalw [hypreal_zero_def] "(0::hypreal) + z = z";
+by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
+by (asm_full_simp_tac (simpset() addsimps 
+    [hypreal_add]) 1);
+qed "hypreal_add_zero_left";
+
+Goal "z + (0::hypreal) = z";
+by (simp_tac (simpset() addsimps 
+    [hypreal_add_zero_left,hypreal_add_commute]) 1);
+qed "hypreal_add_zero_right";
+
+Goalw [hypreal_zero_def] "z + -z = (0::hypreal)";
+by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_minus, hypreal_add]) 1);
+qed "hypreal_add_minus";
+
+Goal "-z + z = (0::hypreal)";
+by (simp_tac (simpset() addsimps [hypreal_add_commute, hypreal_add_minus]) 1);
+qed "hypreal_add_minus_left";
+
+Addsimps [hypreal_add_minus,hypreal_add_minus_left,
+          hypreal_add_zero_left,hypreal_add_zero_right];
+
+Goal "EX y. (x::hypreal) + y = 0";
+by (fast_tac (claset() addIs [hypreal_add_minus]) 1);
+qed "hypreal_minus_ex";
+
+Goal "EX! y. (x::hypreal) + y = 0";
+by (auto_tac (claset() addIs [hypreal_add_minus], simpset()));
+by (dres_inst_tac [("f","%x. ya+x")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
+qed "hypreal_minus_ex1";
+
+Goal "EX! y. y + (x::hypreal) = 0";
+by (auto_tac (claset() addIs [hypreal_add_minus_left], simpset()));
+by (dres_inst_tac [("f","%x. x+ya")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc]) 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
+qed "hypreal_minus_left_ex1";
+
+Goal "x + y = (0::hypreal) ==> x = -y";
+by (cut_inst_tac [("z","y")] hypreal_add_minus_left 1);
+by (res_inst_tac [("x1","y")] (hypreal_minus_left_ex1 RS ex1E) 1);
+by (Blast_tac 1);
+qed "hypreal_add_minus_eq_minus";
+
+Goal "EX y::hypreal. x = -y";
+by (cut_inst_tac [("x","x")] hypreal_minus_ex 1);
+by (etac exE 1 THEN dtac hypreal_add_minus_eq_minus 1);
+by (Fast_tac 1);
+qed "hypreal_as_add_inverse_ex";
+
+Goal "-(x + (y::hypreal)) = -x + -y";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+              simpset() addsimps [hypreal_minus, hypreal_add,
+                                  real_minus_add_distrib]));
+qed "hypreal_minus_add_distrib";
+Addsimps [hypreal_minus_add_distrib];
+
+Goal "-(y + -(x::hypreal)) = x + -y";
+by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
+qed "hypreal_minus_distrib1";
+
+Goal "(x + - (y::hypreal)) + (y + - z) = x + -z";
+by (res_inst_tac [("w1","y")] (hypreal_add_commute RS subst) 1);
+by (simp_tac (simpset() addsimps [hypreal_add_left_commute,
+                                  hypreal_add_assoc]) 1);
+by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
+qed "hypreal_add_minus_cancel1";
+
+Goal "((x::hypreal) + y = x + z) = (y = z)";
+by (Step_tac 1);
+by (dres_inst_tac [("f","%t.-x + t")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
+qed "hypreal_add_left_cancel";
+
+Goal "z + (x + (y + -z)) = x + (y::hypreal)";
+by (simp_tac (simpset() addsimps hypreal_add_ac) 1);
+qed "hypreal_add_minus_cancel2";
+Addsimps [hypreal_add_minus_cancel2];
+
+Goal "y + -(x + y) = -(x::hypreal)";
+by (Full_simp_tac 1);
+by (rtac (hypreal_add_left_commute RS subst) 1);
+by (Full_simp_tac 1);
+qed "hypreal_add_minus_cancel";
+Addsimps [hypreal_add_minus_cancel];
+
+Goal "y + -(y + x) = -(x::hypreal)";
+by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
+qed "hypreal_add_minus_cancelc";
+Addsimps [hypreal_add_minus_cancelc];
+
+Goal "(z + -x) + (y + -z) = (y + -(x::hypreal))";
+by (full_simp_tac
+    (simpset() addsimps [hypreal_minus_add_distrib RS sym, 
+                         hypreal_add_left_cancel] @ hypreal_add_ac 
+               delsimps [hypreal_minus_add_distrib]) 1); 
+qed "hypreal_add_minus_cancel3";
+Addsimps [hypreal_add_minus_cancel3];
+
+Goal "(y + (x::hypreal)= z + x) = (y = z)";
+by (simp_tac (simpset() addsimps [hypreal_add_commute,
+                                  hypreal_add_left_cancel]) 1);
+qed "hypreal_add_right_cancel";
+
+Goal "z + (y + -z) = (y::hypreal)";
+by (simp_tac (simpset() addsimps hypreal_add_ac) 1);
+qed "hypreal_add_minus_cancel4";
+Addsimps [hypreal_add_minus_cancel4];
+
+Goal "z + (w + (x + (-z + y))) = w + x + (y::hypreal)";
+by (simp_tac (simpset() addsimps hypreal_add_ac) 1);
+qed "hypreal_add_minus_cancel5";
+Addsimps [hypreal_add_minus_cancel5];
+
+Goal "z + ((- z) + w) = (w::hypreal)";
+by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
+qed "hypreal_add_minus_cancelA";
+
+Goal "(-z) + (z + w) = (w::hypreal)";
+by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
+qed "hypreal_minus_add_cancelA";
+
+Addsimps [hypreal_add_minus_cancelA, hypreal_minus_add_cancelA];
+
+(**** hyperreal multiplication: hypreal_mult  ****)
+
+Goalw [congruent2_def]
+    "congruent2 hyprel (%X Y. hyprel^^{%n. X n * Y n})";
+by Safe_tac;
+by (ALLGOALS(Ultra_tac));
+qed "hypreal_mult_congruent2";
+
+Goalw [hypreal_mult_def]
+  "Abs_hypreal(hyprel^^{%n. X n}) * Abs_hypreal(hyprel^^{%n. Y n}) = \
+\  Abs_hypreal(hyprel^^{%n. X n * Y n})";
+by (simp_tac (simpset() addsimps 
+      [[equiv_hyprel, hypreal_mult_congruent2] MRS UN_equiv_class2]) 1);
+qed "hypreal_mult";
+
+Goal "(z::hypreal) * w = w * z";
+by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","w")] eq_Abs_hypreal 1);
+by (asm_simp_tac (simpset() addsimps ([hypreal_mult] @ real_mult_ac)) 1);
+qed "hypreal_mult_commute";
+
+Goal "((z1::hypreal) * z2) * z3 = z1 * (z2 * z3)";
+by (res_inst_tac [("z","z1")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","z2")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","z3")] eq_Abs_hypreal 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_mult,real_mult_assoc]) 1);
+qed "hypreal_mult_assoc";
+
+qed_goal "hypreal_mult_left_commute" (the_context ())
+    "(z1::hypreal) * (z2 * z3) = z2 * (z1 * z3)"
+ (fn _ => [rtac (hypreal_mult_commute RS trans) 1, 
+           rtac (hypreal_mult_assoc RS trans) 1,
+           rtac (hypreal_mult_commute RS arg_cong) 1]);
+
+(* hypreal multiplication is an AC operator *)
+bind_thms ("hypreal_mult_ac", [hypreal_mult_assoc, hypreal_mult_commute, 
+                       hypreal_mult_left_commute]);
+
+Goalw [hypreal_one_def] "1hr * z = z";
+by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_mult]) 1);
+qed "hypreal_mult_1";
+
+Goal "z * 1hr = z";
+by (simp_tac (simpset() addsimps [hypreal_mult_commute,
+    hypreal_mult_1]) 1);
+qed "hypreal_mult_1_right";
+
+Goalw [hypreal_zero_def] "0 * z = (0::hypreal)";
+by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_mult,real_mult_0]) 1);
+qed "hypreal_mult_0";
+
+Goal "z * 0 = (0::hypreal)";
+by (simp_tac (simpset() addsimps [hypreal_mult_commute, hypreal_mult_0]) 1);
+qed "hypreal_mult_0_right";
+
+Addsimps [hypreal_mult_0,hypreal_mult_0_right];
+
+Goal "-(x * y) = -x * (y::hypreal)";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+	      simpset() addsimps [hypreal_minus, hypreal_mult] 
+                                 @ real_mult_ac @ real_add_ac));
+qed "hypreal_minus_mult_eq1";
+
+Goal "-(x * y) = (x::hypreal) * -y";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_minus, hypreal_mult] 
+                                           @ real_mult_ac @ real_add_ac));
+qed "hypreal_minus_mult_eq2";
+
+(*Pull negations out*)
+Addsimps [hypreal_minus_mult_eq2 RS sym, hypreal_minus_mult_eq1 RS sym];
+
+Goal "-x*y = (x::hypreal)*-y";
+by Auto_tac;
+qed "hypreal_minus_mult_commute";
+
+(*-----------------------------------------------------------------------------
+    A few more theorems
+ ----------------------------------------------------------------------------*)
+Goal "(z::hypreal) + v = z' + v' ==> z + (v + w) = z' + (v' + w)";
+by (asm_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
+qed "hypreal_add_assoc_cong";
+
+Goal "(z::hypreal) + (v + w) = v + (z + w)";
+by (REPEAT (ares_tac [hypreal_add_commute RS hypreal_add_assoc_cong] 1));
+qed "hypreal_add_assoc_swap";
+
+Goal "((z1::hypreal) + z2) * w = (z1 * w) + (z2 * w)";
+by (res_inst_tac [("z","z1")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","z2")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","w")] eq_Abs_hypreal 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_mult,hypreal_add,
+     real_add_mult_distrib]) 1);
+qed "hypreal_add_mult_distrib";
+
+val hypreal_mult_commute'= read_instantiate [("z","w")] hypreal_mult_commute;
+
+Goal "(w::hypreal) * (z1 + z2) = (w * z1) + (w * z2)";
+by (simp_tac (simpset() addsimps [hypreal_mult_commute',hypreal_add_mult_distrib]) 1);
+qed "hypreal_add_mult_distrib2";
+
+bind_thms ("hypreal_mult_simps", [hypreal_mult_1, hypreal_mult_1_right]);
+Addsimps hypreal_mult_simps;
+
+(* 07/00 *)
+
+Goalw [hypreal_diff_def] "((z1::hypreal) - z2) * w = (z1 * w) - (z2 * w)";
+by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib]) 1);
+qed "hypreal_diff_mult_distrib";
+
+Goal "(w::hypreal) * (z1 - z2) = (w * z1) - (w * z2)";
+by (simp_tac (simpset() addsimps [hypreal_mult_commute', 
+				  hypreal_diff_mult_distrib]) 1);
+qed "hypreal_diff_mult_distrib2";
+
+(*** one and zero are distinct ***)
+Goalw [hypreal_zero_def,hypreal_one_def] "0 ~= 1hr";
+by (auto_tac (claset(), simpset() addsimps [real_zero_not_eq_one]));
+qed "hypreal_zero_not_eq_one";
+
+
+(**** multiplicative inverse on hypreal ****)
+
+Goalw [congruent_def]
+  "congruent hyprel (%X. hyprel^^{%n. if X n = #0 then #0 else inverse(X n)})";
+by (Auto_tac THEN Ultra_tac 1);
+qed "hypreal_inverse_congruent";
+
+Goalw [hypreal_inverse_def]
+      "inverse (Abs_hypreal(hyprel^^{%n. X n})) = \
+\      Abs_hypreal(hyprel ^^ {%n. if X n = #0 then #0 else inverse(X n)})";
+by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
+by (simp_tac (simpset() addsimps 
+   [hyprel_in_hypreal RS Abs_hypreal_inverse,
+    [equiv_hyprel, hypreal_inverse_congruent] MRS UN_equiv_class]) 1);
+qed "hypreal_inverse";
+
+Goal "inverse 0 = (0::hypreal)";
+by (simp_tac (simpset() addsimps [hypreal_inverse, hypreal_zero_def]) 1);
+qed "HYPREAL_INVERSE_ZERO";
+
+Goal "a / (0::hypreal) = 0";
+by (simp_tac
+    (simpset() addsimps [hypreal_divide_def, HYPREAL_INVERSE_ZERO]) 1);
+qed "HYPREAL_DIVISION_BY_ZERO";  (*NOT for adding to default simpset*)
+
+fun hypreal_div_undefined_case_tac s i =
+  case_tac s i THEN 
+  asm_simp_tac 
+       (simpset() addsimps [HYPREAL_DIVISION_BY_ZERO, HYPREAL_INVERSE_ZERO]) i;
+
+Goal "inverse (inverse (z::hypreal)) = z";
+by (hypreal_div_undefined_case_tac "z=0" 1);
+by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
+by (asm_full_simp_tac (simpset() addsimps 
+                       [hypreal_inverse, hypreal_zero_def]) 1);
+qed "hypreal_inverse_inverse";
+Addsimps [hypreal_inverse_inverse];
+
+Goalw [hypreal_one_def] "inverse(1hr) = 1hr";
+by (full_simp_tac (simpset() addsimps [hypreal_inverse,
+                                       real_zero_not_eq_one RS not_sym]) 1);
+qed "hypreal_inverse_1";
+Addsimps [hypreal_inverse_1];
+
+
+(*** existence of inverse ***)
+
+Goalw [hypreal_one_def,hypreal_zero_def] 
+     "x ~= 0 ==> x*inverse(x) = 1hr";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (rotate_tac 1 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_inverse, hypreal_mult]) 1);
+by (dtac FreeUltrafilterNat_Compl_mem 1);
+by (blast_tac (claset() addSIs [real_mult_inv_right,
+    FreeUltrafilterNat_subset]) 1);
+qed "hypreal_mult_inverse";
+
+Goal "x ~= 0 ==> inverse(x)*x = 1hr";
+by (asm_simp_tac (simpset() addsimps [hypreal_mult_inverse,
+				      hypreal_mult_commute]) 1);
+qed "hypreal_mult_inverse_left";
+
+Goal "(c::hypreal) ~= 0 ==> (c*a=c*b) = (a=b)";
+by Auto_tac;
+by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_mult_inverse] @ hypreal_mult_ac)  1);
+qed "hypreal_mult_left_cancel";
+    
+Goal "(c::hypreal) ~= 0 ==> (a*c=b*c) = (a=b)";
+by (Step_tac 1);
+by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_mult_inverse] @ hypreal_mult_ac)  1);
+qed "hypreal_mult_right_cancel";
+
+Goalw [hypreal_zero_def] "x ~= 0 ==> inverse (x::hypreal) ~= 0";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_inverse, hypreal_mult]) 1);
+qed "hypreal_inverse_not_zero";
+
+Addsimps [hypreal_mult_inverse,hypreal_mult_inverse_left];
+
+Goal "[| x ~= 0; y ~= 0 |] ==> x * y ~= (0::hypreal)";
+by (Step_tac 1);
+by (dres_inst_tac [("f","%z. inverse x*z")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym]) 1);
+qed "hypreal_mult_not_0";
+
+Goal "x*y = (0::hypreal) ==> x = 0 | y = 0";
+by (auto_tac (claset() addIs [ccontr] addDs [hypreal_mult_not_0],
+              simpset()));
+qed "hypreal_mult_zero_disj";
+
+Goal "inverse(-x) = -inverse(x::hypreal)";
+by (hypreal_div_undefined_case_tac "x=0" 1);
+by (rtac (hypreal_mult_right_cancel RS iffD1) 1);
+by (stac hypreal_mult_inverse_left 2);
+by Auto_tac;
+qed "hypreal_minus_inverse";
+
+Goal "inverse(x*y) = inverse(x)*inverse(y::hypreal)";
+by (hypreal_div_undefined_case_tac "x=0" 1);
+by (hypreal_div_undefined_case_tac "y=0" 1);
+by (forw_inst_tac [("y","y")] hypreal_mult_not_0 1 THEN assume_tac 1);
+by (res_inst_tac [("c1","x")] (hypreal_mult_left_cancel RS iffD1) 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_mult_assoc RS sym]));
+by (res_inst_tac [("c1","y")] (hypreal_mult_left_cancel RS iffD1) 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_mult_left_commute]));
+by (asm_simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym]) 1);
+qed "hypreal_inverse_distrib";
+
+(*------------------------------------------------------------------
+                   Theorems for ordering 
+ ------------------------------------------------------------------*)
+
+(* prove introduction and elimination rules for hypreal_less *)
+
+Goalw [hypreal_less_def]
+ "P < (Q::hypreal) = (EX X Y. X : Rep_hypreal(P) & \
+\                             Y : Rep_hypreal(Q) & \
+\                             {n. X n < Y n} : FreeUltrafilterNat)";
+by (Fast_tac 1);
+qed "hypreal_less_iff";
+
+Goalw [hypreal_less_def]
+ "[| {n. X n < Y n} : FreeUltrafilterNat; \
+\         X : Rep_hypreal(P); \
+\         Y : Rep_hypreal(Q) |] ==> P < (Q::hypreal)";
+by (Fast_tac 1);
+qed "hypreal_lessI";
+
+
+Goalw [hypreal_less_def]
+     "!! R1. [| R1 < (R2::hypreal); \
+\         !!X Y. {n. X n < Y n} : FreeUltrafilterNat ==> P; \
+\         !!X. X : Rep_hypreal(R1) ==> P; \ 
+\         !!Y. Y : Rep_hypreal(R2) ==> P |] \
+\     ==> P";
+by Auto_tac;
+qed "hypreal_lessE";
+
+Goalw [hypreal_less_def]
+ "R1 < (R2::hypreal) ==> (EX X Y. {n. X n < Y n} : FreeUltrafilterNat & \
+\                                  X : Rep_hypreal(R1) & \
+\                                  Y : Rep_hypreal(R2))";
+by (Fast_tac 1);
+qed "hypreal_lessD";
+
+Goal "~ (R::hypreal) < R";
+by (res_inst_tac [("z","R")] eq_Abs_hypreal 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_less_def]));
+by (Ultra_tac 1);
+qed "hypreal_less_not_refl";
+
+(*** y < y ==> P ***)
+bind_thm("hypreal_less_irrefl",hypreal_less_not_refl RS notE);
+AddSEs [hypreal_less_irrefl];
+
+Goal "!!(x::hypreal). x < y ==> x ~= y";
+by (auto_tac (claset(), simpset() addsimps [hypreal_less_not_refl]));
+qed "hypreal_not_refl2";
+
+Goal "!!(R1::hypreal). [| R1 < R2; R2 < R3 |] ==> R1 < R3";
+by (res_inst_tac [("z","R1")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","R2")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","R3")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addSIs [exI], simpset() addsimps [hypreal_less_def]));
+by (ultra_tac (claset() addIs [order_less_trans], simpset()) 1);
+qed "hypreal_less_trans";
+
+Goal "!! (R1::hypreal). [| R1 < R2; R2 < R1 |] ==> P";
+by (dtac hypreal_less_trans 1 THEN assume_tac 1);
+by (asm_full_simp_tac (simpset() addsimps 
+    [hypreal_less_not_refl]) 1);
+qed "hypreal_less_asym";
+
+(*-------------------------------------------------------
+  TODO: The following theorem should have been proved 
+  first and then used througout the proofs as it probably 
+  makes many of them more straightforward. 
+ -------------------------------------------------------*)
+Goalw [hypreal_less_def]
+      "(Abs_hypreal(hyprel^^{%n. X n}) < \
+\           Abs_hypreal(hyprel^^{%n. Y n})) = \
+\      ({n. X n < Y n} : FreeUltrafilterNat)";
+by (auto_tac (claset() addSIs [lemma_hyprel_refl], simpset()));
+by (Ultra_tac 1);
+qed "hypreal_less";
+
+(*---------------------------------------------------------------------------------
+             Hyperreals as a linearly ordered field
+ ---------------------------------------------------------------------------------*)
+(*** sum order 
+Goalw [hypreal_zero_def] 
+      "[| 0 < x; 0 < y |] ==> (0::hypreal) < x + y";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps
+    [hypreal_less_def,hypreal_add]));
+by (auto_tac (claset() addSIs [exI],simpset() addsimps
+    [hypreal_less_def,hypreal_add]));
+by (ultra_tac (claset() addIs [real_add_order],simpset()) 1);
+qed "hypreal_add_order";
+***)
+
+(*** mult order 
+Goalw [hypreal_zero_def] 
+          "[| 0 < x; 0 < y |] ==> (0::hypreal) < x * y";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addSIs [exI],simpset() addsimps
+    [hypreal_less_def,hypreal_mult]));
+by (ultra_tac (claset() addIs [rename_numerals real_mult_order],
+	       simpset()) 1);
+qed "hypreal_mult_order";
+****)
+
+
+(*---------------------------------------------------------------------------------
+                         Trichotomy of the hyperreals
+  --------------------------------------------------------------------------------*)
+
+Goalw [hyprel_def] "EX x. x: hyprel ^^ {%n. #0}";
+by (res_inst_tac [("x","%n. #0")] exI 1);
+by (Step_tac 1);
+by (auto_tac (claset() addSIs [FreeUltrafilterNat_Nat_set], simpset()));
+qed "lemma_hyprel_0r_mem";
+
+Goalw [hypreal_zero_def]"0 <  x | x = 0 | x < (0::hypreal)";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [hypreal_less_def]));
+by (cut_facts_tac [lemma_hyprel_0r_mem] 1 THEN etac exE 1);
+by (dres_inst_tac [("x","xa")] spec 1);
+by (dres_inst_tac [("x","x")] spec 1);
+by (cut_inst_tac [("x","x")] lemma_hyprel_refl 1);
+by Auto_tac;
+by (dres_inst_tac [("x","x")] spec 1);
+by (dres_inst_tac [("x","xa")] spec 1);
+by Auto_tac;
+by (Ultra_tac 1);
+by (auto_tac (claset() addIs [real_linear_less2],simpset()));
+qed "hypreal_trichotomy";
+
+val prems = Goal "[| (0::hypreal) < x ==> P; \
+\                 x = 0 ==> P; \
+\                 x < 0 ==> P |] ==> P";
+by (cut_inst_tac [("x","x")] hypreal_trichotomy 1);
+by (REPEAT (eresolve_tac (disjE::prems) 1));
+qed "hypreal_trichotomyE";
+
+(*----------------------------------------------------------------------------
+            More properties of <
+ ----------------------------------------------------------------------------*)
+
+Goal "((x::hypreal) < y) = (0 < y + -x)";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [hypreal_add,
+    hypreal_zero_def,hypreal_minus,hypreal_less]));
+by (ALLGOALS(Ultra_tac));
+qed "hypreal_less_minus_iff"; 
+
+Goal "((x::hypreal) < y) = (x + -y < 0)";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [hypreal_add,
+    hypreal_zero_def,hypreal_minus,hypreal_less]));
+by (ALLGOALS(Ultra_tac));
+qed "hypreal_less_minus_iff2";
+
+Goal "((x::hypreal) = y) = (x + - y = 0)";
+by Auto_tac;
+by (res_inst_tac [("x1","-y")] (hypreal_add_right_cancel RS iffD1) 1);
+by Auto_tac;
+qed "hypreal_eq_minus_iff"; 
+
+Goal "((x::hypreal) = y) = (0 = y + - x)";
+by Auto_tac;
+by (res_inst_tac [("x1","-x")] (hypreal_add_right_cancel RS iffD1) 1);
+by Auto_tac;
+qed "hypreal_eq_minus_iff2"; 
+
+(* 07/00 *)
+Goal "(0::hypreal) - x = -x";
+by (simp_tac (simpset() addsimps [hypreal_diff_def]) 1);
+qed "hypreal_diff_zero";
+
+Goal "x - (0::hypreal) = x";
+by (simp_tac (simpset() addsimps [hypreal_diff_def]) 1);
+qed "hypreal_diff_zero_right";
+
+Goal "x - x = (0::hypreal)";
+by (simp_tac (simpset() addsimps [hypreal_diff_def]) 1);
+qed "hypreal_diff_self";
+
+Addsimps [hypreal_diff_zero, hypreal_diff_zero_right, hypreal_diff_self];
+
+Goal "(x = y + z) = (x + -z = (y::hypreal))";
+by (auto_tac (claset(),simpset() addsimps [hypreal_add_assoc]));
+qed "hypreal_eq_minus_iff3";
+
+Goal "(x ~= a) = (x + -a ~= (0::hypreal))";
+by (auto_tac (claset() addDs [hypreal_eq_minus_iff RS iffD2],
+              simpset())); 
+qed "hypreal_not_eq_minus_iff";
+
+Goal "(x+y = (0::hypreal)) = (x = -y)";
+by (stac hypreal_eq_minus_iff 1);
+by Auto_tac;
+qed "hypreal_add_eq_0_iff";
+AddIffs [hypreal_add_eq_0_iff];
+
+
+(*** linearity ***)
+
+Goal "(x::hypreal) < y | x = y | y < x";
+by (stac hypreal_eq_minus_iff2 1);
+by (res_inst_tac [("x1","x")] (hypreal_less_minus_iff RS ssubst) 1);
+by (res_inst_tac [("x1","y")] (hypreal_less_minus_iff2 RS ssubst) 1);
+by (rtac hypreal_trichotomyE 1);
+by Auto_tac;
+qed "hypreal_linear";
+
+Goal "((w::hypreal) ~= z) = (w<z | z<w)";
+by (cut_facts_tac [hypreal_linear] 1);
+by (Blast_tac 1);
+qed "hypreal_neq_iff";
+
+Goal "!!(x::hypreal). [| x < y ==> P;  x = y ==> P; \
+\          y < x ==> P |] ==> P";
+by (cut_inst_tac [("x","x"),("y","y")] hypreal_linear 1);
+by Auto_tac;
+qed "hypreal_linear_less2";
+
+(*------------------------------------------------------------------------------
+                            Properties of <=
+ ------------------------------------------------------------------------------*)
+(*------ hypreal le iff reals le a.e ------*)
+
+Goalw [hypreal_le_def,real_le_def]
+      "(Abs_hypreal(hyprel^^{%n. X n}) <= \
+\           Abs_hypreal(hyprel^^{%n. Y n})) = \
+\      ({n. X n <= Y n} : FreeUltrafilterNat)";
+by (auto_tac (claset(),simpset() addsimps [hypreal_less]));
+by (ALLGOALS(Ultra_tac));
+qed "hypreal_le";
+
+(*---------------------------------------------------------*)
+(*---------------------------------------------------------*)
+Goalw [hypreal_le_def] 
+     "~(w < z) ==> z <= (w::hypreal)";
+by (assume_tac 1);
+qed "hypreal_leI";
+
+Goalw [hypreal_le_def] 
+      "z<=w ==> ~(w<(z::hypreal))";
+by (assume_tac 1);
+qed "hypreal_leD";
+
+bind_thm ("hypreal_leE", make_elim hypreal_leD);
+
+Goal "(~(w < z)) = (z <= (w::hypreal))";
+by (fast_tac (claset() addSIs [hypreal_leI,hypreal_leD]) 1);
+qed "hypreal_less_le_iff";
+
+Goalw [hypreal_le_def] "~ z <= w ==> w<(z::hypreal)";
+by (Fast_tac 1);
+qed "not_hypreal_leE";
+
+Goalw [hypreal_le_def] "!!(x::hypreal). x <= y ==> x < y | x = y";
+by (cut_facts_tac [hypreal_linear] 1);
+by (fast_tac (claset() addEs [hypreal_less_irrefl,hypreal_less_asym]) 1);
+qed "hypreal_le_imp_less_or_eq";
+
+Goalw [hypreal_le_def] "z<w | z=w ==> z <=(w::hypreal)";
+by (cut_facts_tac [hypreal_linear] 1);
+by (fast_tac (claset() addEs [hypreal_less_irrefl,hypreal_less_asym]) 1);
+qed "hypreal_less_or_eq_imp_le";
+
+Goal "(x <= (y::hypreal)) = (x < y | x=y)";
+by (REPEAT(ares_tac [iffI, hypreal_less_or_eq_imp_le, hypreal_le_imp_less_or_eq] 1));
+qed "hypreal_le_eq_less_or_eq";
+val hypreal_le_less = hypreal_le_eq_less_or_eq;
+
+Goal "w <= (w::hypreal)";
+by (simp_tac (simpset() addsimps [hypreal_le_eq_less_or_eq]) 1);
+qed "hypreal_le_refl";
+
+(* Axiom 'linorder_linear' of class 'linorder': *)
+Goal "(z::hypreal) <= w | w <= z";
+by (simp_tac (simpset() addsimps [hypreal_le_less]) 1);
+by (cut_facts_tac [hypreal_linear] 1);
+by (Blast_tac 1);
+qed "hypreal_le_linear";
+
+Goal "[| i <= j; j <= k |] ==> i <= (k::hypreal)";
+by (EVERY1 [dtac hypreal_le_imp_less_or_eq, dtac hypreal_le_imp_less_or_eq,
+            rtac hypreal_less_or_eq_imp_le, 
+            fast_tac (claset() addIs [hypreal_less_trans])]);
+qed "hypreal_le_trans";
+
+Goal "[| z <= w; w <= z |] ==> z = (w::hypreal)";
+by (EVERY1 [dtac hypreal_le_imp_less_or_eq, dtac hypreal_le_imp_less_or_eq,
+            fast_tac (claset() addEs [hypreal_less_irrefl,hypreal_less_asym])]);
+qed "hypreal_le_anti_sym";
+
+Goal "[| ~ y < x; y ~= x |] ==> x < (y::hypreal)";
+by (rtac not_hypreal_leE 1);
+by (fast_tac (claset() addDs [hypreal_le_imp_less_or_eq]) 1);
+qed "not_less_not_eq_hypreal_less";
+
+(* Axiom 'order_less_le' of class 'order': *)
+Goal "(w::hypreal) < z = (w <= z & w ~= z)";
+by (simp_tac (simpset() addsimps [hypreal_le_def, hypreal_neq_iff]) 1);
+by (blast_tac (claset() addIs [hypreal_less_asym]) 1);
+qed "hypreal_less_le";
+
+Goal "(0 < -R) = (R < (0::hypreal))";
+by (res_inst_tac [("z","R")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+       simpset() addsimps [hypreal_zero_def, hypreal_less,hypreal_minus]));
+qed "hypreal_minus_zero_less_iff";
+Addsimps [hypreal_minus_zero_less_iff];
+
+Goal "(-R < 0) = ((0::hypreal) < R)";
+by (res_inst_tac [("z","R")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+         simpset() addsimps [hypreal_zero_def, hypreal_less,hypreal_minus]));
+by (ALLGOALS(Ultra_tac));
+qed "hypreal_minus_zero_less_iff2";
+
+Goalw [hypreal_le_def] "((0::hypreal) <= -r) = (r <= (0::hypreal))";
+by (simp_tac (simpset() addsimps [hypreal_minus_zero_less_iff2]) 1);
+qed "hypreal_minus_zero_le_iff";
+Addsimps [hypreal_minus_zero_le_iff];
+
+(*----------------------------------------------------------
+  hypreal_of_real preserves field and order properties
+ -----------------------------------------------------------*)
+Goalw [hypreal_of_real_def] 
+     "hypreal_of_real (z1 + z2) = hypreal_of_real z1 + hypreal_of_real z2";
+by (simp_tac (simpset() addsimps [hypreal_add, hypreal_add_mult_distrib]) 1);
+qed "hypreal_of_real_add";
+Addsimps [hypreal_of_real_add];
+
+Goalw [hypreal_of_real_def] 
+     "hypreal_of_real (z1 * z2) = hypreal_of_real z1 * hypreal_of_real z2";
+by (simp_tac (simpset() addsimps [hypreal_mult, hypreal_add_mult_distrib2]) 1);
+qed "hypreal_of_real_mult";
+Addsimps [hypreal_of_real_mult];
+
+Goalw [hypreal_less_def,hypreal_of_real_def] 
+     "(hypreal_of_real z1 <  hypreal_of_real z2) = (z1 < z2)";
+by Auto_tac;
+by (res_inst_tac [("x","%n. z1")] exI 2);
+by (Step_tac 1); 
+by (res_inst_tac [("x","%n. z2")] exI 3);
+by Auto_tac;
+by (rtac FreeUltrafilterNat_P 1);
+by (Ultra_tac 1);
+qed "hypreal_of_real_less_iff";
+Addsimps [hypreal_of_real_less_iff];
+
+Goalw [hypreal_le_def,real_le_def] 
+     "(hypreal_of_real z1 <= hypreal_of_real z2) = (z1 <= z2)";
+by Auto_tac;
+qed "hypreal_of_real_le_iff";
+Addsimps [hypreal_of_real_le_iff];
+
+Goal "(hypreal_of_real z1 = hypreal_of_real z2) = (z1 = z2)";
+by (force_tac (claset() addIs [order_antisym, hypreal_of_real_le_iff RS iffD1],
+               simpset()) 1); 
+qed "hypreal_of_real_eq_iff";
+Addsimps [hypreal_of_real_eq_iff];
+
+Goalw [hypreal_of_real_def] "hypreal_of_real (-r) = - hypreal_of_real  r";
+by (auto_tac (claset(),simpset() addsimps [hypreal_minus]));
+qed "hypreal_of_real_minus";
+Addsimps [hypreal_of_real_minus];
+
+(*DON'T insert this or the next one as default simprules.
+  They are used in both orientations and anyway aren't the ones we finally
+  need, which would use binary literals.*)
+Goalw [hypreal_of_real_def,hypreal_one_def] "hypreal_of_real  #1 = 1hr";
+by (Step_tac 1);
+qed "hypreal_of_real_one";
+
+Goalw [hypreal_of_real_def,hypreal_zero_def] "hypreal_of_real #0 = 0";
+by (Step_tac 1);
+qed "hypreal_of_real_zero";
+
+Goal "(hypreal_of_real r = 0) = (r = #0)";
+by (auto_tac (claset() addIs [FreeUltrafilterNat_P],
+    simpset() addsimps [hypreal_of_real_def,
+                        hypreal_zero_def,FreeUltrafilterNat_Nat_set]));
+qed "hypreal_of_real_zero_iff";
+
+Goal "hypreal_of_real (inverse r) = inverse (hypreal_of_real r)";
+by (case_tac "r=#0" 1);
+by (asm_simp_tac (simpset() addsimps [REAL_DIVIDE_ZERO, INVERSE_ZERO, 
+                              HYPREAL_INVERSE_ZERO, hypreal_of_real_zero]) 1);
+by (res_inst_tac [("c1","hypreal_of_real r")] 
+    (hypreal_mult_left_cancel RS iffD1) 1);
+by (stac (hypreal_of_real_mult RS sym) 2); 
+by (auto_tac (claset(), 
+         simpset() addsimps [hypreal_of_real_one, hypreal_of_real_zero_iff]));
+qed "hypreal_of_real_inverse";
+Addsimps [hypreal_of_real_inverse];
+
+Goal "hypreal_of_real (z1 / z2) = hypreal_of_real z1 / hypreal_of_real z2";
+by (simp_tac (simpset() addsimps [hypreal_divide_def, real_divide_def]) 1);
+qed "hypreal_of_real_divide"; 
+Addsimps [hypreal_of_real_divide];
+
+
+(*** Division lemmas ***)
+
+Goal "(0::hypreal)/x = 0";
+by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); 
+qed "hypreal_zero_divide";
+
+Goal "x/1hr = x";
+by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); 
+qed "hypreal_divide_one";
+Addsimps [hypreal_zero_divide, hypreal_divide_one];
+
+Goal "(x::hypreal) * (y/z) = (x*y)/z";
+by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 1); 
+qed "hypreal_times_divide1_eq";
+
+Goal "(y/z) * (x::hypreal) = (y*x)/z";
+by (simp_tac (simpset() addsimps [hypreal_divide_def]@hypreal_mult_ac) 1); 
+qed "hypreal_times_divide2_eq";
+
+Addsimps [hypreal_times_divide1_eq, hypreal_times_divide2_eq];
+
+Goal "(x::hypreal) / (y/z) = (x*z)/y";
+by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_inverse_distrib]@
+                                 hypreal_mult_ac) 1); 
+qed "hypreal_divide_divide1_eq";
+
+Goal "((x::hypreal) / y) / z = x/(y*z)";
+by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_inverse_distrib, 
+                                  hypreal_mult_assoc]) 1); 
+qed "hypreal_divide_divide2_eq";
+
+Addsimps [hypreal_divide_divide1_eq, hypreal_divide_divide2_eq];
+
+(** As with multiplication, pull minus signs OUT of the / operator **)
+
+Goal "(-x) / (y::hypreal) = - (x/y)";
+by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); 
+qed "hypreal_minus_divide_eq";
+Addsimps [hypreal_minus_divide_eq];
+
+Goal "(x / -(y::hypreal)) = - (x/y)";
+by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_minus_inverse]) 1); 
+qed "hypreal_divide_minus_eq";
+Addsimps [hypreal_divide_minus_eq];
+
+Goal "(x+y)/(z::hypreal) = x/z + y/z";
+by (simp_tac (simpset() addsimps [hypreal_divide_def, 
+                                  hypreal_add_mult_distrib]) 1); 
+qed "hypreal_add_divide_distrib";
+
+Goal "[|(x::hypreal) ~= 0;  y ~= 0 |]  \
+\     ==> inverse(x) + inverse(y) = (x + y)*inverse(x*y)";
+by (asm_full_simp_tac (simpset() addsimps [hypreal_inverse_distrib,
+                    hypreal_add_mult_distrib,hypreal_mult_assoc RS sym]) 1);
+by (stac hypreal_mult_assoc 1);
+by (rtac (hypreal_mult_left_commute RS subst) 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
+qed "hypreal_inverse_add";
+
+Goal "x = -x ==> x = (0::hypreal)";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_minus, hypreal_zero_def]));
+by (Ultra_tac 1);
+qed "hypreal_self_eq_minus_self_zero";
+
+Goal "(x + x = 0) = (x = (0::hypreal))";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_add, hypreal_zero_def]));
+qed "hypreal_add_self_zero_cancel";
+Addsimps [hypreal_add_self_zero_cancel];
+
+Goal "(x + x + y = y) = (x = (0::hypreal))";
+by Auto_tac;
+by (dtac (hypreal_eq_minus_iff RS iffD1) 1);
+by (auto_tac (claset(), 
+     simpset() addsimps [hypreal_add_assoc, hypreal_self_eq_minus_self_zero]));
+qed "hypreal_add_self_zero_cancel2";
+Addsimps [hypreal_add_self_zero_cancel2];
+
+Goal "(x + (x + y) = y) = (x = (0::hypreal))";
+by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
+qed "hypreal_add_self_zero_cancel2a";
+Addsimps [hypreal_add_self_zero_cancel2a];
+
+Goal "(b = -a) = (-b = (a::hypreal))";
+by Auto_tac;
+qed "hypreal_minus_eq_swap";
+
+Goal "(-b = -a) = (b = (a::hypreal))";
+by (asm_full_simp_tac (simpset() addsimps 
+    [hypreal_minus_eq_swap]) 1);
+qed "hypreal_minus_eq_cancel";
+Addsimps [hypreal_minus_eq_cancel];
+
+Goal "x < x + 1hr";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+              simpset() addsimps [hypreal_add, hypreal_one_def,hypreal_less]));
+qed "hypreal_less_self_add_one";
+Addsimps [hypreal_less_self_add_one];
+
+(*??DELETE MOST OF THE FOLLOWING??*)
+Goal "((x::hypreal) + x = y + y) = (x = y)";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [hypreal_add]));
+by (ALLGOALS(Ultra_tac));
+qed "hypreal_add_self_cancel";
+Addsimps [hypreal_add_self_cancel];
+
+Goal "(y = x + - y + x) = (y = (x::hypreal))";
+by Auto_tac;
+by (dres_inst_tac [("x1","y")] 
+    (hypreal_add_right_cancel RS iffD2) 1);
+by (auto_tac (claset(),simpset() addsimps hypreal_add_ac));
+qed "hypreal_add_self_minus_cancel";
+Addsimps [hypreal_add_self_minus_cancel];
+
+Goal "(y = x + (- y + x)) = (y = (x::hypreal))";
+by (asm_full_simp_tac (simpset() addsimps 
+         [hypreal_add_assoc RS sym])1);
+qed "hypreal_add_self_minus_cancel2";
+Addsimps [hypreal_add_self_minus_cancel2];
+
+(* of course, can prove this by "transfer" as well *)
+Goal "z + -x = y + (y + (-x + -z)) = (y = (z::hypreal))";
+by Auto_tac;
+by (dres_inst_tac [("x1","z")] 
+    (hypreal_add_right_cancel RS iffD2) 1);
+by (asm_full_simp_tac (simpset() addsimps 
+    [hypreal_minus_add_distrib RS sym] @ hypreal_add_ac
+    delsimps [hypreal_minus_add_distrib]) 1);
+by (asm_full_simp_tac (simpset() addsimps 
+     [hypreal_add_assoc RS sym,hypreal_add_right_cancel]) 1);
+qed "hypreal_add_self_minus_cancel3";
+Addsimps [hypreal_add_self_minus_cancel3];
+
+Goalw [hypreal_diff_def] "(x<y) = (x-y < (0::hypreal))";
+by (rtac hypreal_less_minus_iff2 1);
+qed "hypreal_less_eq_diff";
+
+(*** Subtraction laws ***)
+
+Goal "x + (y - z) = (x + y) - (z::hypreal)";
+by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
+qed "hypreal_add_diff_eq";
+
+Goal "(x - y) + z = (x + z) - (y::hypreal)";
+by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
+qed "hypreal_diff_add_eq";
+
+Goal "(x - y) - z = x - (y + (z::hypreal))";
+by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
+qed "hypreal_diff_diff_eq";
+
+Goal "x - (y - z) = (x + z) - (y::hypreal)";
+by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
+qed "hypreal_diff_diff_eq2";
+
+Goal "(x-y < z) = (x < z + (y::hypreal))";
+by (stac hypreal_less_eq_diff 1);
+by (res_inst_tac [("y1", "z")] (hypreal_less_eq_diff RS ssubst) 1);
+by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
+qed "hypreal_diff_less_eq";
+
+Goal "(x < z-y) = (x + (y::hypreal) < z)";
+by (stac hypreal_less_eq_diff 1);
+by (res_inst_tac [("y1", "z-y")] (hypreal_less_eq_diff RS ssubst) 1);
+by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
+qed "hypreal_less_diff_eq";
+
+Goalw [hypreal_le_def] "(x-y <= z) = (x <= z + (y::hypreal))";
+by (simp_tac (simpset() addsimps [hypreal_less_diff_eq]) 1);
+qed "hypreal_diff_le_eq";
+
+Goalw [hypreal_le_def] "(x <= z-y) = (x + (y::hypreal) <= z)";
+by (simp_tac (simpset() addsimps [hypreal_diff_less_eq]) 1);
+qed "hypreal_le_diff_eq";
+
+Goalw [hypreal_diff_def] "(x-y = z) = (x = z + (y::hypreal))";
+by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc]));
+qed "hypreal_diff_eq_eq";
+
+Goalw [hypreal_diff_def] "(x = z-y) = (x + (y::hypreal) = z)";
+by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc]));
+qed "hypreal_eq_diff_eq";
+
+(*This list of rewrites simplifies (in)equalities by bringing subtractions
+  to the top and then moving negative terms to the other side.  
+  Use with hypreal_add_ac*)
+val hypreal_compare_rls = 
+  [symmetric hypreal_diff_def,
+   hypreal_add_diff_eq, hypreal_diff_add_eq, hypreal_diff_diff_eq, 
+   hypreal_diff_diff_eq2, 
+   hypreal_diff_less_eq, hypreal_less_diff_eq, hypreal_diff_le_eq, 
+   hypreal_le_diff_eq, hypreal_diff_eq_eq, hypreal_eq_diff_eq];
+
+
+(** For the cancellation simproc.
+    The idea is to cancel like terms on opposite sides by subtraction **)
+
+Goal "(x::hypreal) - y = x' - y' ==> (x<y) = (x'<y')";
+by (stac hypreal_less_eq_diff 1);
+by (res_inst_tac [("y1", "y")] (hypreal_less_eq_diff RS ssubst) 1);
+by (Asm_simp_tac 1);
+qed "hypreal_less_eqI";
+
+Goal "(x::hypreal) - y = x' - y' ==> (y<=x) = (y'<=x')";
+by (dtac hypreal_less_eqI 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_le_def]) 1);
+qed "hypreal_le_eqI";
+
+Goal "(x::hypreal) - y = x' - y' ==> (x=y) = (x'=y')";
+by Safe_tac;
+by (ALLGOALS
+    (asm_full_simp_tac
+     (simpset() addsimps [hypreal_eq_diff_eq, hypreal_diff_eq_eq])));
+qed "hypreal_eq_eqI";
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HyperDef.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,98 @@
+(*  Title       : HOL/Real/Hyperreal/HyperDef.thy
+    ID          : $Id$
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Construction of hyperreals using ultrafilters
+*) 
+
+HyperDef = Filter + Real +
+
+consts 
+ 
+    FreeUltrafilterNat   :: nat set set    ("\\<U>")
+
+defs
+
+    FreeUltrafilterNat_def
+    "FreeUltrafilterNat    ==   (@U. U : FreeUltrafilter (UNIV:: nat set))"
+
+
+constdefs
+    hyprel :: "((nat=>real)*(nat=>real)) set"
+    "hyprel == {p. ? X Y. p = ((X::nat=>real),Y) & 
+                   {n::nat. X(n) = Y(n)}: FreeUltrafilterNat}"
+
+typedef hypreal = "UNIV//hyprel"   (Equiv.quotient_def)
+
+instance
+   hypreal  :: {ord, zero, plus, times, minus, inverse}
+
+consts 
+
+  "1hr"       :: hypreal               ("1hr")  
+  "whr"       :: hypreal               ("whr")  
+  "ehr"       :: hypreal               ("ehr")  
+
+
+defs
+
+  hypreal_zero_def
+  "0 == Abs_hypreal(hyprel^^{%n::nat. (#0::real)})"
+
+  hypreal_one_def
+  "1hr == Abs_hypreal(hyprel^^{%n::nat. (#1::real)})"
+
+  (* an infinite number = [<1,2,3,...>] *)
+  omega_def
+  "whr == Abs_hypreal(hyprel^^{%n::nat. real_of_posnat n})"
+    
+  (* an infinitesimal number = [<1,1/2,1/3,...>] *)
+  epsilon_def
+  "ehr == Abs_hypreal(hyprel^^{%n::nat. inverse (real_of_posnat n)})"
+
+  hypreal_minus_def
+  "- P == Abs_hypreal(UN X: Rep_hypreal(P). hyprel^^{%n::nat. - (X n)})"
+
+  hypreal_diff_def 
+  "x - y == x + -(y::hypreal)"
+
+  hypreal_inverse_def
+  "inverse P == Abs_hypreal(UN X: Rep_hypreal(P). 
+                    hyprel^^{%n. if X n = #0 then #0 else inverse (X n)})"
+
+  hypreal_divide_def
+  "P / Q::hypreal == P * inverse Q"
+  
+constdefs
+
+  hypreal_of_real  :: real => hypreal                 
+  "hypreal_of_real r         == Abs_hypreal(hyprel^^{%n::nat. r})"
+  
+  (* n::nat --> (n+1)::hypreal *)
+  hypreal_of_posnat :: nat => hypreal                
+  "hypreal_of_posnat n  == (hypreal_of_real(real_of_preal
+                            (preal_of_prat(prat_of_pnat(pnat_of_nat n)))))"
+
+  hypreal_of_nat :: nat => hypreal                   
+  "hypreal_of_nat n      == hypreal_of_posnat n + -1hr"
+
+defs 
+
+  hypreal_add_def  
+  "P + Q == Abs_hypreal(UN X:Rep_hypreal(P). UN Y:Rep_hypreal(Q).
+                hyprel^^{%n::nat. X n + Y n})"
+
+  hypreal_mult_def  
+  "P * Q == Abs_hypreal(UN X:Rep_hypreal(P). UN Y:Rep_hypreal(Q).
+                hyprel^^{%n::nat. X n * Y n})"
+
+  hypreal_less_def
+  "P < (Q::hypreal) == EX X Y. X: Rep_hypreal(P) & 
+                               Y: Rep_hypreal(Q) & 
+                               {n::nat. X n < Y n} : FreeUltrafilterNat"
+  hypreal_le_def
+  "P <= (Q::hypreal) == ~(Q < P)" 
+
+end
+ 
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HyperNat.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,1335 @@
+(*  Title       : HyperNat.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Explicit construction of hypernaturals using 
+                  ultrafilters
+*) 
+       
+(*------------------------------------------------------------------------
+                       Properties of hypnatrel
+ ------------------------------------------------------------------------*)
+
+(** Proving that hyprel is an equivalence relation       **)
+(** Natural deduction for hypnatrel - similar to hyprel! **)
+
+Goalw [hypnatrel_def]
+     "((X,Y): hypnatrel) = ({n. X n = Y n}: FreeUltrafilterNat)";
+by (Fast_tac 1);
+qed "hypnatrel_iff";
+
+Goalw [hypnatrel_def] 
+     "{n. X n = Y n}: FreeUltrafilterNat ==> (X,Y): hypnatrel";
+by (Fast_tac 1);
+qed "hypnatrelI";
+
+Goalw [hypnatrel_def]
+  "p: hypnatrel --> (EX X Y. \
+\                 p = (X,Y) & {n. X n = Y n} : FreeUltrafilterNat)";
+by (Fast_tac 1);
+qed "hypnatrelE_lemma";
+
+val [major,minor] = Goal
+  "[| p: hypnatrel;  \
+\     !!X Y. [| p = (X,Y); {n. X n = Y n}: FreeUltrafilterNat |] \
+\            ==> Q |] \
+\  ==> Q";
+by (cut_facts_tac [major RS (hypnatrelE_lemma RS mp)] 1);
+by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
+qed "hypnatrelE";
+
+AddSIs [hypnatrelI];
+AddSEs [hypnatrelE];
+
+Goalw [hypnatrel_def] "(x,x): hypnatrel";
+by (Auto_tac);
+qed "hypnatrel_refl";
+
+Goalw [hypnatrel_def] "(x,y): hypnatrel --> (y,x):hypnatrel";
+by (auto_tac (claset() addIs [lemma_perm RS subst], simpset()));
+qed_spec_mp "hypnatrel_sym";
+
+Goalw [hypnatrel_def]
+     "(x,y): hypnatrel --> (y,z):hypnatrel --> (x,z):hypnatrel";
+by (Auto_tac);
+by (Fuf_tac 1);
+qed_spec_mp "hypnatrel_trans";
+
+Goalw [equiv_def, refl_def, sym_def, trans_def]
+     "equiv UNIV hypnatrel";
+by (auto_tac (claset() addSIs [hypnatrel_refl] 
+                       addSEs [hypnatrel_sym,hypnatrel_trans] 
+                       delrules [hypnatrelI,hypnatrelE],
+              simpset()));
+qed "equiv_hypnatrel";
+
+val equiv_hypnatrel_iff =
+    [UNIV_I, UNIV_I] MRS (equiv_hypnatrel RS eq_equiv_class_iff);
+
+Goalw  [hypnat_def,hypnatrel_def,quotient_def] "hypnatrel^^{x}:hypnat";
+by (Blast_tac 1);
+qed "hypnatrel_in_hypnat";
+
+Goal "inj_on Abs_hypnat hypnat";
+by (rtac inj_on_inverseI 1);
+by (etac Abs_hypnat_inverse 1);
+qed "inj_on_Abs_hypnat";
+
+Addsimps [equiv_hypnatrel_iff,inj_on_Abs_hypnat RS inj_on_iff,
+          hypnatrel_iff, hypnatrel_in_hypnat, Abs_hypnat_inverse];
+
+Addsimps [equiv_hypnatrel RS eq_equiv_class_iff];
+val eq_hypnatrelD = equiv_hypnatrel RSN (2,eq_equiv_class);
+
+Goal "inj(Rep_hypnat)";
+by (rtac inj_inverseI 1);
+by (rtac Rep_hypnat_inverse 1);
+qed "inj_Rep_hypnat";
+
+Goalw [hypnatrel_def] "x: hypnatrel ^^ {x}";
+by (Step_tac 1);
+by (Auto_tac);
+qed "lemma_hypnatrel_refl";
+
+Addsimps [lemma_hypnatrel_refl];
+
+Goalw [hypnat_def] "{} ~: hypnat";
+by (auto_tac (claset() addSEs [quotientE],simpset()));
+qed "hypnat_empty_not_mem";
+
+Addsimps [hypnat_empty_not_mem];
+
+Goal "Rep_hypnat x ~= {}";
+by (cut_inst_tac [("x","x")] Rep_hypnat 1);
+by (Auto_tac);
+qed "Rep_hypnat_nonempty";
+
+Addsimps [Rep_hypnat_nonempty];
+
+(*------------------------------------------------------------------------
+   hypnat_of_nat: the injection from nat to hypnat
+ ------------------------------------------------------------------------*)
+Goal "inj(hypnat_of_nat)";
+by (rtac injI 1);
+by (rewtac hypnat_of_nat_def);
+by (dtac (inj_on_Abs_hypnat RS inj_onD) 1);
+by (REPEAT (rtac hypnatrel_in_hypnat 1));
+by (dtac eq_equiv_class 1);
+by (rtac equiv_hypnatrel 1);
+by (Fast_tac 1);
+by (rtac ccontr 1 THEN rotate_tac 1 1);
+by (Auto_tac);
+qed "inj_hypnat_of_nat";
+
+val [prem] = Goal
+    "(!!x. z = Abs_hypnat(hypnatrel^^{x}) ==> P) ==> P";
+by (res_inst_tac [("x1","z")] 
+    (rewrite_rule [hypnat_def] Rep_hypnat RS quotientE) 1);
+by (dres_inst_tac [("f","Abs_hypnat")] arg_cong 1);
+by (res_inst_tac [("x","x")] prem 1);
+by (asm_full_simp_tac (simpset() addsimps [Rep_hypnat_inverse]) 1);
+qed "eq_Abs_hypnat";
+
+(*-----------------------------------------------------------
+   Addition for hyper naturals: hypnat_add 
+ -----------------------------------------------------------*)
+Goalw [congruent2_def]
+     "congruent2 hypnatrel (%X Y. hypnatrel^^{%n. X n + Y n})";
+by Safe_tac;
+by (ALLGOALS(Fuf_tac));
+qed "hypnat_add_congruent2";
+
+Goalw [hypnat_add_def]
+  "Abs_hypnat(hypnatrel^^{%n. X n}) + Abs_hypnat(hypnatrel^^{%n. Y n}) = \
+\  Abs_hypnat(hypnatrel^^{%n. X n + Y n})";
+by (asm_simp_tac
+    (simpset() addsimps [[equiv_hypnatrel, hypnat_add_congruent2] 
+     MRS UN_equiv_class2]) 1);
+qed "hypnat_add";
+
+Goal "(z::hypnat) + w = w + z";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","w")] eq_Abs_hypnat 1);
+by (asm_simp_tac (simpset() addsimps (add_ac @ [hypnat_add])) 1);
+qed "hypnat_add_commute";
+
+Goal "((z1::hypnat) + z2) + z3 = z1 + (z2 + z3)";
+by (res_inst_tac [("z","z1")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","z2")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","z3")] eq_Abs_hypnat 1);
+by (asm_simp_tac (simpset() addsimps [hypnat_add,add_assoc]) 1);
+qed "hypnat_add_assoc";
+
+(*For AC rewriting*)
+Goal "(x::hypnat)+(y+z)=y+(x+z)";
+by (rtac (hypnat_add_commute RS trans) 1);
+by (rtac (hypnat_add_assoc RS trans) 1);
+by (rtac (hypnat_add_commute RS arg_cong) 1);
+qed "hypnat_add_left_commute";
+
+(* hypnat addition is an AC operator *)
+val hypnat_add_ac = [hypnat_add_assoc,hypnat_add_commute,
+                      hypnat_add_left_commute];
+
+Goalw [hypnat_zero_def] "(0::hypnat) + z = z";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (asm_full_simp_tac (simpset() addsimps [hypnat_add]) 1);
+qed "hypnat_add_zero_left";
+
+Goal "z + (0::hypnat) = z";
+by (simp_tac (simpset() addsimps 
+    [hypnat_add_zero_left,hypnat_add_commute]) 1);
+qed "hypnat_add_zero_right";
+
+Addsimps [hypnat_add_zero_left,hypnat_add_zero_right];
+
+(*-----------------------------------------------------------
+   Subtraction for hyper naturals: hypnat_minus
+ -----------------------------------------------------------*)
+Goalw [congruent2_def]
+    "congruent2 hypnatrel (%X Y. hypnatrel^^{%n. X n - Y n})";
+by Safe_tac;
+by (ALLGOALS(Fuf_tac));
+qed "hypnat_minus_congruent2";
+ 
+Goalw [hypnat_minus_def]
+  "Abs_hypnat(hypnatrel^^{%n. X n}) - Abs_hypnat(hypnatrel^^{%n. Y n}) = \
+\  Abs_hypnat(hypnatrel^^{%n. X n - Y n})";
+by (asm_simp_tac
+    (simpset() addsimps [[equiv_hypnatrel, hypnat_minus_congruent2] 
+     MRS UN_equiv_class2]) 1);
+qed "hypnat_minus";
+
+Goalw [hypnat_zero_def] "z - z = (0::hypnat)";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (asm_full_simp_tac (simpset() addsimps [hypnat_minus]) 1);
+qed "hypnat_minus_zero";
+
+Goalw [hypnat_zero_def] "(0::hypnat) - n = 0";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (asm_full_simp_tac (simpset() addsimps [hypnat_minus]) 1);
+qed "hypnat_diff_0_eq_0";
+
+Addsimps [hypnat_minus_zero,hypnat_diff_0_eq_0];
+
+Goalw [hypnat_zero_def] "(m+n = (0::hypnat)) = (m=0 & n=0)";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (auto_tac (claset() addIs [FreeUltrafilterNat_subset]
+                       addDs [FreeUltrafilterNat_Int],
+              simpset() addsimps [hypnat_add] ));
+qed "hypnat_add_is_0";
+
+AddIffs [hypnat_add_is_0];
+
+Goal "!!i::hypnat. i-j-k = i - (j+k)";
+by (res_inst_tac [("z","i")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","j")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","k")] eq_Abs_hypnat 1);
+by (asm_full_simp_tac (simpset() addsimps 
+    [hypnat_minus,hypnat_add,diff_diff_left]) 1);
+qed "hypnat_diff_diff_left";
+
+Goal "!!i::hypnat. i-j-k = i-k-j";
+by (simp_tac (simpset() addsimps 
+    [hypnat_diff_diff_left, hypnat_add_commute]) 1);
+qed "hypnat_diff_commute";
+
+Goal "!!n::hypnat. (n+m) - n = m";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (asm_full_simp_tac (simpset() addsimps [hypnat_minus,hypnat_add]) 1);
+qed "hypnat_diff_add_inverse";
+Addsimps [hypnat_diff_add_inverse];
+
+Goal "!!n::hypnat.(m+n) - n = m";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (asm_full_simp_tac (simpset() addsimps [hypnat_minus,hypnat_add]) 1);
+qed "hypnat_diff_add_inverse2";
+Addsimps [hypnat_diff_add_inverse2];
+
+Goal "!!k::hypnat. (k+m) - (k+n) = m - n";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","k")] eq_Abs_hypnat 1);
+by (asm_full_simp_tac (simpset() addsimps [hypnat_minus,hypnat_add]) 1);
+qed "hypnat_diff_cancel";
+Addsimps [hypnat_diff_cancel];
+
+Goal "!!m::hypnat. (m+k) - (n+k) = m - n";
+val hypnat_add_commute_k = read_instantiate [("w","k")] hypnat_add_commute;
+by (asm_simp_tac (simpset() addsimps ([hypnat_add_commute_k])) 1);
+qed "hypnat_diff_cancel2";
+Addsimps [hypnat_diff_cancel2];
+
+Goalw [hypnat_zero_def] "!!n::hypnat. n - (n+m) = (0::hypnat)";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (asm_full_simp_tac (simpset() addsimps [hypnat_minus,hypnat_add]) 1);
+qed "hypnat_diff_add_0";
+Addsimps [hypnat_diff_add_0];
+
+(*-----------------------------------------------------------
+   Multiplication for hyper naturals: hypnat_mult
+ -----------------------------------------------------------*)
+Goalw [congruent2_def]
+    "congruent2 hypnatrel (%X Y. hypnatrel^^{%n. X n * Y n})";
+by Safe_tac;
+by (ALLGOALS(Fuf_tac));
+qed "hypnat_mult_congruent2";
+
+Goalw [hypnat_mult_def]
+  "Abs_hypnat(hypnatrel^^{%n. X n}) * Abs_hypnat(hypnatrel^^{%n. Y n}) = \
+\  Abs_hypnat(hypnatrel^^{%n. X n * Y n})";
+by (asm_simp_tac
+    (simpset() addsimps [[equiv_hypnatrel,hypnat_mult_congruent2] MRS
+     UN_equiv_class2]) 1);
+qed "hypnat_mult";
+
+Goal "(z::hypnat) * w = w * z";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","w")] eq_Abs_hypnat 1);
+by (asm_simp_tac (simpset() addsimps ([hypnat_mult] @ mult_ac)) 1);
+qed "hypnat_mult_commute";
+
+Goal "((z1::hypnat) * z2) * z3 = z1 * (z2 * z3)";
+by (res_inst_tac [("z","z1")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","z2")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","z3")] eq_Abs_hypnat 1);
+by (asm_simp_tac (simpset() addsimps [hypnat_mult,mult_assoc]) 1);
+qed "hypnat_mult_assoc";
+
+
+Goal "(z1::hypnat) * (z2 * z3) = z2 * (z1 * z3)";
+by (rtac (hypnat_mult_commute RS trans) 1);
+by (rtac (hypnat_mult_assoc RS trans) 1);
+by (rtac (hypnat_mult_commute RS arg_cong) 1);
+qed "hypnat_mult_left_commute";
+
+(* hypnat multiplication is an AC operator *)
+val hypnat_mult_ac = [hypnat_mult_assoc, hypnat_mult_commute, 
+                      hypnat_mult_left_commute];
+
+Goalw [hypnat_one_def] "1hn * z = z";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (asm_full_simp_tac (simpset() addsimps [hypnat_mult]) 1);
+qed "hypnat_mult_1";
+Addsimps [hypnat_mult_1];
+
+Goal "z * 1hn = z";
+by (simp_tac (simpset() addsimps [hypnat_mult_commute]) 1);
+qed "hypnat_mult_1_right";
+Addsimps [hypnat_mult_1_right];
+
+Goalw [hypnat_zero_def] "(0::hypnat) * z = 0";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (asm_full_simp_tac (simpset() addsimps [hypnat_mult]) 1);
+qed "hypnat_mult_0";
+Addsimps [hypnat_mult_0];
+
+Goal "z * (0::hypnat) = 0";
+by (simp_tac (simpset() addsimps [hypnat_mult_commute]) 1);
+qed "hypnat_mult_0_right";
+Addsimps [hypnat_mult_0_right];
+
+Goal "!!m::hypnat. (m - n) * k = (m * k) - (n * k)";
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","k")] eq_Abs_hypnat 1);
+by (asm_simp_tac (simpset() addsimps [hypnat_mult,
+    hypnat_minus,diff_mult_distrib]) 1);
+qed "hypnat_diff_mult_distrib" ;
+
+Goal "!!m::hypnat. k * (m - n) = (k * m) - (k * n)";
+val hypnat_mult_commute_k = read_instantiate [("z","k")] hypnat_mult_commute;
+by (simp_tac (simpset() addsimps [hypnat_diff_mult_distrib, 
+                                  hypnat_mult_commute_k]) 1);
+qed "hypnat_diff_mult_distrib2" ;
+
+(*--------------------------
+    A Few more theorems
+ -------------------------*)
+
+Goal "(z::hypnat) + v = z' + v' ==> z + (v + w) = z' + (v' + w)";
+by (asm_simp_tac (simpset() addsimps [hypnat_add_assoc RS sym]) 1);
+qed "hypnat_add_assoc_cong";
+
+Goal "(z::hypnat) + (v + w) = v + (z + w)";
+by (REPEAT (ares_tac [hypnat_add_commute RS hypnat_add_assoc_cong] 1));
+qed "hypnat_add_assoc_swap";
+
+Goal "((z1::hypnat) + z2) * w = (z1 * w) + (z2 * w)";
+by (res_inst_tac [("z","z1")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","z2")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","w")] eq_Abs_hypnat 1);
+by (asm_simp_tac (simpset() addsimps [hypnat_mult,hypnat_add,
+                                      add_mult_distrib]) 1);
+qed "hypnat_add_mult_distrib";
+Addsimps [hypnat_add_mult_distrib];
+
+val hypnat_mult_commute'= read_instantiate [("z","w")] hypnat_mult_commute;
+
+Goal "(w::hypnat) * (z1 + z2) = (w * z1) + (w * z2)";
+by (simp_tac (simpset() addsimps [hypnat_mult_commute']) 1);
+qed "hypnat_add_mult_distrib2";
+Addsimps [hypnat_add_mult_distrib2];
+
+(*** (hypnat) one and zero are distinct ***)
+Goalw [hypnat_zero_def,hypnat_one_def] "(0::hypnat) ~= 1hn";
+by (Auto_tac);
+qed "hypnat_zero_not_eq_one";
+Addsimps [hypnat_zero_not_eq_one];
+Addsimps [hypnat_zero_not_eq_one RS not_sym];
+
+Goalw [hypnat_zero_def] "(m*n = (0::hypnat)) = (m=0 | n=0)";
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem],
+    simpset() addsimps [hypnat_mult]));
+by (ALLGOALS(Fuf_tac));
+qed "hypnat_mult_is_0";
+Addsimps [hypnat_mult_is_0];
+
+(*------------------------------------------------------------------
+                   Theorems for ordering 
+ ------------------------------------------------------------------*)
+
+(* prove introduction and elimination rules for hypnat_less *)
+
+Goalw [hypnat_less_def]
+ "P < (Q::hypnat) = (EX X Y. X : Rep_hypnat(P) & \
+\                             Y : Rep_hypnat(Q) & \
+\                             {n. X n < Y n} : FreeUltrafilterNat)";
+by (Fast_tac 1);
+qed "hypnat_less_iff";
+
+Goalw [hypnat_less_def]
+ "!!P. [| {n. X n < Y n} : FreeUltrafilterNat; \
+\         X : Rep_hypnat(P); \
+\         Y : Rep_hypnat(Q) |] ==> P < (Q::hypnat)";
+by (Fast_tac 1);
+qed "hypnat_lessI";
+
+Goalw [hypnat_less_def]
+     "!! R1. [| R1 < (R2::hypnat); \
+\         !!X Y. {n. X n < Y n} : FreeUltrafilterNat ==> P; \
+\         !!X. X : Rep_hypnat(R1) ==> P; \ 
+\         !!Y. Y : Rep_hypnat(R2) ==> P |] \
+\     ==> P";
+by (Auto_tac);
+qed "hypnat_lessE";
+
+Goalw [hypnat_less_def]
+ "!!R1. R1 < (R2::hypnat) ==> (EX X Y. {n. X n < Y n} : FreeUltrafilterNat & \
+\                                  X : Rep_hypnat(R1) & \
+\                                  Y : Rep_hypnat(R2))";
+by (Fast_tac 1);
+qed "hypnat_lessD";
+
+Goal "~ (R::hypnat) < R";
+by (res_inst_tac [("z","R")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),simpset() addsimps [hypnat_less_def]));
+by (Fuf_empty_tac 1);
+qed "hypnat_less_not_refl";
+Addsimps [hypnat_less_not_refl];
+
+bind_thm("hypnat_less_irrefl",hypnat_less_not_refl RS notE);
+
+Goalw [hypnat_less_def,hypnat_zero_def] "~ n<(0::hypnat)";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (Auto_tac);
+by (Fuf_empty_tac 1);
+qed "hypnat_not_less0";
+AddIffs [hypnat_not_less0];
+
+(* n<(0::hypnat) ==> R *)
+bind_thm ("hypnat_less_zeroE", hypnat_not_less0 RS notE);
+
+Goalw [hypnat_less_def,hypnat_zero_def,hypnat_one_def]
+      "(n<1hn) = (n=0)";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset() addSIs [exI] addEs 
+    [FreeUltrafilterNat_subset],simpset()));
+by (Fuf_tac 1);
+qed "hypnat_less_one";
+AddIffs [hypnat_less_one];
+
+Goal "!!(R1::hypnat). [| R1 < R2; R2 < R3 |] ==> R1 < R3";
+by (res_inst_tac [("z","R1")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","R2")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","R3")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),simpset() addsimps [hypnat_less_def]));
+by (res_inst_tac [("x","X")] exI 1);
+by (Auto_tac);
+by (res_inst_tac [("x","Ya")] exI 1);
+by (Auto_tac);
+by (Fuf_tac 1);
+qed "hypnat_less_trans";
+
+Goal "!! (R1::hypnat). [| R1 < R2; R2 < R1 |] ==> P";
+by (dtac hypnat_less_trans 1 THEN assume_tac 1);
+by (Asm_full_simp_tac 1);
+qed "hypnat_less_asym";
+
+(*----- hypnat less iff less a.e -----*)
+(* See comments in HYPER for corresponding thm *)
+
+Goalw [hypnat_less_def]
+      "(Abs_hypnat(hypnatrel^^{%n. X n}) < \
+\           Abs_hypnat(hypnatrel^^{%n. Y n})) = \
+\      ({n. X n < Y n} : FreeUltrafilterNat)";
+by (auto_tac (claset() addSIs [lemma_hypnatrel_refl],simpset()));
+by (Fuf_tac 1);
+qed "hypnat_less";
+
+Goal "~ m<n --> n+(m-n) = (m::hypnat)";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [hypnat_minus,hypnat_add,hypnat_less]));
+by (dtac FreeUltrafilterNat_Compl_mem 1);
+by (Fuf_tac 1);
+qed_spec_mp "hypnat_add_diff_inverse";
+
+Goal "n<=m ==> n+(m-n) = (m::hypnat)";
+by (asm_full_simp_tac (simpset() addsimps 
+    [hypnat_add_diff_inverse, hypnat_le_def]) 1);
+qed "hypnat_le_add_diff_inverse";
+
+Goal "n<=m ==> (m-n)+n = (m::hypnat)";
+by (asm_simp_tac (simpset() addsimps [hypnat_le_add_diff_inverse, 
+    hypnat_add_commute]) 1);
+qed "hypnat_le_add_diff_inverse2";
+
+(*---------------------------------------------------------------------------------
+                    Hyper naturals as a linearly ordered field
+ ---------------------------------------------------------------------------------*)
+Goalw [hypnat_zero_def] 
+     "[| (0::hypnat) < x; 0 < y |] ==> 0 < x + y";
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),simpset() addsimps
+   [hypnat_less_def,hypnat_add]));
+by (REPEAT(Step_tac 1));
+by (Fuf_tac 1);
+qed "hypnat_add_order";
+
+Goalw [hypnat_zero_def] 
+      "!!(x::hypnat). [| (0::hypnat) < x; 0 < y |] ==> 0 < x * y";
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [hypnat_less_def,hypnat_mult]));
+by (REPEAT(Step_tac 1));
+by (Fuf_tac 1);
+qed "hypnat_mult_order";
+
+(*---------------------------------------------------------------------------------
+                   Trichotomy of the hyper naturals
+  --------------------------------------------------------------------------------*)
+Goalw [hypnatrel_def] "EX x. x: hypnatrel ^^ {%n. 0}";
+by (res_inst_tac [("x","%n. 0")] exI 1);
+by (Step_tac 1);
+by (Auto_tac);
+qed "lemma_hypnatrel_0_mem";
+
+(* linearity is actually proved further down! *)
+Goalw [hypnat_zero_def,
+       hypnat_less_def]"(0::hypnat) <  x | x = 0 | x < 0";
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (Auto_tac );
+by (REPEAT(Step_tac 1));
+by (REPEAT(dtac FreeUltrafilterNat_Compl_mem 1));
+by (Fuf_tac 1);
+qed "hypnat_trichotomy";
+
+Goal "!!x. [| (0::hypnat) < x ==> P; \
+\                 x = 0 ==> P; \
+\                 x < 0 ==> P |] ==> P";
+by (cut_inst_tac [("x","x")] hypnat_trichotomy 1);
+by (Auto_tac);
+qed "hypnat_trichotomyE";
+
+(*------------------------------------------------------------------------------
+            More properties of <
+ ------------------------------------------------------------------------------*)
+Goal "!!(A::hypnat). A < B ==> A + C < B + C";
+by (res_inst_tac [("z","A")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","B")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","C")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [hypnat_less_def,hypnat_add]));
+by (REPEAT(Step_tac 1));
+by (Fuf_tac 1);
+qed "hypnat_add_less_mono1";
+
+Goal "!!(A::hypnat). A < B ==> C + A < C + B";
+by (auto_tac (claset() addIs [hypnat_add_less_mono1],
+    simpset() addsimps [hypnat_add_commute]));
+qed "hypnat_add_less_mono2";
+
+Goal "!!k l::hypnat. [|i<j;  k<l |] ==> i + k < j + l";
+by (etac (hypnat_add_less_mono1 RS hypnat_less_trans) 1);
+by (simp_tac (simpset() addsimps [hypnat_add_commute]) 1);
+(*j moves to the end because it is free while k, l are bound*)
+by (etac hypnat_add_less_mono1 1);
+qed "hypnat_add_less_mono";
+
+(*---------------------------------------
+        hypnat_minus_less
+ ---------------------------------------*)
+Goalw [hypnat_less_def,hypnat_zero_def] 
+      "((x::hypnat) < y) = ((0::hypnat) < y - x)";
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [hypnat_minus]));
+by (REPEAT(Step_tac 1));
+by (REPEAT(Step_tac 2));
+by (ALLGOALS(fuf_tac (claset() addDs [sym],simpset())));
+
+(*** linearity ***)
+Goalw [hypnat_less_def] "(x::hypnat) < y | x = y | y < x";
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
+by (Auto_tac );
+by (REPEAT(Step_tac 1));
+by (REPEAT(dtac FreeUltrafilterNat_Compl_mem 1));
+by (Fuf_tac 1);
+qed "hypnat_linear";
+
+Goal
+    "!!(x::hypnat). [| x < y ==> P;  x = y ==> P; \
+\          y < x ==> P |] ==> P";
+by (cut_inst_tac [("x","x"),("y","y")] hypnat_linear 1);
+by (Auto_tac);
+qed "hypnat_linear_less2";
+
+(*------------------------------------------------------------------------------
+                            Properties of <=
+ ------------------------------------------------------------------------------*)
+(*------ hypnat le iff nat le a.e ------*)
+Goalw [hypnat_le_def,le_def]
+      "(Abs_hypnat(hypnatrel^^{%n. X n}) <= \
+\           Abs_hypnat(hypnatrel^^{%n. Y n})) = \
+\      ({n. X n <= Y n} : FreeUltrafilterNat)";
+by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem],
+    simpset() addsimps [hypnat_less]));
+by (Fuf_tac 1 THEN Fuf_empty_tac 1);
+qed "hypnat_le";
+
+(*---------------------------------------------------------*)
+(*---------------------------------------------------------*)
+Goalw [hypnat_le_def] "!!w. ~(w < z) ==> z <= (w::hypnat)";
+by (assume_tac 1);
+qed "hypnat_leI";
+
+Goalw [hypnat_le_def] "!!w. z<=w ==> ~(w<(z::hypnat))";
+by (assume_tac 1);
+qed "hypnat_leD";
+
+val hypnat_leE = make_elim hypnat_leD;
+
+Goal "!!w. (~(w < z)) = (z <= (w::hypnat))";
+by (fast_tac (claset() addSIs [hypnat_leI,hypnat_leD]) 1);
+qed "hypnat_less_le_iff";
+
+Goalw [hypnat_le_def] "!!z. ~ z <= w ==> w<(z::hypnat)";
+by (Fast_tac 1);
+qed "not_hypnat_leE";
+
+Goalw [hypnat_le_def] "!!z. z < w ==> z <= (w::hypnat)";
+by (fast_tac (claset() addEs [hypnat_less_asym]) 1);
+qed "hypnat_less_imp_le";
+
+Goalw [hypnat_le_def] "!!(x::hypnat). x <= y ==> x < y | x = y";
+by (cut_facts_tac [hypnat_linear] 1);
+by (fast_tac (claset() addEs [hypnat_less_irrefl,hypnat_less_asym]) 1);
+qed "hypnat_le_imp_less_or_eq";
+
+Goalw [hypnat_le_def] "!!z. z<w | z=w ==> z <=(w::hypnat)";
+by (cut_facts_tac [hypnat_linear] 1);
+by (fast_tac (claset() addEs [hypnat_less_irrefl,hypnat_less_asym]) 1);
+qed "hypnat_less_or_eq_imp_le";
+
+Goal "(x <= (y::hypnat)) = (x < y | x=y)";
+by (REPEAT(ares_tac [iffI, hypnat_less_or_eq_imp_le, hypnat_le_imp_less_or_eq] 1));
+qed "hypnat_le_eq_less_or_eq";
+
+Goal "w <= (w::hypnat)";
+by (simp_tac (simpset() addsimps [hypnat_le_eq_less_or_eq]) 1);
+qed "hypnat_le_refl";
+Addsimps [hypnat_le_refl];
+
+val prems = goal thy "!!i. [| i <= j; j < k |] ==> i < (k::hypnat)";
+by (dtac hypnat_le_imp_less_or_eq 1);
+by (fast_tac (claset() addIs [hypnat_less_trans]) 1);
+qed "hypnat_le_less_trans";
+
+Goal "!! (i::hypnat). [| i < j; j <= k |] ==> i < k";
+by (dtac hypnat_le_imp_less_or_eq 1);
+by (fast_tac (claset() addIs [hypnat_less_trans]) 1);
+qed "hypnat_less_le_trans";
+
+Goal "!!i. [| i <= j; j <= k |] ==> i <= (k::hypnat)";
+by (EVERY1 [dtac hypnat_le_imp_less_or_eq, dtac hypnat_le_imp_less_or_eq,
+            rtac hypnat_less_or_eq_imp_le, fast_tac (claset() addIs [hypnat_less_trans])]);
+qed "hypnat_le_trans";
+
+Goal "!!z. [| z <= w; w <= z |] ==> z = (w::hypnat)";
+by (EVERY1 [dtac hypnat_le_imp_less_or_eq, dtac hypnat_le_imp_less_or_eq,
+            fast_tac (claset() addEs [hypnat_less_irrefl,hypnat_less_asym])]);
+qed "hypnat_le_anti_sym";
+
+Goal "!!x. [| ~ y < x; y ~= x |] ==> x < (y::hypnat)";
+by (rtac not_hypnat_leE 1);
+by (fast_tac (claset() addDs [hypnat_le_imp_less_or_eq]) 1);
+qed "not_less_not_eq_hypnat_less";
+
+Goal "!!x. [| (0::hypnat) <= x; 0 <= y |] ==> 0 <= x * y";
+by (REPEAT(dtac hypnat_le_imp_less_or_eq 1));
+by (auto_tac (claset() addIs [hypnat_mult_order,
+    hypnat_less_imp_le],simpset() addsimps [hypnat_le_refl]));
+qed "hypnat_le_mult_order";
+
+Goalw [hypnat_one_def,hypnat_zero_def,hypnat_less_def] 
+      "(0::hypnat) < 1hn";
+by (res_inst_tac [("x","%n. 0")] exI 1);
+by (res_inst_tac [("x","%n. 1")] exI 1);
+by (Auto_tac);
+qed "hypnat_zero_less_one";
+
+Goal "!!x. [| (0::hypnat) <= x; 0 <= y |] ==> 0 <= x + y";
+by (REPEAT(dtac hypnat_le_imp_less_or_eq 1));
+by (auto_tac (claset() addIs [hypnat_add_order,
+    hypnat_less_imp_le],simpset() addsimps [hypnat_le_refl]));
+qed "hypnat_le_add_order";
+
+Goal "!!(q1::hypnat). q1 <= q2  ==> x + q1 <= x + q2";
+by (dtac hypnat_le_imp_less_or_eq 1);
+by (Step_tac 1);
+by (auto_tac (claset() addSIs [hypnat_le_refl,
+    hypnat_less_imp_le,hypnat_add_less_mono1],
+    simpset() addsimps [hypnat_add_commute]));
+qed "hypnat_add_le_mono2";
+
+Goal "!!(q1::hypnat). q1 <= q2  ==> q1 + x <= q2 + x";
+by (auto_tac (claset() addDs [hypnat_add_le_mono2],
+    simpset() addsimps [hypnat_add_commute]));
+qed "hypnat_add_le_mono1";
+
+Goal "!!k l::hypnat. [|i<=j;  k<=l |] ==> i + k <= j + l";
+by (etac (hypnat_add_le_mono1 RS hypnat_le_trans) 1);
+by (simp_tac (simpset() addsimps [hypnat_add_commute]) 1);
+(*j moves to the end because it is free while k, l are bound*)
+by (etac hypnat_add_le_mono1 1);
+qed "hypnat_add_le_mono";
+
+Goalw [hypnat_zero_def]
+     "!!x::hypnat. [| (0::hypnat) < z; x < y |] ==> x * z < y * z";
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [hypnat_less,hypnat_mult]));
+by (Fuf_tac 1);
+qed "hypnat_mult_less_mono1";
+
+Goal "!!x::hypnat. [| 0 < z; x < y |] ==> z * x < z * y";
+by (auto_tac (claset() addIs [hypnat_mult_less_mono1],
+              simpset() addsimps [hypnat_mult_commute]));
+qed "hypnat_mult_less_mono2";
+
+Addsimps [hypnat_mult_less_mono2,hypnat_mult_less_mono1];
+
+Goal "(x::hypnat) <= n + x";
+by (res_inst_tac [("x","n")] hypnat_trichotomyE 1);
+by (auto_tac (claset() addDs [(hypnat_less_imp_le RS
+    hypnat_add_le_mono1)],simpset() addsimps [hypnat_le_refl]));
+qed "hypnat_add_self_le";
+Addsimps [hypnat_add_self_le];
+
+Goal "(x::hypnat) < x + 1hn";
+by (cut_facts_tac [hypnat_zero_less_one 
+         RS hypnat_add_less_mono2] 1);
+by (Auto_tac);
+qed "hypnat_add_one_self_less";
+Addsimps [hypnat_add_one_self_less];
+
+Goalw [hypnat_le_def] "~ x + 1hn <= x";
+by (Simp_tac 1);
+qed "not_hypnat_add_one_le_self";
+Addsimps [not_hypnat_add_one_le_self];
+
+Goal "((0::hypnat) < n) = (1hn <= n)";
+by (res_inst_tac [("x","n")] hypnat_trichotomyE 1);
+by (auto_tac (claset(),simpset() addsimps [hypnat_le_def]));
+qed "hypnat_gt_zero_iff";
+
+Addsimps  [hypnat_le_add_diff_inverse, hypnat_le_add_diff_inverse2,
+           hypnat_less_imp_le RS hypnat_le_add_diff_inverse2];
+
+Goal "(0 < n) = (EX m. n = m + 1hn)";
+by (Step_tac 1);
+by (res_inst_tac [("x","m")] hypnat_trichotomyE 2);
+by (rtac hypnat_less_trans 2);
+by (res_inst_tac [("x","n - 1hn")] exI 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [hypnat_gt_zero_iff,hypnat_add_commute]));
+qed "hypnat_gt_zero_iff2";
+
+Goalw [hypnat_zero_def] "(0::hypnat) <= n";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (asm_simp_tac (simpset() addsimps [hypnat_le]) 1);
+qed "hypnat_le_zero";
+Addsimps [hypnat_le_zero];
+
+(*------------------------------------------------------------------
+     hypnat_of_nat: properties embedding of naturals in hypernaturals
+ -----------------------------------------------------------------*)
+    (** hypnat_of_nat preserves field and order properties **)
+
+Goalw [hypnat_of_nat_def] 
+      "hypnat_of_nat ((z1::nat) + z2) = \
+\      hypnat_of_nat z1 + hypnat_of_nat z2";
+by (asm_simp_tac (simpset() addsimps [hypnat_add]) 1);
+qed "hypnat_of_nat_add";
+
+Goalw [hypnat_of_nat_def] 
+      "hypnat_of_nat ((z1::nat) - z2) = \
+\      hypnat_of_nat z1 - hypnat_of_nat z2";
+by (asm_simp_tac (simpset() addsimps [hypnat_minus]) 1);
+qed "hypnat_of_nat_minus";
+
+Goalw [hypnat_of_nat_def] 
+      "hypnat_of_nat (z1 * z2) = hypnat_of_nat z1 * hypnat_of_nat z2";
+by (full_simp_tac (simpset() addsimps [hypnat_mult]) 1);
+qed "hypnat_of_nat_mult";
+
+Goalw [hypnat_less_def,hypnat_of_nat_def] 
+      "(z1 < z2) = (hypnat_of_nat z1 < hypnat_of_nat z2)";
+by (auto_tac (claset() addSIs [exI] addIs 
+   [FreeUltrafilterNat_all],simpset()));
+by (rtac FreeUltrafilterNat_P 1 THEN Fuf_tac 1);
+qed "hypnat_of_nat_less_iff";
+Addsimps [hypnat_of_nat_less_iff RS sym];
+
+Goalw [hypnat_le_def,le_def] 
+      "(z1 <= z2) = (hypnat_of_nat z1 <= hypnat_of_nat z2)";
+by (Auto_tac);
+qed "hypnat_of_nat_le_iff";
+
+Goalw [hypnat_of_nat_def,hypnat_one_def] "hypnat_of_nat  1 = 1hn";
+by (Simp_tac 1);
+qed "hypnat_of_nat_one";
+
+Goalw [hypnat_of_nat_def,hypnat_zero_def] "hypnat_of_nat  0 = 0";
+by (Simp_tac 1);
+qed "hypnat_of_nat_zero";
+
+Goal "(hypnat_of_nat  n = 0) = (n = 0)";
+by (auto_tac (claset() addIs [FreeUltrafilterNat_P],
+    simpset() addsimps [hypnat_of_nat_def,
+    hypnat_zero_def]));
+qed "hypnat_of_nat_zero_iff";
+
+Goal "(hypnat_of_nat  n ~= 0) = (n ~= 0)";
+by (full_simp_tac (simpset() addsimps [hypnat_of_nat_zero_iff]) 1);
+qed "hypnat_of_nat_not_zero_iff";
+
+goalw HyperNat.thy [hypnat_of_nat_def,hypnat_one_def]
+      "hypnat_of_nat (Suc n) = hypnat_of_nat n + 1hn";
+by (auto_tac (claset(),simpset() addsimps [hypnat_add]));
+qed "hypnat_of_nat_Suc";
+
+(*---------------------------------------------------------------------------------
+              Existence of infinite hypernatural number
+ ---------------------------------------------------------------------------------*)
+
+Goal "hypnatrel^^{%n::nat. n} : hypnat";
+by (Auto_tac);
+qed "hypnat_omega";
+
+Goalw [hypnat_omega_def] "Rep_hypnat(whn) : hypnat";
+by (rtac Rep_hypnat 1);
+qed "Rep_hypnat_omega";
+
+(* See Hyper.thy for similar argument*)
+(* existence of infinite number not corresponding to any natural number *)
+(* use assumption that member FreeUltrafilterNat is not finite       *)
+(* a few lemmas first *)
+
+Goalw [hypnat_omega_def,hypnat_of_nat_def]
+      "~ (EX x. hypnat_of_nat x = whn)";
+by (auto_tac (claset() addDs [FreeUltrafilterNat_not_finite], 
+              simpset()));
+qed "not_ex_hypnat_of_nat_eq_omega";
+
+Goal "hypnat_of_nat x ~= whn";
+by (cut_facts_tac [not_ex_hypnat_of_nat_eq_omega] 1);
+by (Auto_tac);
+qed "hypnat_of_nat_not_eq_omega";
+Addsimps [hypnat_of_nat_not_eq_omega RS not_sym];
+
+(*-----------------------------------------------------------
+    Properties of the set SHNat of embedded natural numbers
+              (cf. set SReal in NSA.thy/NSA.ML)
+ ----------------------------------------------------------*)
+
+(* Infinite hypernatural not in embedded SHNat *)
+Goalw [SHNat_def] "whn ~: SHNat";
+by (Auto_tac);
+qed "SHNAT_omega_not_mem";
+Addsimps [SHNAT_omega_not_mem];
+
+(*-----------------------------------------------------------------------
+     Closure laws for members of (embedded) set standard naturals SHNat
+ -----------------------------------------------------------------------*)
+Goalw [SHNat_def] 
+      "!!x. [| x: SHNat; y: SHNat |] ==> x + y: SHNat";
+by (Step_tac 1);
+by (res_inst_tac [("x","N + Na")] exI 1);
+by (simp_tac (simpset() addsimps [hypnat_of_nat_add]) 1);
+qed "SHNat_add";
+
+Goalw [SHNat_def] 
+      "!!x. [| x: SHNat; y: SHNat |] ==> x - y: SHNat";
+by (Step_tac 1);
+by (res_inst_tac [("x","N - Na")] exI 1);
+by (simp_tac (simpset() addsimps [hypnat_of_nat_minus]) 1);
+qed "SHNat_minus";
+
+Goalw [SHNat_def] 
+      "!!x. [| x: SHNat; y: SHNat |] ==> x * y: SHNat";
+by (Step_tac 1);
+by (res_inst_tac [("x","N * Na")] exI 1);
+by (simp_tac (simpset() addsimps [hypnat_of_nat_mult]) 1);
+qed "SHNat_mult";
+
+Goal "!!x. [| x + y : SHNat; y: SHNat |] ==> x: SHNat";
+by (dres_inst_tac [("x","x+y")] SHNat_minus 1);
+by (Auto_tac);
+qed "SHNat_add_cancel";
+
+Goalw [SHNat_def] "hypnat_of_nat x : SHNat";
+by (Blast_tac 1);
+qed "SHNat_hypnat_of_nat";
+Addsimps [SHNat_hypnat_of_nat];
+
+Goal "hypnat_of_nat 1 : SHNat";
+by (Simp_tac 1);
+qed "SHNat_hypnat_of_nat_one";
+
+Goal "hypnat_of_nat 0 : SHNat";
+by (Simp_tac 1);
+qed "SHNat_hypnat_of_nat_zero";
+
+Goal "1hn : SHNat";
+by (simp_tac (simpset() addsimps [SHNat_hypnat_of_nat_one,
+    hypnat_of_nat_one RS sym]) 1);
+qed "SHNat_one";
+
+Goal "0 : SHNat";
+by (simp_tac (simpset() addsimps [SHNat_hypnat_of_nat_zero,
+    hypnat_of_nat_zero RS sym]) 1);
+qed "SHNat_zero";
+
+Addsimps [SHNat_hypnat_of_nat_one,SHNat_hypnat_of_nat_zero,
+          SHNat_one,SHNat_zero];
+
+Goal "1hn + 1hn : SHNat";
+by (rtac ([SHNat_one,SHNat_one] MRS SHNat_add) 1);
+qed "SHNat_two";
+
+Goalw [SHNat_def] "{x. hypnat_of_nat x : SHNat} = (UNIV::nat set)";
+by (Auto_tac);
+qed "SHNat_UNIV_nat";
+
+Goalw [SHNat_def] "(x: SHNat) = (EX y. x = hypnat_of_nat  y)";
+by (Auto_tac);
+qed "SHNat_iff";
+
+Goalw [SHNat_def] "hypnat_of_nat ``(UNIV::nat set) = SHNat";
+by (Auto_tac);
+qed "hypnat_of_nat_image";
+
+Goalw [SHNat_def] "inv hypnat_of_nat ``SHNat = (UNIV::nat set)";
+by (Auto_tac);
+by (rtac (inj_hypnat_of_nat RS inv_f_f RS subst) 1);
+by (Blast_tac 1);
+qed "inv_hypnat_of_nat_image";
+
+Goalw [SHNat_def] 
+      "!!P. [| EX x. x: P; P <= SHNat |] ==> \
+\           EX Q. P = hypnat_of_nat `` Q";
+by (Best_tac 1); 
+qed "SHNat_hypnat_of_nat_image";
+
+Goalw [SHNat_def] 
+      "SHNat = hypnat_of_nat `` (UNIV::nat set)";
+by (Auto_tac);
+qed "SHNat_hypnat_of_nat_iff";
+
+Goalw [SHNat_def] "SHNat <= (UNIV::hypnat set)";
+by (Auto_tac);
+qed "SHNat_subset_UNIV";
+
+Goal "{n. n <= Suc m} = {n. n <= m} Un {n. n = Suc m}";
+by (auto_tac (claset(),simpset() addsimps [le_Suc_eq]));
+qed "leSuc_Un_eq";
+
+Goal "finite {n::nat. n <= m}";
+by (nat_ind_tac "m" 1);
+by (auto_tac (claset(),simpset() addsimps [leSuc_Un_eq]));
+qed "finite_nat_le_segment";
+
+Goal "{n::nat. m < n} : FreeUltrafilterNat";
+by (cut_inst_tac [("m2","m")] (finite_nat_le_segment RS 
+    FreeUltrafilterNat_finite RS FreeUltrafilterNat_Compl_mem) 1);
+by (Fuf_tac 1);
+qed "lemma_unbounded_set";
+Addsimps [lemma_unbounded_set];
+
+Goalw [SHNat_def,hypnat_of_nat_def,
+           hypnat_less_def,hypnat_omega_def] 
+           "ALL n: SHNat. n < whn";
+by (Clarify_tac 1);
+by (auto_tac (claset() addSIs [exI],simpset()));
+qed "hypnat_omega_gt_SHNat";
+
+Goal "hypnat_of_nat n < whn";
+by (cut_facts_tac [hypnat_omega_gt_SHNat] 1);
+by (dres_inst_tac [("x","hypnat_of_nat n")] bspec 1);
+by (Auto_tac);
+qed "hypnat_of_nat_less_whn";
+Addsimps [hypnat_of_nat_less_whn];
+
+Goal "hypnat_of_nat n <= whn";
+by (rtac (hypnat_of_nat_less_whn RS hypnat_less_imp_le) 1);
+qed "hypnat_of_nat_le_whn";
+Addsimps [hypnat_of_nat_le_whn];
+
+Goal "0 < whn";
+by (rtac (hypnat_omega_gt_SHNat RS ballE) 1);
+by (Auto_tac);
+qed "hypnat_zero_less_hypnat_omega";
+Addsimps [hypnat_zero_less_hypnat_omega];
+
+Goal "1hn < whn";
+by (rtac (hypnat_omega_gt_SHNat RS ballE) 1);
+by (Auto_tac);
+qed "hypnat_one_less_hypnat_omega";
+Addsimps [hypnat_one_less_hypnat_omega];
+
+(*--------------------------------------------------------------------------
+     Theorems about infinite hypernatural numbers -- HNatInfinite
+ -------------------------------------------------------------------------*)
+Goalw [HNatInfinite_def,SHNat_def] "whn : HNatInfinite";
+by (Auto_tac);
+qed "HNatInfinite_whn";
+Addsimps [HNatInfinite_whn];
+
+Goalw [HNatInfinite_def] "!!x. x: SHNat ==> x ~: HNatInfinite";
+by (Simp_tac 1);
+qed "SHNat_not_HNatInfinite";
+
+Goalw [HNatInfinite_def] "!!x. x ~: HNatInfinite ==> x: SHNat";
+by (Asm_full_simp_tac 1);
+qed "not_HNatInfinite_SHNat";
+
+Goalw [HNatInfinite_def] "!!x. x ~: SHNat ==> x: HNatInfinite";
+by (Simp_tac 1);
+qed "not_SHNat_HNatInfinite";
+
+Goalw [HNatInfinite_def] "!!x. x: HNatInfinite ==> x ~: SHNat";
+by (Asm_full_simp_tac 1);
+qed "HNatInfinite_not_SHNat";
+
+Goal "(x: SHNat) = (x ~: HNatInfinite)";
+by (blast_tac (claset() addSIs [SHNat_not_HNatInfinite,
+    not_HNatInfinite_SHNat]) 1);
+qed "SHNat_not_HNatInfinite_iff";
+
+Goal "(x ~: SHNat) = (x: HNatInfinite)";
+by (blast_tac (claset() addSIs [not_SHNat_HNatInfinite,
+    HNatInfinite_not_SHNat]) 1);
+qed "not_SHNat_HNatInfinite_iff";
+
+Goal "x : SHNat | x : HNatInfinite";
+by (simp_tac (simpset() addsimps [SHNat_not_HNatInfinite_iff]) 1);
+qed "SHNat_HNatInfinite_disj";
+
+(*-------------------------------------------------------------------
+   Proof of alternative definition for set of Infinite hypernatural 
+   numbers --- HNatInfinite = {N. ALL n: SHNat. n < N}
+ -------------------------------------------------------------------*)
+Goal "!!N (xa::nat=>nat). \
+\         (ALL N. {n. xa n ~= N} : FreeUltrafilterNat) ==> \
+\         ({n. N < xa n} : FreeUltrafilterNat)";
+by (nat_ind_tac "N" 1);
+by (dres_inst_tac [("x","0")] spec 1);
+by (rtac ccontr 1 THEN dtac FreeUltrafilterNat_Compl_mem 1
+    THEN dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
+by (Asm_full_simp_tac 1);
+by (dres_inst_tac [("x","Suc N")] spec 1);
+by (fuf_tac (claset() addSDs [Suc_leI],simpset() addsimps 
+    [le_eq_less_or_eq]) 1);
+qed "HNatInfinite_FreeUltrafilterNat_lemma";
+
+(*** alternative definition ***)
+Goalw [HNatInfinite_def,SHNat_def,hypnat_of_nat_def] 
+      "HNatInfinite = {N. ALL n:SHNat. n < N}";
+by (Step_tac 1);
+by (dres_inst_tac [("x","Abs_hypnat \
+\        (hypnatrel ^^ {%n. N})")] bspec 2);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),simpset() addsimps [hypnat_less_iff]));
+by (auto_tac (claset() addSIs [exI] addEs 
+    [HNatInfinite_FreeUltrafilterNat_lemma],
+    simpset() addsimps [FreeUltrafilterNat_Compl_iff1, 
+     CLAIM "- {n. xa n = N} = {n. xa n ~= N}"]));
+qed "HNatInfinite_iff";
+
+(*--------------------------------------------------------------------
+   Alternative definition for HNatInfinite using Free ultrafilter
+ --------------------------------------------------------------------*)
+Goal "!!x. x : HNatInfinite ==> EX X: Rep_hypnat x. \
+\          ALL u. {n. u < X n}:  FreeUltrafilterNat";
+by (auto_tac (claset(),simpset() addsimps [hypnat_less_def,
+    HNatInfinite_iff,SHNat_iff,hypnat_of_nat_def]));
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (EVERY[Auto_tac, rtac bexI 1, 
+    rtac lemma_hypnatrel_refl 2, Step_tac 1]);
+by (dres_inst_tac [("x","hypnat_of_nat u")] bspec 1);
+by (Simp_tac 1);
+by (auto_tac (claset(),
+    simpset() addsimps [hypnat_of_nat_def]));
+by (Fuf_tac 1);
+qed "HNatInfinite_FreeUltrafilterNat";
+
+Goal "!!x. EX X: Rep_hypnat x. \
+\          ALL u. {n. u < X n}:  FreeUltrafilterNat \
+\          ==> x: HNatInfinite";
+by (auto_tac (claset(),simpset() addsimps [hypnat_less_def,
+    HNatInfinite_iff,SHNat_iff,hypnat_of_nat_def]));
+by (rtac exI 1 THEN Auto_tac);
+qed "FreeUltrafilterNat_HNatInfinite";
+
+Goal "!!x. (x : HNatInfinite) = (EX X: Rep_hypnat x. \
+\          ALL u. {n. u < X n}:  FreeUltrafilterNat)";
+by (blast_tac (claset() addIs [HNatInfinite_FreeUltrafilterNat,
+    FreeUltrafilterNat_HNatInfinite]) 1);
+qed "HNatInfinite_FreeUltrafilterNat_iff";
+
+Goal "!!x. x : HNatInfinite ==> 1hn < x";
+by (auto_tac (claset(),simpset() addsimps [HNatInfinite_iff]));
+qed "HNatInfinite_gt_one";
+Addsimps [HNatInfinite_gt_one];
+
+Goal "!!x. 0 ~: HNatInfinite";
+by (auto_tac (claset(),simpset() 
+    addsimps [HNatInfinite_iff]));
+by (dres_inst_tac [("a","1hn")] equals0D 1);
+by (Asm_full_simp_tac 1);
+qed "zero_not_mem_HNatInfinite";
+Addsimps [zero_not_mem_HNatInfinite];
+
+Goal "!!x. x : HNatInfinite ==> x ~= 0";
+by (Auto_tac);
+qed "HNatInfinite_not_eq_zero";
+
+Goal "!!x. x : HNatInfinite ==> 1hn <= x";
+by (blast_tac (claset() addIs [hypnat_less_imp_le,
+         HNatInfinite_gt_one]) 1);
+qed "HNatInfinite_ge_one";
+Addsimps [HNatInfinite_ge_one];
+
+(*--------------------------------------------------
+                   Closure Rules
+ --------------------------------------------------*)
+Goal "!!x. [| x: HNatInfinite; y: HNatInfinite |] \
+\           ==> x + y: HNatInfinite";
+by (auto_tac (claset(),simpset() addsimps [HNatInfinite_iff]));
+by (dtac bspec 1 THEN assume_tac 1);
+by (dtac (SHNat_zero RSN (2,bspec)) 1);
+by (dtac hypnat_add_less_mono 1 THEN assume_tac 1);
+by (Asm_full_simp_tac 1);
+qed "HNatInfinite_add";
+
+Goal "!!x. [| x: HNatInfinite; y: SHNat |] \
+\                       ==> x + y: HNatInfinite";
+by (rtac ccontr 1 THEN dtac not_HNatInfinite_SHNat 1);
+by (dres_inst_tac [("x","x + y")] SHNat_minus 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [SHNat_not_HNatInfinite_iff]));
+qed "HNatInfinite_SHNat_add";
+
+goal HyperNat.thy "!!x. [| x: HNatInfinite; y: SHNat |] \
+\                       ==> x - y: HNatInfinite";
+by (rtac ccontr 1 THEN dtac not_HNatInfinite_SHNat 1);
+by (dres_inst_tac [("x","x - y")] SHNat_add 1);
+by (subgoal_tac "y <= x" 2);
+by (auto_tac (claset() addSDs [hypnat_le_add_diff_inverse2],
+    simpset() addsimps [not_SHNat_HNatInfinite_iff RS sym]));
+by (auto_tac (claset() addSIs [hypnat_less_imp_le],
+    simpset() addsimps [not_SHNat_HNatInfinite_iff,HNatInfinite_iff]));
+qed "HNatInfinite_SHNat_diff";
+
+Goal 
+     "!!x. x: HNatInfinite ==> x + 1hn: HNatInfinite";
+by (auto_tac (claset() addIs [HNatInfinite_SHNat_add],
+              simpset()));
+qed "HNatInfinite_add_one";
+
+Goal 
+     "!!x. x: HNatInfinite ==> x - 1hn: HNatInfinite";
+by (rtac ccontr 1 THEN dtac not_HNatInfinite_SHNat 1);
+by (dres_inst_tac [("x","x - 1hn"),("y","1hn")] SHNat_add 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [not_SHNat_HNatInfinite_iff RS sym]));
+qed "HNatInfinite_minus_one";
+
+Goal "!!x. x : HNatInfinite ==> EX y. x = y + 1hn";
+by (res_inst_tac [("x","x - 1hn")] exI 1);
+by (Auto_tac);
+qed "HNatInfinite_is_Suc";
+
+(*---------------------------------------------------------------
+       HNat : the hypernaturals embedded in the hyperreals
+       Obtained using the NS extension of the naturals
+ --------------------------------------------------------------*)
+ 
+Goalw [HNat_def,starset_def,
+         hypreal_of_posnat_def,hypreal_of_real_def] 
+         "hypreal_of_posnat N : HNat";
+by (Auto_tac);
+by (Ultra_tac 1);
+by (res_inst_tac [("x","Suc N")] exI 1);
+by (dtac sym 1 THEN auto_tac (claset(),simpset() 
+    addsimps [real_of_preal_real_of_nat_Suc]));
+qed "HNat_hypreal_of_posnat";
+Addsimps [HNat_hypreal_of_posnat];
+
+Goalw [HNat_def,starset_def]
+     "[| x: HNat; y: HNat |] ==> x + y: HNat";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addSDs [bspec] addIs [lemma_hyprel_refl],
+    simpset() addsimps [hypreal_add]));
+by (Ultra_tac 1);
+by (dres_inst_tac [("t","Y xb")] sym 1);
+by (auto_tac (claset(),simpset() addsimps [real_of_nat_add RS sym]));
+qed "HNat_add";
+
+Goalw [HNat_def,starset_def]
+     "[| x: HNat; y: HNat |] ==> x * y: HNat";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addSDs [bspec] addIs [lemma_hyprel_refl],
+    simpset() addsimps [hypreal_mult]));
+by (Ultra_tac 1);
+by (dres_inst_tac [("t","Y xb")] sym 1);
+by (auto_tac (claset(),simpset() addsimps [real_of_nat_mult RS sym]));
+qed "HNat_mult";
+
+(*---------------------------------------------------------------
+    Embedding of the hypernaturals into the hyperreal
+ --------------------------------------------------------------*)
+(*** A lemma that should have been derived a long time ago! ***)
+Goal "(Ya : hyprel ^^{%n. f(n)}) = \
+\         ({n. f n = Ya n} : FreeUltrafilterNat)";
+by (Auto_tac);
+qed "lemma_hyprel_FUFN";
+
+Goalw [hypreal_of_hypnat_def]
+      "hypreal_of_hypnat (Abs_hypnat(hypnatrel^^{%n. X n})) = \
+\         Abs_hypreal(hyprel ^^ {%n. real_of_nat (X n)})";
+by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
+by (auto_tac (claset() addEs [FreeUltrafilterNat_Int RS
+    FreeUltrafilterNat_subset],simpset() addsimps 
+    [lemma_hyprel_FUFN]));
+qed "hypreal_of_hypnat";
+
+Goal "inj(hypreal_of_hypnat)";
+by (rtac injI 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [hypreal_of_hypnat,real_of_nat_eq_cancel]));
+qed "inj_hypreal_of_hypnat";
+
+Goal "(hypreal_of_hypnat n = hypreal_of_hypnat m) = (n = m)";
+by (auto_tac (claset(),simpset() addsimps [inj_hypreal_of_hypnat RS injD]));
+qed "hypreal_of_hypnat_eq_cancel";
+Addsimps [hypreal_of_hypnat_eq_cancel];
+
+Goal "(hypnat_of_nat n = hypnat_of_nat m) = (n = m)";
+by (auto_tac (claset() addDs [inj_hypnat_of_nat RS injD],
+              simpset()));
+qed "hypnat_of_nat_eq_cancel";
+Addsimps [hypnat_of_nat_eq_cancel];
+
+Goalw [hypnat_zero_def] 
+     "hypreal_of_hypnat 0 = #0";
+by (simp_tac (HOL_ss addsimps
+             [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1); 
+by (simp_tac (simpset() addsimps [hypreal_of_hypnat, real_of_nat_zero]) 1);
+qed "hypreal_of_hypnat_zero";
+
+Goalw [hypnat_one_def] 
+     "hypreal_of_hypnat 1hn = #1";
+by (simp_tac (HOL_ss addsimps
+             [one_eq_numeral_1 RS sym, hypreal_one_def]) 1); 
+by (simp_tac (simpset() addsimps [hypreal_of_hypnat, real_of_nat_one]) 1);
+qed "hypreal_of_hypnat_one";
+
+Goal "hypreal_of_hypnat n1 + hypreal_of_hypnat n2 = hypreal_of_hypnat (n1 + n2)";
+by (res_inst_tac [("z","n1")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","n2")] eq_Abs_hypnat 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_of_hypnat,
+        hypreal_add,hypnat_add,real_of_nat_add]) 1);
+qed "hypreal_of_hypnat_add";
+
+Goal "hypreal_of_hypnat n1 * hypreal_of_hypnat n2 = hypreal_of_hypnat (n1 * n2)";
+by (res_inst_tac [("z","n1")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","n2")] eq_Abs_hypnat 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_of_hypnat,
+        hypreal_mult,hypnat_mult,real_of_nat_mult]) 1);
+qed "hypreal_of_hypnat_mult";
+
+Goal "(hypreal_of_hypnat n < hypreal_of_hypnat m) = (n < m)";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (asm_simp_tac (simpset() addsimps 
+    [hypreal_of_hypnat,hypreal_less,hypnat_less]) 1);
+qed "hypreal_of_hypnat_less_iff";
+Addsimps [hypreal_of_hypnat_less_iff];
+
+Goal "(hypreal_of_hypnat N = #0) = (N = 0)";
+by (simp_tac (simpset() addsimps [hypreal_of_hypnat_zero RS sym]) 1);
+qed "hypreal_of_hypnat_eq_zero_iff";
+Addsimps [hypreal_of_hypnat_eq_zero_iff];
+
+Goal "ALL n. N <= n ==> N = (0::hypnat)";
+by (dres_inst_tac [("x","0")] spec 1);
+by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),simpset() addsimps [hypnat_le,hypnat_zero_def]));
+qed "hypnat_eq_zero";
+Addsimps [hypnat_eq_zero];
+
+Goal "~ (ALL n. n = (0::hypnat))";
+by Auto_tac;
+by (res_inst_tac [("x","1hn")] exI 1);
+by (Simp_tac 1);
+qed "hypnat_not_all_eq_zero";
+Addsimps [hypnat_not_all_eq_zero];
+
+Goal "n ~= 0 ==> (n <= 1hn) = (n = 1hn)";
+by (auto_tac (claset(),simpset() addsimps [hypnat_le_eq_less_or_eq]));
+qed "hypnat_le_one_eq_one";
+Addsimps [hypnat_le_one_eq_one];
+
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HyperNat.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,84 @@
+(*  Title       : HyperNat.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Explicit construction of hypernaturals using 
+                  ultrafilters
+*) 
+
+HyperNat = Star + 
+
+constdefs
+    hypnatrel :: "((nat=>nat)*(nat=>nat)) set"
+    "hypnatrel == {p. EX X Y. p = ((X::nat=>nat),Y) & 
+                       {n::nat. X(n) = Y(n)} : FreeUltrafilterNat}"
+
+typedef hypnat = "UNIV//hypnatrel"              (Equiv.quotient_def)
+
+instance
+   hypnat  :: {ord,zero,plus,times,minus}
+
+consts
+  "1hn"       :: hypnat               ("1hn")  
+  "whn"       :: hypnat               ("whn")  
+
+defs
+
+  hypnat_zero_def      "0 == Abs_hypnat(hypnatrel^^{%n::nat. 0})"
+  hypnat_one_def       "1hn == Abs_hypnat(hypnatrel^^{%n::nat. 1})"
+
+  (* omega is in fact an infinite hypernatural number = [<1,2,3,...>] *)
+  hypnat_omega_def     "whn == Abs_hypnat(hypnatrel^^{%n::nat. n})"
+ 
+
+constdefs
+
+  (* embedding the naturals in the hypernaturals *)
+  hypnat_of_nat   :: nat => hypnat
+  "hypnat_of_nat m  == Abs_hypnat(hypnatrel^^{%n::nat. m})"
+
+  (* set of naturals embedded in the hypernaturals*)
+  SHNat         :: "hypnat set"
+  "SHNat        == {n. EX N. n = hypnat_of_nat N}"  
+ 
+  (* embedding the naturals in the hyperreals*)
+  SNat         :: "hypreal set"
+  "SNat        == {n. EX N. n = hypreal_of_nat N}"  
+
+  (* hypernaturals as members of the hyperreals; the set is defined as  *)
+  (* the nonstandard extension of set of naturals embedded in the reals *)
+  HNat         :: "hypreal set"
+  "HNat         == *s* {n. EX no. n = real_of_nat no}"
+
+  (* the set of infinite hypernatural numbers *)
+  HNatInfinite :: "hypnat set"
+  "HNatInfinite == {n. n ~: SHNat}"
+
+  (* explicit embedding of the hypernaturals in the hyperreals *)    
+  hypreal_of_hypnat :: hypnat => hypreal
+  "hypreal_of_hypnat N  == Abs_hypreal(UN X: Rep_hypnat(N). 
+                            hyprel^^{%n::nat. real_of_nat (X n)})"
+  
+defs
+  hypnat_add_def  
+  "P + Q == Abs_hypnat(UN X:Rep_hypnat(P). UN Y:Rep_hypnat(Q).
+                hypnatrel^^{%n::nat. X n + Y n})"
+
+  hypnat_mult_def  
+  "P * Q == Abs_hypnat(UN X:Rep_hypnat(P). UN Y:Rep_hypnat(Q).
+                hypnatrel^^{%n::nat. X n * Y n})"
+
+  hypnat_minus_def  
+  "P - Q == Abs_hypnat(UN X:Rep_hypnat(P). UN Y:Rep_hypnat(Q).
+                hypnatrel^^{%n::nat. X n - Y n})"
+
+  hypnat_less_def
+  "P < (Q::hypnat) == EX X Y. X: Rep_hypnat(P) & 
+                               Y: Rep_hypnat(Q) & 
+                               {n::nat. X n < Y n} : FreeUltrafilterNat"
+  hypnat_le_def
+  "P <= (Q::hypnat) == ~(Q < P)" 
+
+end
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HyperOrd.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,715 @@
+(*  Title:	 Real/Hyperreal/HyperOrd.ML
+    Author:      Jacques D. Fleuriot
+    Copyright:   1998 University of Cambridge
+                 2000 University of Edinburgh
+    Description: Type "hypreal" is a linear order and also 
+                 satisfies plus_ac0: + is an AC-operator with zero
+*)
+
+(**** The simproc abel_cancel ****)
+
+(*** Two lemmas needed for the simprocs ***)
+
+(*Deletion of other terms in the formula, seeking the -x at the front of z*)
+Goal "((x::hypreal) + (y + z) = y + u) = ((x + z) = u)";
+by (stac hypreal_add_left_commute 1);
+by (rtac hypreal_add_left_cancel 1);
+qed "hypreal_add_cancel_21";
+
+(*A further rule to deal with the case that
+  everything gets cancelled on the right.*)
+Goal "((x::hypreal) + (y + z) = y) = (x = -z)";
+by (stac hypreal_add_left_commute 1);
+by (res_inst_tac [("t", "y")] (hypreal_add_zero_right RS subst) 1
+    THEN stac hypreal_add_left_cancel 1);
+by (simp_tac (simpset() addsimps [hypreal_eq_diff_eq RS sym]) 1);
+qed "hypreal_add_cancel_end";
+
+
+structure Hyperreal_Cancel_Data =
+struct
+  val ss		= HOL_ss
+  val eq_reflection	= eq_reflection
+
+  val sg_ref		= Sign.self_ref (Theory.sign_of (the_context ()))
+  val T			= Type("HyperDef.hypreal",[])
+  val zero		= Const ("0", T)
+  val restrict_to_left  = restrict_to_left
+  val add_cancel_21	= hypreal_add_cancel_21
+  val add_cancel_end	= hypreal_add_cancel_end
+  val add_left_cancel	= hypreal_add_left_cancel
+  val add_assoc		= hypreal_add_assoc
+  val add_commute	= hypreal_add_commute
+  val add_left_commute	= hypreal_add_left_commute
+  val add_0		= hypreal_add_zero_left
+  val add_0_right	= hypreal_add_zero_right
+
+  val eq_diff_eq	= hypreal_eq_diff_eq
+  val eqI_rules		= [hypreal_less_eqI, hypreal_eq_eqI, hypreal_le_eqI]
+  fun dest_eqI th = 
+      #1 (HOLogic.dest_bin "op =" HOLogic.boolT 
+	      (HOLogic.dest_Trueprop (concl_of th)))
+
+  val diff_def		= hypreal_diff_def
+  val minus_add_distrib	= hypreal_minus_add_distrib
+  val minus_minus	= hypreal_minus_minus
+  val minus_0		= hypreal_minus_zero
+  val add_inverses	= [hypreal_add_minus, hypreal_add_minus_left]
+  val cancel_simps	= [hypreal_add_minus_cancelA, hypreal_minus_add_cancelA]
+end;
+
+structure Hyperreal_Cancel = Abel_Cancel (Hyperreal_Cancel_Data);
+
+Addsimprocs [Hyperreal_Cancel.sum_conv, Hyperreal_Cancel.rel_conv];
+
+Goal "- (z - y) = y - (z::hypreal)";
+by (Simp_tac 1);
+qed "hypreal_minus_diff_eq";
+Addsimps [hypreal_minus_diff_eq];
+
+
+Goal "((x::hypreal) < y) = (-y < -x)";
+by (stac hypreal_less_minus_iff 1);
+by (res_inst_tac [("x1","x")] (hypreal_less_minus_iff RS ssubst) 1);
+by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
+qed "hypreal_less_swap_iff";
+
+Goalw [hypreal_zero_def] 
+      "((0::hypreal) < x) = (-x < x)";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_less, hypreal_minus]));
+by (ALLGOALS(Ultra_tac));
+qed "hypreal_gt_zero_iff";
+
+Goal "(A::hypreal) < B ==> A + C < B + C";
+by (res_inst_tac [("z","A")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","B")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","C")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addSIs [exI], 
+              simpset() addsimps [hypreal_less_def,hypreal_add]));
+by (Ultra_tac 1);
+qed "hypreal_add_less_mono1";
+
+Goal "!!(A::hypreal). A < B ==> C + A < C + B";
+by (auto_tac (claset() addIs [hypreal_add_less_mono1],
+    simpset() addsimps [hypreal_add_commute]));
+qed "hypreal_add_less_mono2";
+
+Goal "(x < (0::hypreal)) = (x < -x)";
+by (rtac (hypreal_minus_zero_less_iff RS subst) 1);
+by (stac hypreal_gt_zero_iff 1);
+by (Full_simp_tac 1);
+qed "hypreal_lt_zero_iff";
+
+Goalw [hypreal_le_def] "((0::hypreal) <= x) = (-x <= x)";
+by (auto_tac (claset(),simpset() addsimps [hypreal_lt_zero_iff RS sym]));
+qed "hypreal_ge_zero_iff";
+
+Goalw [hypreal_le_def] "(x <= (0::hypreal)) = (x <= -x)";
+by (auto_tac (claset(),simpset() addsimps [hypreal_gt_zero_iff RS sym]));
+qed "hypreal_le_zero_iff";
+
+Goalw [hypreal_zero_def] 
+      "[| 0 < x; 0 < y |] ==> (0::hypreal) < x + y";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps
+    [hypreal_less_def,hypreal_add]));
+by (auto_tac (claset() addSIs [exI],simpset() addsimps
+    [hypreal_less_def,hypreal_add]));
+by (ultra_tac (claset() addIs [real_add_order],simpset()) 1);
+qed "hypreal_add_order";
+
+Goal "[| 0 < x; 0 <= y |] ==> (0::hypreal) < x + y";
+by (auto_tac (claset() addDs [sym, order_le_imp_less_or_eq]
+                       addIs [hypreal_add_order],
+              simpset()));
+qed "hypreal_add_order_le";            
+
+Goalw [hypreal_zero_def] 
+          "[| 0 < x; 0 < y |] ==> (0::hypreal) < x * y";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addSIs [exI],simpset() addsimps
+    [hypreal_less_def,hypreal_mult]));
+by (ultra_tac (claset() addIs [rename_numerals real_mult_order],
+	       simpset()) 1);
+qed "hypreal_mult_order";
+
+Goal "[| x < 0; y < 0 |] ==> (0::hypreal) < x * y";
+by (REPEAT(dtac (hypreal_minus_zero_less_iff RS iffD2) 1));
+by (dtac hypreal_mult_order 1 THEN assume_tac 1);
+by (Asm_full_simp_tac 1);
+qed "hypreal_mult_less_zero1";
+
+Goal "[| 0 <= x; 0 <= y |] ==> (0::hypreal) <= x * y";
+by (REPEAT(dtac order_le_imp_less_or_eq 1));
+by (auto_tac (claset() addIs [hypreal_mult_order, order_less_imp_le],
+              simpset()));
+qed "hypreal_le_mult_order";
+
+
+Goal "[| x <= 0; y <= 0 |] ==> (0::hypreal) <= x * y";
+by (rtac hypreal_less_or_eq_imp_le 1);
+by (dtac order_le_imp_less_or_eq 1 THEN etac disjE 1);
+by Auto_tac;
+by (dtac order_le_imp_less_or_eq 1);
+by (auto_tac (claset() addDs [hypreal_mult_less_zero1],simpset()));
+qed "hypreal_mult_le_zero1";
+
+Goal "[| 0 <= x; y < 0 |] ==> x * y <= (0::hypreal)";
+by (rtac hypreal_less_or_eq_imp_le 1);
+by (dtac order_le_imp_less_or_eq 1 THEN etac disjE 1);
+by Auto_tac;
+by (dtac (hypreal_minus_zero_less_iff RS iffD2) 1);
+by (rtac (hypreal_minus_zero_less_iff RS subst) 1);
+by (blast_tac (claset() addDs [hypreal_mult_order] 
+    addIs [hypreal_minus_mult_eq2 RS ssubst]) 1);
+qed "hypreal_mult_le_zero";
+
+Goal "[| 0 < x; y < 0 |] ==> x*y < (0::hypreal)";
+by (dtac (hypreal_minus_zero_less_iff RS iffD2) 1);
+by (dtac hypreal_mult_order 1 THEN assume_tac 1);
+by (rtac (hypreal_minus_zero_less_iff RS iffD1) 1);
+by (Asm_full_simp_tac 1);
+qed "hypreal_mult_less_zero";
+
+Goalw [hypreal_one_def,hypreal_zero_def,hypreal_less_def] "0 < 1hr";
+by (res_inst_tac [("x","%n. #0")] exI 1);
+by (res_inst_tac [("x","%n. #1")] exI 1);
+by (auto_tac (claset(),
+        simpset() addsimps [real_zero_less_one, FreeUltrafilterNat_Nat_set]));
+qed "hypreal_zero_less_one";
+
+Goal "1hr < 1hr + 1hr";
+by (rtac (hypreal_less_minus_iff RS iffD2) 1);
+by (full_simp_tac (simpset() addsimps [hypreal_zero_less_one,
+    hypreal_add_assoc]) 1);
+qed "hypreal_one_less_two";
+
+Goal "0 < 1hr + 1hr";
+by (rtac ([hypreal_zero_less_one,
+          hypreal_one_less_two] MRS order_less_trans) 1);
+qed "hypreal_zero_less_two";
+
+Goal "1hr + 1hr ~= 0";
+by (rtac (hypreal_zero_less_two RS hypreal_not_refl2 RS not_sym) 1);
+qed "hypreal_two_not_zero";
+Addsimps [hypreal_two_not_zero];
+
+Goal "[| 0 <= x; 0 <= y |] ==> (0::hypreal) <= x + y";
+by (REPEAT(dtac order_le_imp_less_or_eq 1));
+by (auto_tac (claset() addIs [hypreal_add_order, order_less_imp_le],
+              simpset()));
+qed "hypreal_le_add_order";
+
+(*** Monotonicity results ***)
+
+Goal "(v+z < w+z) = (v < (w::hypreal))";
+by (Simp_tac 1);
+qed "hypreal_add_right_cancel_less";
+
+Goal "(z+v < z+w) = (v < (w::hypreal))";
+by (Simp_tac 1);
+qed "hypreal_add_left_cancel_less";
+
+Addsimps [hypreal_add_right_cancel_less, 
+          hypreal_add_left_cancel_less];
+
+Goal "(v+z <= w+z) = (v <= (w::hypreal))";
+by (Simp_tac 1);
+qed "hypreal_add_right_cancel_le";
+
+Goal "(z+v <= z+w) = (v <= (w::hypreal))";
+by (Simp_tac 1);
+qed "hypreal_add_left_cancel_le";
+
+Addsimps [hypreal_add_right_cancel_le, hypreal_add_left_cancel_le];
+
+Goal  "[| (z1::hypreal) < y1; z2 < y2 |] ==> z1 + z2 < y1 + y2";
+by (dtac (hypreal_less_minus_iff RS iffD1) 1);
+by (dtac (hypreal_less_minus_iff RS iffD1) 1);
+by (dtac hypreal_add_order 1 THEN assume_tac 1);
+by (thin_tac "0 < y2 + - z2" 1);
+by (dres_inst_tac [("C","z1 + z2")] hypreal_add_less_mono1 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [hypreal_minus_add_distrib RS sym] @ hypreal_add_ac
+    delsimps [hypreal_minus_add_distrib]));
+qed "hypreal_add_less_mono";
+
+Goal "(q1::hypreal) <= q2  ==> x + q1 <= x + q2";
+by (dtac order_le_imp_less_or_eq 1);
+by (Step_tac 1);
+by (auto_tac (claset() addSIs [order_less_imp_le,hypreal_add_less_mono1],
+              simpset() addsimps [hypreal_add_commute]));
+qed "hypreal_add_left_le_mono1";
+
+Goal "(q1::hypreal) <= q2  ==> q1 + x <= q2 + x";
+by (auto_tac (claset() addDs [hypreal_add_left_le_mono1],
+              simpset() addsimps [hypreal_add_commute]));
+qed "hypreal_add_le_mono1";
+
+Goal "[|(i::hypreal)<=j;  k<=l |] ==> i + k <= j + l";
+by (etac (hypreal_add_le_mono1 RS order_trans) 1);
+by (Simp_tac 1);
+qed "hypreal_add_le_mono";
+
+Goal "[|(i::hypreal)<j;  k<=l |] ==> i + k < j + l";
+by (auto_tac (claset() addSDs [order_le_imp_less_or_eq] 
+                       addIs [hypreal_add_less_mono1,hypreal_add_less_mono],
+    simpset()));
+qed "hypreal_add_less_le_mono";
+
+Goal "[|(i::hypreal)<=j;  k<l |] ==> i + k < j + l";
+by (auto_tac (claset() addSDs [order_le_imp_less_or_eq] 
+                       addIs [hypreal_add_less_mono2,hypreal_add_less_mono],
+              simpset()));
+qed "hypreal_add_le_less_mono";
+
+Goal "(A::hypreal) + C < B + C ==> A < B";
+by (Full_simp_tac 1);
+qed "hypreal_less_add_right_cancel";
+
+Goal "(C::hypreal) + A < C + B ==> A < B";
+by (Full_simp_tac 1);
+qed "hypreal_less_add_left_cancel";
+
+Goal "[|r < x; (0::hypreal) <= y|] ==> r < x + y";
+by (auto_tac (claset() addDs [hypreal_add_less_le_mono],
+    simpset()));
+qed "hypreal_add_zero_less_le_mono";
+
+Goal "!!(A::hypreal). A + C <= B + C ==> A <= B";
+by (dres_inst_tac [("x","-C")] hypreal_add_le_mono1 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc]) 1);
+qed "hypreal_le_add_right_cancel";
+
+Goal "!!(A::hypreal). C + A <= C + B ==> A <= B";
+by (dres_inst_tac [("x","-C")] hypreal_add_left_le_mono1 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
+qed "hypreal_le_add_left_cancel";
+
+Goal "(0::hypreal) <= x*x";
+by (res_inst_tac [("x","0"),("y","x")] hypreal_linear_less2 1);
+by (auto_tac (claset() addIs [hypreal_mult_order,
+                              hypreal_mult_less_zero1,order_less_imp_le],
+              simpset()));
+qed "hypreal_le_square";
+Addsimps [hypreal_le_square];
+
+Goalw [hypreal_le_def] "- (x*x) <= (0::hypreal)";
+by (auto_tac (claset() addSDs [hypreal_le_square RS order_le_less_trans],
+      simpset() addsimps [hypreal_minus_zero_less_iff]));
+qed "hypreal_less_minus_square";
+Addsimps [hypreal_less_minus_square];
+
+Goal "(0*x<r)=((0::hypreal)<r)";
+by (Simp_tac 1);
+qed "hypreal_mult_0_less";
+
+Goal "[| (0::hypreal) < z; x < y |] ==> x*z < y*z";       
+by (rotate_tac 1 1);
+by (dtac (hypreal_less_minus_iff RS iffD1) 1);
+by (rtac (hypreal_less_minus_iff RS iffD2) 1);
+by (dtac hypreal_mult_order 1 THEN assume_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_add_mult_distrib2,
+					   hypreal_mult_commute ]) 1);
+qed "hypreal_mult_less_mono1";
+
+Goal "[| (0::hypreal)<z; x<y |] ==> z*x<z*y";       
+by (asm_simp_tac (simpset() addsimps [hypreal_mult_commute,hypreal_mult_less_mono1]) 1);
+qed "hypreal_mult_less_mono2";
+
+Goal "[| (0::hypreal)<=z; x<y |] ==> x*z<=y*z";
+by (EVERY1 [rtac hypreal_less_or_eq_imp_le, dtac order_le_imp_less_or_eq]);
+by (auto_tac (claset() addIs [hypreal_mult_less_mono1],simpset()));
+qed "hypreal_mult_le_less_mono1";
+
+Goal "[| (0::hypreal)<=z; x<y |] ==> z*x<=z*y";
+by (asm_simp_tac (simpset() addsimps [hypreal_mult_commute,
+				      hypreal_mult_le_less_mono1]) 1);
+qed "hypreal_mult_le_less_mono2";
+
+Goal "[| (0::hypreal)<=z; x<=y |] ==> z*x<=z*y";
+by (dres_inst_tac [("x","x")] order_le_imp_less_or_eq 1);
+by (auto_tac (claset() addIs [hypreal_mult_le_less_mono2], simpset()));
+qed "hypreal_mult_le_le_mono1";
+
+val prem1::prem2::prem3::rest = goal thy
+     "[| (0::hypreal)<y; x<r; y*r<t*s |] ==> y*x<t*s";
+by (rtac ([[prem1,prem2] MRS hypreal_mult_less_mono2, prem3] 
+          MRS order_less_trans) 1);
+qed "hypreal_mult_less_trans";
+
+Goal "[| 0<=y; x<r; y*r<t*s; (0::hypreal)<t*s|] ==> y*x<t*s";
+by (dtac order_le_imp_less_or_eq 1);
+by (fast_tac (HOL_cs addEs [hypreal_mult_0_less RS iffD2,
+                            hypreal_mult_less_trans]) 1);
+qed "hypreal_mult_le_less_trans";
+
+Goal "[| 0 <= y; x <= r; y*r < t*s; (0::hypreal) < t*s|] ==> y*x < t*s";
+by (dres_inst_tac [("x","x")] order_le_imp_less_or_eq 1);
+by (fast_tac (claset() addIs [hypreal_mult_le_less_trans]) 1);
+qed "hypreal_mult_le_le_trans";
+
+Goal "[| u<v;  x<y;  (0::hypreal) < v;  0 < x |] ==> u*x < v* y";
+by (etac (hypreal_mult_less_mono1 RS order_less_trans) 1);
+by (assume_tac 1);
+by (etac hypreal_mult_less_mono2 1);
+by (assume_tac 1);
+qed "hypreal_mult_less_mono";
+
+(*UNUSED at present but possibly more useful than hypreal_mult_less_mono*)
+Goal "[| x < y;  r1 < r2;  (0::hypreal) <= r1;  0 <= x|] ==> r1 * x < r2 * y";
+by (subgoal_tac "0<r2" 1);
+by (blast_tac (claset() addIs [order_le_less_trans]) 2); 
+by (case_tac "x=0" 1);
+by (auto_tac (claset() addSDs [order_le_imp_less_or_eq] 
+	               addIs [hypreal_mult_less_mono, hypreal_mult_order],
+	      simpset()));
+qed "hypreal_mult_less_mono'";
+
+Goal "0 < x ==> 0 < inverse (x::hypreal)";
+by (EVERY1[rtac ccontr, dtac hypreal_leI]);
+by (forward_tac [hypreal_minus_zero_less_iff2 RS iffD2] 1);
+by (forward_tac [hypreal_not_refl2 RS not_sym] 1);
+by (dtac (hypreal_not_refl2 RS not_sym RS hypreal_inverse_not_zero) 1);
+by (EVERY1[dtac order_le_imp_less_or_eq, Step_tac]); 
+by (dtac hypreal_mult_less_zero1 1 THEN assume_tac 1);
+by (auto_tac (claset() addIs [hypreal_zero_less_one RS hypreal_less_asym],
+                 simpset() addsimps [hypreal_minus_zero_less_iff]));
+qed "hypreal_inverse_gt_zero";
+
+Goal "x < 0 ==> inverse (x::hypreal) < 0";
+by (ftac hypreal_not_refl2 1);
+by (dtac (hypreal_minus_zero_less_iff RS iffD2) 1);
+by (rtac (hypreal_minus_zero_less_iff RS iffD1) 1);
+by (stac (hypreal_minus_inverse RS sym) 1);
+by (auto_tac (claset() addIs [hypreal_inverse_gt_zero],  simpset()));
+qed "hypreal_inverse_less_zero";
+
+Goal "[| 0 <= x; 0 <= y |] ==> (x+y = 0) = (x = 0 & y = (0::hypreal))";
+by (auto_tac (claset() addIs [order_antisym], simpset()));
+qed "hypreal_add_nonneg_eq_0_iff";
+
+Goal "(x*y = 0) = (x = 0 | y = (0::hypreal))";
+by Auto_tac;
+by (blast_tac (claset() addDs [hypreal_mult_zero_disj]) 1);
+qed "hypreal_mult_is_0";
+
+Goal "(x*x + y*y = 0) = (x = 0 & y = (0::hypreal))";
+by (simp_tac (HOL_ss addsimps [hypreal_le_square, hypreal_add_nonneg_eq_0_iff, 
+                               hypreal_mult_is_0]) 1);
+qed "hypreal_squares_add_zero_iff";
+Addsimps [hypreal_squares_add_zero_iff];
+
+Goal "(x*x + y*y + z*z = 0) = (x = 0 & y = 0 & z = (0::hypreal))";
+by (simp_tac (HOL_ss addsimps [hypreal_le_square, hypreal_le_add_order, 
+                         hypreal_add_nonneg_eq_0_iff, hypreal_mult_is_0]) 1);
+qed "hypreal_three_squares_add_zero_iff";
+
+Goal "(x::hypreal)*x <= x*x + y*y";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+              simpset() addsimps [hypreal_mult,hypreal_add,hypreal_le]));
+qed "hypreal_self_le_add_pos";
+Addsimps [hypreal_self_le_add_pos];
+
+(*lcp: new lemma unfortunately needed...*)
+Goal "-(x*x) <= (y*y::real)";
+by (rtac order_trans 1);
+by (rtac real_le_square 2);
+by Auto_tac;
+qed "minus_square_le_square";
+
+Goal "(x::hypreal)*x <= x*x + y*y + z*z";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+	      simpset() addsimps [hypreal_mult, hypreal_add, hypreal_le,
+				  minus_square_le_square]));
+qed "hypreal_self_le_add_pos2";
+Addsimps [hypreal_self_le_add_pos2];
+
+
+(*----------------------------------------------------------------------------
+             Embedding of the naturals in the hyperreals
+ ----------------------------------------------------------------------------*)
+Goalw [hypreal_of_posnat_def] "hypreal_of_posnat 0 = 1hr";
+by (full_simp_tac (simpset() addsimps 
+    [pnat_one_iff RS sym,real_of_preal_def]) 1);
+by (fold_tac [real_one_def]);
+by (simp_tac (simpset() addsimps [hypreal_of_real_one]) 1);
+qed "hypreal_of_posnat_one";
+
+Goalw [hypreal_of_posnat_def] "hypreal_of_posnat 1 = 1hr + 1hr";
+by (full_simp_tac (simpset() addsimps 
+		   [real_of_preal_def,
+		    rename_numerals (real_one_def RS meta_eq_to_obj_eq),
+		    hypreal_add,hypreal_of_real_def,pnat_two_eq,
+		    hypreal_one_def, real_add,
+		    prat_of_pnat_add RS sym, preal_of_prat_add RS sym] @ 
+		    pnat_add_ac) 1);
+qed "hypreal_of_posnat_two";
+
+(*FIXME: delete this and all posnat*)
+Goalw [hypreal_of_posnat_def]
+     "hypreal_of_posnat n1 + hypreal_of_posnat n2 = \
+\     hypreal_of_posnat (n1 + n2) + 1hr";
+by (full_simp_tac (HOL_ss addsimps [hypreal_of_posnat_one RS sym,
+    hypreal_of_real_add RS sym,hypreal_of_posnat_def,real_of_preal_add RS sym,
+    preal_of_prat_add RS sym,prat_of_pnat_add RS sym,pnat_of_nat_add]) 1);
+qed "hypreal_of_posnat_add";
+
+Goal "hypreal_of_posnat (n + 1) = hypreal_of_posnat n + 1hr";
+by (res_inst_tac [("x1","1hr")] (hypreal_add_right_cancel RS iffD1) 1);
+by (rtac (hypreal_of_posnat_add RS subst) 1);
+by (full_simp_tac (simpset() addsimps [hypreal_of_posnat_two,hypreal_add_assoc]) 1);
+qed "hypreal_of_posnat_add_one";
+
+Goalw [real_of_posnat_def,hypreal_of_posnat_def] 
+      "hypreal_of_posnat n = hypreal_of_real (real_of_posnat n)";
+by (rtac refl 1);
+qed "hypreal_of_real_of_posnat";
+
+Goalw [hypreal_of_posnat_def] 
+      "(n < m) = (hypreal_of_posnat n < hypreal_of_posnat m)";
+by Auto_tac;
+qed "hypreal_of_posnat_less_iff";
+
+Addsimps [hypreal_of_posnat_less_iff RS sym];
+(*---------------------------------------------------------------------------------
+               Existence of infinite hyperreal number
+ ---------------------------------------------------------------------------------*)
+
+Goal "hyprel^^{%n::nat. real_of_posnat n} : hypreal";
+by Auto_tac;
+qed "hypreal_omega";
+
+Goalw [omega_def] "Rep_hypreal(whr) : hypreal";
+by (rtac Rep_hypreal 1);
+qed "Rep_hypreal_omega";
+
+(* existence of infinite number not corresponding to any real number *)
+(* use assumption that member FreeUltrafilterNat is not finite       *)
+(* a few lemmas first *)
+
+Goal "{n::nat. x = real_of_posnat n} = {} | \
+\     (EX y. {n::nat. x = real_of_posnat n} = {y})";
+by (auto_tac (claset() addDs [inj_real_of_posnat RS injD],simpset()));
+qed "lemma_omega_empty_singleton_disj";
+
+Goal "finite {n::nat. x = real_of_posnat n}";
+by (cut_inst_tac [("x","x")] lemma_omega_empty_singleton_disj 1);
+by Auto_tac;
+qed "lemma_finite_omega_set";
+
+Goalw [omega_def,hypreal_of_real_def] 
+      "~ (EX x. hypreal_of_real x = whr)";
+by (auto_tac (claset(),simpset() addsimps [lemma_finite_omega_set 
+    RS FreeUltrafilterNat_finite]));
+qed "not_ex_hypreal_of_real_eq_omega";
+
+Goal "hypreal_of_real x ~= whr";
+by (cut_facts_tac [not_ex_hypreal_of_real_eq_omega] 1);
+by Auto_tac;
+qed "hypreal_of_real_not_eq_omega";
+
+(* existence of infinitesimal number also not *)
+(* corresponding to any real number *)
+
+Goal "{n::nat. x = inverse(real_of_posnat n)} = {} | \
+\     (EX y. {n::nat. x = inverse(real_of_posnat n)} = {y})";
+by (Step_tac 1 THEN Step_tac 1);
+by (auto_tac (claset() addIs [real_of_posnat_inverse_inj],simpset()));
+qed "lemma_epsilon_empty_singleton_disj";
+
+Goal "finite {n::nat. x = inverse(real_of_posnat n)}";
+by (cut_inst_tac [("x","x")] lemma_epsilon_empty_singleton_disj 1);
+by Auto_tac;
+qed "lemma_finite_epsilon_set";
+
+Goalw [epsilon_def,hypreal_of_real_def] 
+      "~ (EX x. hypreal_of_real x = ehr)";
+by (auto_tac (claset(),simpset() addsimps [lemma_finite_epsilon_set 
+    RS FreeUltrafilterNat_finite]));
+qed "not_ex_hypreal_of_real_eq_epsilon";
+
+Goal "hypreal_of_real x ~= ehr";
+by (cut_facts_tac [not_ex_hypreal_of_real_eq_epsilon] 1);
+by Auto_tac;
+qed "hypreal_of_real_not_eq_epsilon";
+
+Goalw [epsilon_def,hypreal_zero_def] "ehr ~= 0";
+by (auto_tac (claset(),
+     simpset() addsimps [rename_numerals real_of_posnat_not_eq_zero]));
+qed "hypreal_epsilon_not_zero";
+
+Addsimps [rename_numerals real_of_posnat_not_eq_zero];
+Goalw [omega_def,hypreal_zero_def] "whr ~= 0";
+by (Simp_tac 1);
+qed "hypreal_omega_not_zero";
+
+Goal "ehr = inverse(whr)";
+by (asm_full_simp_tac (simpset() addsimps 
+                     [hypreal_inverse, omega_def, epsilon_def]) 1);
+qed "hypreal_epsilon_inverse_omega";
+
+(*----------------------------------------------------------------
+     Another embedding of the naturals in the 
+    hyperreals (see hypreal_of_posnat)
+ ----------------------------------------------------------------*)
+Goalw [hypreal_of_nat_def] "hypreal_of_nat 0 = 0";
+by (full_simp_tac (simpset() addsimps [hypreal_of_posnat_one]) 1);
+qed "hypreal_of_nat_zero";
+
+Goalw [hypreal_of_nat_def] "hypreal_of_nat 1 = 1hr";
+by (full_simp_tac (simpset() addsimps [hypreal_of_posnat_two,
+                                       hypreal_add_assoc]) 1);
+qed "hypreal_of_nat_one";
+
+Goalw [hypreal_of_nat_def]
+     "hypreal_of_nat n1 + hypreal_of_nat n2 = hypreal_of_nat (n1 + n2)";
+by (full_simp_tac (simpset() addsimps hypreal_add_ac) 1);
+by (simp_tac (simpset() addsimps [hypreal_of_posnat_add,
+                                  hypreal_add_assoc RS sym]) 1);
+qed "hypreal_of_nat_add";
+
+Goal "hypreal_of_nat 2 = 1hr + 1hr";
+by (simp_tac (simpset() addsimps [hypreal_of_nat_one 
+    RS sym,hypreal_of_nat_add]) 1);
+qed "hypreal_of_nat_two";
+
+Goalw [hypreal_of_nat_def] 
+      "(n < m) = (hypreal_of_nat n < hypreal_of_nat m)";
+by (auto_tac (claset() addIs [hypreal_add_less_mono1],simpset()));
+qed "hypreal_of_nat_less_iff";
+Addsimps [hypreal_of_nat_less_iff RS sym];
+
+(*------------------------------------------------------------*)
+(* naturals embedded in hyperreals                            *)
+(* is a hyperreal c.f. NS extension                           *)
+(*------------------------------------------------------------*)
+
+Goalw [hypreal_of_nat_def,real_of_nat_def] 
+      "hypreal_of_nat  m = Abs_hypreal(hyprel^^{%n. real_of_nat m})";
+by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_def,
+    hypreal_of_real_of_posnat,hypreal_minus,hypreal_one_def,hypreal_add]));
+qed "hypreal_of_nat_iff";
+
+Goal "inj hypreal_of_nat";
+by (rtac injI 1);
+by (auto_tac (claset() addSDs [FreeUltrafilterNat_P],
+        simpset() addsimps [split_if_mem1, hypreal_of_nat_iff,
+        real_add_right_cancel,inj_real_of_nat RS injD]));
+qed "inj_hypreal_of_nat";
+
+Goalw [hypreal_of_nat_def,hypreal_of_real_def,hypreal_of_posnat_def,
+       real_of_posnat_def,hypreal_one_def,real_of_nat_def] 
+       "hypreal_of_nat n = hypreal_of_real (real_of_nat n)";
+by (simp_tac (simpset() addsimps [hypreal_add,hypreal_minus]) 1);
+qed "hypreal_of_nat_real_of_nat";
+
+Goal "hypreal_of_posnat (Suc n) = hypreal_of_posnat n + 1hr";
+by (stac (hypreal_of_posnat_add_one RS sym) 1);
+by (Simp_tac 1);
+qed "hypreal_of_posnat_Suc";
+
+Goalw [hypreal_of_nat_def] 
+      "hypreal_of_nat (Suc n) = hypreal_of_nat n + 1hr";
+by (simp_tac (simpset() addsimps [hypreal_of_posnat_Suc] @ hypreal_add_ac) 1);
+qed "hypreal_of_nat_Suc";
+
+(* this proof is so much simpler than one for reals!! *)
+Goal "[| 0 < r; r < x |] ==> inverse x < inverse (r::hypreal)";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [hypreal_inverse,
+    hypreal_less,hypreal_zero_def]));
+by (ultra_tac (claset() addIs [real_inverse_less_swap],simpset()) 1);
+qed "hypreal_inverse_less_swap";
+
+Goal "[| 0 < r; 0 < x|] ==> (r < x) = (inverse x < inverse (r::hypreal))";
+by (auto_tac (claset() addIs [hypreal_inverse_less_swap],simpset()));
+by (res_inst_tac [("t","r")] (hypreal_inverse_inverse RS subst) 1);
+by (res_inst_tac [("t","x")] (hypreal_inverse_inverse RS subst) 1);
+by (rtac hypreal_inverse_less_swap 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_inverse_gt_zero]));
+qed "hypreal_inverse_less_iff";
+
+Goal "[| 0 < z; x < y |] ==> x * inverse z < y * inverse (z::hypreal)";
+by (blast_tac (claset() addSIs [hypreal_mult_less_mono1,
+                                hypreal_inverse_gt_zero]) 1);
+qed "hypreal_mult_inverse_less_mono1";
+
+Goal "[| 0 < z; x < y |] ==> inverse z * x < inverse (z::hypreal) * y";
+by (blast_tac (claset() addSIs [hypreal_mult_less_mono2,
+                                hypreal_inverse_gt_zero]) 1);
+qed "hypreal_mult_inverse_less_mono2";
+
+Goal "[| (0::hypreal) < z; x*z < y*z |] ==> x < y";
+by (forw_inst_tac [("x","x*z")] hypreal_mult_inverse_less_mono1 1);
+by (dtac (hypreal_not_refl2 RS not_sym) 2);
+by (auto_tac (claset() addSDs [hypreal_mult_inverse],
+              simpset() addsimps hypreal_mult_ac));
+qed "hypreal_less_mult_right_cancel";
+
+Goal "[| (0::hypreal) < z; z*x < z*y |] ==> x < y";
+by (auto_tac (claset() addIs [hypreal_less_mult_right_cancel],
+    simpset() addsimps [hypreal_mult_commute]));
+qed "hypreal_less_mult_left_cancel";
+
+Goal "[| 0 < r; (0::hypreal) < ra; r < x; ra < y |] ==> r*ra < x*y";
+by (forw_inst_tac [("y","r")] order_less_trans 1);
+by (dres_inst_tac [("z","ra"),("x","r")] hypreal_mult_less_mono1 2);
+by (dres_inst_tac [("z","x"),("x","ra")] hypreal_mult_less_mono2 3);
+by (auto_tac (claset() addIs [order_less_trans], simpset()));
+qed "hypreal_mult_less_gt_zero"; 
+
+Goal "[| 0 < r; (0::hypreal) < ra; r <= x; ra <= y |] ==> r*ra <= x*y";
+by (REPEAT(dtac order_le_imp_less_or_eq 1));
+by (rtac hypreal_less_or_eq_imp_le 1);
+by (auto_tac (claset() addIs [hypreal_mult_less_mono1,
+                         hypreal_mult_less_mono2,hypreal_mult_less_gt_zero],
+    simpset()));
+qed "hypreal_mult_le_ge_zero"; 
+
+(*----------------------------------------------------------------------------
+     Some convenient biconditionals for products of signs 
+ ----------------------------------------------------------------------------*)
+
+Goal "((0::hypreal) < x*y) = (0 < x & 0 < y | x < 0 & y < 0)";
+ by (auto_tac (claset(), 
+               simpset() addsimps [order_le_less, linorder_not_less, 
+                              hypreal_mult_order, hypreal_mult_less_zero1]));
+by (ALLGOALS (rtac ccontr)); 
+by (auto_tac (claset(), 
+              simpset() addsimps [order_le_less, linorder_not_less]));
+by (ALLGOALS (etac rev_mp)); 
+by (ALLGOALS (dtac hypreal_mult_less_zero THEN' assume_tac));
+by (auto_tac (claset() addDs [order_less_not_sym], 
+              simpset() addsimps [hypreal_mult_commute]));  
+qed "hypreal_zero_less_mult_iff";
+
+Goal "((0::hypreal) <= x*y) = (0 <= x & 0 <= y | x <= 0 & y <= 0)";
+by (auto_tac (claset() addDs [hypreal_mult_zero_disj],
+              simpset() addsimps [order_le_less, linorder_not_less,
+                                  hypreal_zero_less_mult_iff]));
+qed "hypreal_zero_le_mult_iff";
+
+Goal "(x*y < (0::hypreal)) = (0 < x & y < 0 | x < 0 & 0 < y)";
+by (auto_tac (claset(), 
+              simpset() addsimps [hypreal_zero_le_mult_iff, 
+                                  linorder_not_le RS sym]));
+by (auto_tac (claset() addDs [order_less_not_sym],  
+              simpset() addsimps [linorder_not_le]));
+qed "hypreal_mult_less_zero_iff";
+
+Goal "(x*y <= (0::hypreal)) = (0 <= x & y <= 0 | x <= 0 & 0 <= y)";
+by (auto_tac (claset() addDs [order_less_not_sym], 
+              simpset() addsimps [hypreal_zero_less_mult_iff, 
+                                  linorder_not_less RS sym]));
+qed "hypreal_mult_le_zero_iff";
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HyperOrd.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,15 @@
+(*  Title:	 Real/Hyperreal/HyperOrd.thy
+    Author:      Jacques D. Fleuriot
+    Copyright:   2000 University of Edinburgh
+    Description: Type "hypreal" is a linear order and also 
+                 satisfies plus_ac0: + is an AC-operator with zero
+*)
+
+HyperOrd = HyperDef +
+
+instance hypreal :: order (hypreal_le_refl,hypreal_le_trans,hypreal_le_anti_sym,hypreal_less_le)
+instance hypreal :: linorder (hypreal_le_linear)
+
+instance hypreal :: plus_ac0(hypreal_add_commute,hypreal_add_assoc,hypreal_add_zero_left)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HyperPow.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,521 @@
+(*  Title       : HyperPow.ML
+    Author      : Jacques D. Fleuriot  
+    Copyright   : 1998  University of Cambridge
+    Description : Natural Powers of hyperreals theory
+
+*)
+
+Goal "(#0::hypreal) ^ (Suc n) = 0";
+by (Auto_tac);
+qed "hrealpow_zero";
+Addsimps [hrealpow_zero];
+
+Goal "r ~= (#0::hypreal) --> r ^ n ~= 0";
+by (induct_tac "n" 1);
+by Auto_tac;
+qed_spec_mp "hrealpow_not_zero";
+
+Goal "r ~= (#0::hypreal) --> inverse(r ^ n) = (inverse r) ^ n";
+by (induct_tac "n" 1);
+by (Auto_tac);
+by (forw_inst_tac [("n","n")] hrealpow_not_zero 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_inverse_distrib]));
+qed_spec_mp "hrealpow_inverse";
+
+Goal "abs (r::hypreal) ^ n = abs (r ^ n)";
+by (induct_tac "n" 1);
+by (auto_tac (claset(), simpset() addsimps [hrabs_mult]));
+qed "hrealpow_hrabs";
+
+Goal "(r::hypreal) ^ (n + m) = (r ^ n) * (r ^ m)";
+by (induct_tac "n" 1);
+by (auto_tac (claset(), simpset() addsimps hypreal_mult_ac));
+qed "hrealpow_add";
+
+Goal "(r::hypreal) ^ 1 = r";
+by (Simp_tac 1);
+qed "hrealpow_one";
+Addsimps [hrealpow_one];
+
+Goal "(r::hypreal) ^ 2 = r * r";
+by (Simp_tac 1);
+qed "hrealpow_two";
+
+Goal "(#0::hypreal) <= r --> #0 <= r ^ n";
+by (induct_tac "n" 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_0_le_mult_iff]));
+qed_spec_mp "hrealpow_ge_zero";
+
+Goal "(#0::hypreal) < r --> #0 < r ^ n";
+by (induct_tac "n" 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_0_less_mult_iff]));
+qed_spec_mp "hrealpow_gt_zero";
+
+Goal "x <= y & (#0::hypreal) < x --> x ^ n <= y ^ n";
+by (induct_tac "n" 1);
+by (auto_tac (claset() addSIs [hypreal_mult_le_mono], simpset()));
+by (asm_simp_tac (simpset() addsimps [hrealpow_ge_zero]) 1);
+qed_spec_mp "hrealpow_le";
+
+Goal "x < y & (#0::hypreal) < x & 0 < n --> x ^ n < y ^ n";
+by (induct_tac "n" 1);
+by (auto_tac (claset() addIs [hypreal_mult_less_mono,gr0I],
+              simpset() addsimps [hrealpow_gt_zero]));
+qed "hrealpow_less";
+
+Goal "#1 ^ n = (#1::hypreal)";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "hrealpow_eq_one";
+Addsimps [hrealpow_eq_one];
+
+Goal "abs(-(#1 ^ n)) = (#1::hypreal)";
+by Auto_tac;  
+qed "hrabs_minus_hrealpow_one";
+Addsimps [hrabs_minus_hrealpow_one];
+
+Goal "abs(#-1 ^ n) = (#1::hypreal)";
+by (induct_tac "n" 1);
+by Auto_tac;  
+qed "hrabs_hrealpow_minus_one";
+Addsimps [hrabs_hrealpow_minus_one];
+
+Goal "((r::hypreal) * s) ^ n = (r ^ n) * (s ^ n)";
+by (induct_tac "n" 1);
+by (auto_tac (claset(), simpset() addsimps hypreal_mult_ac));
+qed "hrealpow_mult";
+
+Goal "(#0::hypreal) <= r ^ 2";
+by (auto_tac (claset(), simpset() addsimps [hypreal_0_le_mult_iff]));
+qed "hrealpow_two_le";
+Addsimps [hrealpow_two_le];
+
+Goal "(#0::hypreal) <= u ^ 2 + v ^ 2";
+by (simp_tac (HOL_ss addsimps [hrealpow_two_le, 
+                    rename_numerals hypreal_le_add_order]) 1); 
+qed "hrealpow_two_le_add_order";
+Addsimps [hrealpow_two_le_add_order];
+
+Goal "(#0::hypreal) <= u ^ 2 + v ^ 2 + w ^ 2";
+by (simp_tac (HOL_ss addsimps [hrealpow_two_le, 
+                    rename_numerals hypreal_le_add_order]) 1); 
+qed "hrealpow_two_le_add_order2";
+Addsimps [hrealpow_two_le_add_order2];
+
+Goal "(x ^ 2 + y ^ 2 + z ^ 2 = (#0::hypreal)) = (x = #0 & y = #0 & z = #0)";
+by (simp_tac (HOL_ss addsimps
+      [rename_numerals hypreal_three_squares_add_zero_iff, hrealpow_two]) 1);
+qed "hrealpow_three_squares_add_zero_iff";
+Addsimps [hrealpow_three_squares_add_zero_iff];
+
+Goal "abs(x ^ 2) = (x::hypreal) ^ 2";
+by (auto_tac (claset(), 
+     simpset() addsimps [hrabs_def, hypreal_0_le_mult_iff])); 
+qed "hrabs_hrealpow_two";
+Addsimps [hrabs_hrealpow_two];
+
+Goal "abs(x) ^ 2 = (x::hypreal) ^ 2";
+by (simp_tac (simpset() addsimps [hrealpow_hrabs, hrabs_eqI1] 
+                        delsimps [hpowr_Suc]) 1);
+qed "hrealpow_two_hrabs";
+Addsimps [hrealpow_two_hrabs];
+
+Goal "(#1::hypreal) < r ==> #1 < r ^ 2";
+by (auto_tac (claset(), simpset() addsimps [hrealpow_two]));
+by (res_inst_tac [("y","#1*#1")] order_le_less_trans 1); 
+by (rtac hypreal_mult_less_mono 2); 
+by Auto_tac;  
+qed "hrealpow_two_gt_one";
+
+Goal "(#1::hypreal) <= r ==> #1 <= r ^ 2";
+by (etac (order_le_imp_less_or_eq RS disjE) 1);
+by (etac (hrealpow_two_gt_one RS order_less_imp_le) 1);
+by Auto_tac;  
+qed "hrealpow_two_ge_one";
+
+Goal "(#1::hypreal) <= #2 ^ n";
+by (res_inst_tac [("y","#1 ^ n")] order_trans 1);
+by (rtac hrealpow_le 2);
+by (auto_tac (claset() addIs [order_less_imp_le],
+         simpset() addsimps [hypreal_zero_less_one, hypreal_one_less_two]));
+qed "two_hrealpow_ge_one";
+
+Goal "hypreal_of_nat n < #2 ^ n";
+by (induct_tac "n" 1);
+by (auto_tac (claset(),
+        simpset() addsimps [hypreal_of_nat_Suc,hypreal_of_nat_zero,
+          hypreal_zero_less_one,hypreal_add_mult_distrib,hypreal_of_nat_one]));
+by (cut_inst_tac [("n","n")] two_hrealpow_ge_one 1);
+by (arith_tac 1);
+qed "two_hrealpow_gt";
+Addsimps [two_hrealpow_gt,two_hrealpow_ge_one];
+
+Goal "#-1 ^ (2*n) = (#1::hypreal)";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "hrealpow_minus_one";
+
+Goal "#-1 ^ (n + n) = (#1::hypreal)";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "hrealpow_minus_one2";
+Addsimps [hrealpow_minus_one2];
+
+Goal "(-(x::hypreal)) ^ 2 = x ^ 2";
+by (Auto_tac);
+qed "hrealpow_minus_two";
+Addsimps [hrealpow_minus_two];
+
+Goal "(#0::hypreal) < r & r < #1 --> r ^ Suc n < r ^ n";
+by (induct_tac "n" 1);
+by (auto_tac (claset(),
+              simpset() addsimps [hypreal_mult_less_mono2]));
+qed_spec_mp "hrealpow_Suc_less";
+
+Goal "(#0::hypreal) <= r & r < #1 --> r ^ Suc n <= r ^ n";
+by (induct_tac "n" 1);
+by (auto_tac (claset() addIs [order_less_imp_le]
+                       addSDs [order_le_imp_less_or_eq,hrealpow_Suc_less],
+              simpset() addsimps [hypreal_mult_less_mono2]));
+qed_spec_mp "hrealpow_Suc_le";
+
+Goal "Abs_hypreal(hyprel^^{%n. X n}) ^ m = Abs_hypreal(hyprel^^{%n. (X n) ^ m})";
+by (induct_tac "m" 1);
+by (auto_tac (claset(),
+              simpset() delsimps [one_eq_numeral_1]
+			addsimps [hypreal_one_def, hypreal_mult,
+                                  one_eq_numeral_1 RS sym]));
+qed "hrealpow";
+
+Goal "(x + (y::hypreal)) ^ 2 = \
+\     x ^ 2 + y ^ 2 + (hypreal_of_nat 2)*x*y";
+by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib2,
+               hypreal_add_mult_distrib,hypreal_of_nat_two] 
+                @ hypreal_add_ac @ hypreal_mult_ac) 1);
+qed "hrealpow_sum_square_expand";
+
+(*---------------------------------------------------------------
+   we'll prove the following theorem by going down to the
+   level of the ultrafilter and relying on the analogous
+   property for the real rather than prove it directly 
+   using induction: proof is much simpler this way!
+ ---------------------------------------------------------------*)
+Goal "[|(#0::hypreal) <= x; #0 <= y;x ^ Suc n <= y ^ Suc n |] ==> x <= y";
+by (full_simp_tac
+    (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+              simpset() addsimps [hrealpow,hypreal_le,hypreal_mult]));
+by (ultra_tac (claset() addIs [realpow_increasing], simpset()) 1);
+qed "hrealpow_increasing";
+
+(*By antisymmetry with the above we conclude x=y, replacing the deleted 
+  theorem hrealpow_Suc_cancel_eq*)
+
+Goal "x : HFinite --> x ^ n : HFinite";
+by (induct_tac "n" 1);
+by (auto_tac (claset() addIs [HFinite_mult], simpset()));
+qed_spec_mp "hrealpow_HFinite";
+
+(*---------------------------------------------------------------
+                  Hypernaturals Powers
+ --------------------------------------------------------------*)
+Goalw [congruent_def]
+     "congruent hyprel \
+\    (%X Y. hyprel^^{%n. ((X::nat=>real) n ^ (Y::nat=>nat) n)})";
+by (safe_tac (claset() addSIs [ext]));
+by (ALLGOALS(Fuf_tac));
+qed "hyperpow_congruent";
+
+Goalw [hyperpow_def]
+  "Abs_hypreal(hyprel^^{%n. X n}) pow Abs_hypnat(hypnatrel^^{%n. Y n}) = \
+\  Abs_hypreal(hyprel^^{%n. X n ^ Y n})";
+by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
+by (auto_tac (claset() addSIs [lemma_hyprel_refl,bexI],
+    simpset() addsimps [hyprel_in_hypreal RS 
+    Abs_hypreal_inverse,equiv_hyprel,hyperpow_congruent]));
+by (Fuf_tac 1);
+qed "hyperpow";
+
+Goalw [hypnat_one_def] "(#0::hypreal) pow (n + 1hn) = #0";
+by (simp_tac (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [hyperpow,hypnat_add]));
+qed "hyperpow_zero";
+Addsimps [hyperpow_zero];
+
+Goal "r ~= (#0::hypreal) --> r pow n ~= #0";
+by (simp_tac (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
+by (auto_tac (claset(), simpset() addsimps [hyperpow]));
+by (dtac FreeUltrafilterNat_Compl_mem 1);
+by (fuf_empty_tac (claset() addIs [realpow_not_zero RS notE],
+    simpset()) 1);
+qed_spec_mp "hyperpow_not_zero";
+
+Goal "r ~= (#0::hypreal) --> inverse(r pow n) = (inverse r) pow n";
+by (simp_tac (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem],
+              simpset() addsimps [hypreal_inverse,hyperpow]));
+by (rtac FreeUltrafilterNat_subset 1);
+by (auto_tac (claset() addDs [realpow_not_zero] 
+                       addIs [realpow_inverse], 
+              simpset()));
+qed "hyperpow_inverse";
+
+Goal "abs r pow n = abs (r pow n)";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+              simpset() addsimps [hypreal_hrabs, hyperpow,realpow_abs]));
+qed "hyperpow_hrabs";
+
+Goal "r pow (n + m) = (r pow n) * (r pow m)";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+          simpset() addsimps [hyperpow,hypnat_add, hypreal_mult,realpow_add]));
+qed "hyperpow_add";
+
+Goalw [hypnat_one_def] "r pow 1hn = r";
+by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
+by (auto_tac (claset(), simpset() addsimps [hyperpow]));
+qed "hyperpow_one";
+Addsimps [hyperpow_one];
+
+Goalw [hypnat_one_def] 
+      "r pow (1hn + 1hn) = r * r";
+by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+           simpset() addsimps [hyperpow,hypnat_add, hypreal_mult,realpow_two]));
+qed "hyperpow_two";
+
+Goal "(#0::hypreal) < r --> #0 < r pow n";
+by (simp_tac (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addSEs [FreeUltrafilterNat_subset, realpow_gt_zero],
+              simpset() addsimps [hyperpow,hypreal_less, hypreal_le]));
+qed_spec_mp "hyperpow_gt_zero";
+
+Goal "(#0::hypreal) < r --> #0 <= r pow n";
+by (blast_tac (claset() addSIs [hyperpow_gt_zero, order_less_imp_le]) 1);
+qed_spec_mp "hyperpow_ge_zero";
+
+Goal "(#0::hypreal) <= r --> #0 <= r pow n";
+by (simp_tac (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addSEs [FreeUltrafilterNat_subset, realpow_ge_zero2],
+              simpset() addsimps [hyperpow,hypreal_le]));
+qed "hyperpow_ge_zero2";
+
+Goal "(#0::hypreal) < x & x <= y --> x pow n <= y pow n";
+by (full_simp_tac
+    (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addIs [realpow_le,
+    (FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset)],
+    simpset() addsimps [hyperpow,hypreal_le,hypreal_less]));
+qed_spec_mp "hyperpow_le";
+
+Goal "#1 pow n = (#1::hypreal)";
+by (simp_tac (HOL_ss addsimps [one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [hyperpow]));
+qed "hyperpow_eq_one";
+Addsimps [hyperpow_eq_one];
+
+Goal "abs(-(#1 pow n)) = (#1::hypreal)";
+by (simp_tac (HOL_ss addsimps [one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [hyperpow,hypreal_hrabs]));
+qed "hrabs_minus_hyperpow_one";
+Addsimps [hrabs_minus_hyperpow_one];
+
+Goal "abs(#-1 pow n) = (#1::hypreal)";
+by (subgoal_tac "abs((-1hr) pow n) = 1hr" 1);
+by (Asm_full_simp_tac 1); 
+by (simp_tac (HOL_ss addsimps [hypreal_one_def]) 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+       simpset() addsimps [hyperpow,hypreal_minus,hypreal_hrabs]));
+qed "hrabs_hyperpow_minus_one";
+Addsimps [hrabs_hyperpow_minus_one];
+
+Goal "(r * s) pow n = (r pow n) * (s pow n)";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","s")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+       simpset() addsimps [hyperpow, hypreal_mult,realpow_mult]));
+qed "hyperpow_mult";
+
+Goal "(#0::hypreal) <= r pow (1hn + 1hn)";
+by (auto_tac (claset(), 
+              simpset() addsimps [hyperpow_two, hypreal_0_le_mult_iff]));
+qed "hyperpow_two_le";
+Addsimps [hyperpow_two_le];
+
+Goal "abs(x pow (1hn + 1hn)) = x pow (1hn + 1hn)";
+by (simp_tac (simpset() addsimps [hrabs_eqI1]) 1);
+qed "hrabs_hyperpow_two";
+Addsimps [hrabs_hyperpow_two];
+
+Goal "abs(x) pow (1hn + 1hn)  = x pow (1hn + 1hn)";
+by (simp_tac (simpset() addsimps [hyperpow_hrabs,hrabs_eqI1]) 1);
+qed "hyperpow_two_hrabs";
+Addsimps [hyperpow_two_hrabs];
+
+(*? very similar to hrealpow_two_gt_one *)
+Goal "(#1::hypreal) < r ==> #1 < r pow (1hn + 1hn)";
+by (auto_tac (claset(), simpset() addsimps [hyperpow_two]));
+by (res_inst_tac [("y","#1*#1")] order_le_less_trans 1); 
+by (rtac hypreal_mult_less_mono 2); 
+by Auto_tac;  
+qed "hyperpow_two_gt_one";
+
+Goal "(#1::hypreal) <= r ==> #1 <= r pow (1hn + 1hn)";
+by (auto_tac (claset() addSDs [order_le_imp_less_or_eq] 
+                       addIs [hyperpow_two_gt_one,order_less_imp_le],
+              simpset()));
+qed "hyperpow_two_ge_one";
+
+Goal "(#1::hypreal) <= #2 pow n";
+by (res_inst_tac [("y","#1 pow n")] order_trans 1);
+by (rtac hyperpow_le 2);
+by (auto_tac (claset() addIs [order_less_imp_le],
+          simpset() addsimps [hypreal_zero_less_one, hypreal_one_less_two]));
+qed "two_hyperpow_ge_one";
+Addsimps [two_hyperpow_ge_one];
+
+Addsimps [simplify (simpset()) realpow_minus_one];
+
+Goal "#-1 pow ((1hn + 1hn)*n) = (#1::hypreal)";
+by (subgoal_tac "(-(1hr)) pow ((1hn + 1hn)*n) = 1hr" 1);
+by (Asm_full_simp_tac 1); 
+by (simp_tac (HOL_ss addsimps [hypreal_one_def]) 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+              simpset() addsimps [hyperpow, hypnat_add,hypreal_minus]));
+qed "hyperpow_minus_one2";
+Addsimps [hyperpow_minus_one2];
+
+Goalw [hypnat_one_def]
+     "(#0::hypreal) < r & r < #1 --> r pow (n + 1hn) < r pow n";
+by (full_simp_tac
+    (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def,
+                      one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addSDs [conjI RS realpow_Suc_less] 
+                  addEs [FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset],
+              simpset() addsimps [hyperpow,hypreal_less,hypnat_add]));
+qed_spec_mp "hyperpow_Suc_less";
+
+Goalw [hypnat_one_def]
+     "#0 <= r & r < (#1::hypreal) --> r pow (n + 1hn) <= r pow n";
+by (full_simp_tac
+    (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def,
+                      one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addSDs [conjI RS realpow_Suc_le] addEs
+    [FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset ],
+    simpset() addsimps [hyperpow,hypreal_le,hypnat_add,
+    hypreal_less]));
+qed_spec_mp "hyperpow_Suc_le";
+
+Goalw [hypnat_one_def]
+     "(#0::hypreal) <= r & r < #1 & n < N --> r pow N <= r pow n";
+by (full_simp_tac
+    (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def,
+                      one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+        simpset() addsimps [hyperpow, hypreal_le,hypreal_less,hypnat_less]));
+by (etac (FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset) 1);
+by (etac FreeUltrafilterNat_Int 1);
+by (auto_tac (claset() addSDs [conjI RS realpow_less_le],
+    simpset()));
+qed_spec_mp "hyperpow_less_le";
+
+Goal "[| (#0::hypreal) <= r; r < #1 |]  \
+\     ==> ALL N n. n < N --> r pow N <= r pow n";
+by (blast_tac (claset() addSIs [hyperpow_less_le]) 1);
+qed "hyperpow_less_le2";
+
+Goal "[| #0 <= r;  r < (#1::hypreal);  N : HNatInfinite |]  \
+\     ==> ALL n:SHNat. r pow N <= r pow n";
+by (auto_tac (claset() addSIs [hyperpow_less_le],
+              simpset() addsimps [HNatInfinite_iff]));
+qed "hyperpow_SHNat_le";
+
+Goalw [hypreal_of_real_def,hypnat_of_nat_def] 
+      "(hypreal_of_real r) pow (hypnat_of_nat n) = hypreal_of_real (r ^ n)";
+by (auto_tac (claset(), simpset() addsimps [hyperpow]));
+qed "hyperpow_realpow";
+
+Goalw [SReal_def]
+     "(hypreal_of_real r) pow (hypnat_of_nat n) : SReal";
+by (auto_tac (claset(), simpset() addsimps [hyperpow_realpow]));
+qed "hyperpow_SReal";
+Addsimps [hyperpow_SReal];
+
+Goal "N : HNatInfinite ==> (#0::hypreal) pow N = 0";
+by (dtac HNatInfinite_is_Suc 1);
+by (Auto_tac);
+qed "hyperpow_zero_HNatInfinite";
+Addsimps [hyperpow_zero_HNatInfinite];
+
+Goal "[| (#0::hypreal) <= r; r < #1; n <= N |] ==> r pow N <= r pow n";
+by (dres_inst_tac [("y","N")] hypnat_le_imp_less_or_eq 1);
+by (auto_tac (claset() addIs [hyperpow_less_le], simpset()));
+qed "hyperpow_le_le";
+
+Goal "[| (#0::hypreal) < r; r < #1 |] ==> r pow (n + 1hn) <= r";
+by (dres_inst_tac [("n","1hn")] (order_less_imp_le RS hyperpow_le_le) 1);
+by (Auto_tac);
+qed "hyperpow_Suc_le_self";
+
+Goal "[| (#0::hypreal) <= r; r < #1 |] ==> r pow (n + 1hn) <= r";
+by (dres_inst_tac [("n","1hn")] hyperpow_le_le 1);
+by (Auto_tac);
+qed "hyperpow_Suc_le_self2";
+
+Goalw [Infinitesimal_def]
+     "[| x : Infinitesimal; 0 < N |] ==> (abs x) pow N <= abs x";
+by (auto_tac (claset() addSIs [hyperpow_Suc_le_self2],
+    simpset() addsimps [hyperpow_hrabs RS sym,
+                        hypnat_gt_zero_iff2, hrabs_ge_zero,
+                        hypreal_zero_less_one]));
+qed "lemma_Infinitesimal_hyperpow";
+
+Goal "[| x : Infinitesimal; 0 < N |] ==> x pow N : Infinitesimal";
+by (rtac hrabs_le_Infinitesimal 1);
+by (dtac Infinitesimal_hrabs 1);
+by (auto_tac (claset() addSIs [lemma_Infinitesimal_hyperpow],
+    simpset() addsimps [hyperpow_hrabs RS sym]));
+qed "Infinitesimal_hyperpow";
+
+Goalw [hypnat_of_nat_def] 
+     "(x ^ n : Infinitesimal) = (x pow (hypnat_of_nat n) : Infinitesimal)";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(), simpset() addsimps [hrealpow, hyperpow]));
+qed "hrealpow_hyperpow_Infinitesimal_iff";
+
+Goal "[| x : Infinitesimal; 0 < n |] ==> x ^ n : Infinitesimal";
+by (auto_tac (claset() addSIs [Infinitesimal_hyperpow],
+    simpset() addsimps [hrealpow_hyperpow_Infinitesimal_iff,
+                        hypnat_of_nat_less_iff,hypnat_of_nat_zero] 
+              delsimps [hypnat_of_nat_less_iff RS sym]));
+qed "Infinitesimal_hrealpow";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/HyperPow.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,26 @@
+(*  Title       : HyperPow.thy
+    Author      : Jacques D. Fleuriot  
+    Copyright   : 1998  University of Cambridge
+    Description : Powers theory for hyperreals
+*)
+
+HyperPow = HRealAbs + HyperNat + RealPow +  
+
+instance hypreal :: {power}
+
+consts hpowr :: "[hypreal,nat] => hypreal"  
+primrec
+     hpowr_0   "r ^ 0       = 1hr"
+     hpowr_Suc "r ^ (Suc n) = (r::hypreal) * (r ^ n)"
+
+consts
+  "pow"  :: [hypreal,hypnat] => hypreal     (infixr 80)
+
+defs
+
+  (* hypernatural powers of hyperreals *)
+  hyperpow_def
+  "(R::hypreal) pow (N::hypnat) 
+      == Abs_hypreal(UN X:Rep_hypreal(R). UN Y: Rep_hypnat(N).
+             hyprel^^{%n::nat. (X n) ^ (Y n)})"
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/Hyperreal.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,4 @@
+
+theory Hyperreal = HSeries:
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/Lim.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,2324 @@
+(*  Title       : Lim.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Theory of limits, continuity and 
+                  differentiation of real=>real functions
+*)
+
+
+fun ARITH_PROVE str = prove_goal thy str 
+                      (fn prems => [cut_facts_tac prems 1,arith_tac 1]);
+
+
+(*---------------------------------------------------------------
+   Theory of limits, continuity and differentiation of 
+   real=>real functions 
+ ----------------------------------------------------------------*)
+
+Goalw [LIM_def] "(%x. k) -- x --> k";
+by Auto_tac;
+qed "LIM_const";
+Addsimps [LIM_const];
+
+(***-----------------------------------------------------------***)
+(***  Some Purely Standard Proofs - Can be used for comparison ***)
+(***-----------------------------------------------------------***)
+ 
+(*--------------- 
+    LIM_add    
+ ---------------*)
+Goalw [LIM_def] 
+     "[| f -- x --> l; g -- x --> m |] \
+\     ==> (%x. f(x) + g(x)) -- x --> (l + m)";
+by (Step_tac 1);
+by (REPEAT(dres_inst_tac [("x","r/#2")] spec 1));
+by (dtac (rename_numerals (real_zero_less_two RS real_inverse_gt_zero 
+          RSN (2,real_mult_less_mono2))) 1);
+by (Asm_full_simp_tac 1);
+by (Clarify_tac 1);
+by (res_inst_tac [("R1.0","s"),("R2.0","sa")] 
+    real_linear_less2 1);
+by (res_inst_tac [("x","s")] exI 1);
+by (res_inst_tac [("x","sa")] exI 2);
+by (res_inst_tac [("x","sa")] exI 3);
+by (Step_tac 1);
+by (REPEAT(dres_inst_tac [("x","xa")] spec 1) 
+    THEN step_tac (claset() addSEs [order_less_trans]) 1);
+by (REPEAT(dres_inst_tac [("x","xa")] spec 2) 
+    THEN step_tac (claset() addSEs [order_less_trans]) 2);
+by (REPEAT(dres_inst_tac [("x","xa")] spec 3) 
+    THEN step_tac (claset() addSEs [order_less_trans]) 3);
+by (ALLGOALS(rtac (abs_sum_triangle_ineq RS order_le_less_trans)));
+by (ALLGOALS(rtac (real_sum_of_halves RS subst)));
+by (auto_tac (claset() addIs [real_add_less_mono],simpset()));
+qed "LIM_add";
+
+Goalw [LIM_def] "f -- a --> L ==> (%x. -f(x)) -- a --> -L";
+by (full_simp_tac (simpset() addsimps [real_minus_add_distrib RS sym] 
+                    delsimps [real_minus_add_distrib, real_minus_minus]) 1);
+qed "LIM_minus";
+
+(*----------------------------------------------
+     LIM_add_minus
+ ----------------------------------------------*)
+Goal "[| f -- x --> l; g -- x --> m |] \
+\     ==> (%x. f(x) + -g(x)) -- x --> (l + -m)";
+by (blast_tac (claset() addDs [LIM_add,LIM_minus]) 1);
+qed "LIM_add_minus";
+
+(*----------------------------------------------
+     LIM_zero
+ ----------------------------------------------*)
+Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> #0";
+by (res_inst_tac [("z1","l")] (rename_numerals (real_add_minus RS subst)) 1);
+by (rtac LIM_add_minus 1 THEN Auto_tac);
+qed "LIM_zero";
+
+(*--------------------------
+   Limit not zero
+ --------------------------*)
+Goalw [LIM_def] "k ~= #0 ==> ~ ((%x. k) -- x --> #0)";
+by (res_inst_tac [("R1.0","k"),("R2.0","#0")] real_linear_less2 1);
+by (auto_tac (claset(), simpset() addsimps [real_abs_def]));
+by (res_inst_tac [("x","-k")] exI 1);
+by (res_inst_tac [("x","k")] exI 2);
+by Auto_tac;
+by (ALLGOALS(dres_inst_tac [("y","s")] real_dense));
+by Safe_tac;
+by (ALLGOALS(res_inst_tac [("x","r + x")] exI));
+by Auto_tac;  
+qed "LIM_not_zero";
+
+(* [| k ~= #0; (%x. k) -- x --> #0 |] ==> R *)
+bind_thm("LIM_not_zeroE", LIM_not_zero RS notE);
+
+Goal "(%x. k) -- x --> L ==> k = L";
+by (rtac ccontr 1);
+by (dtac LIM_zero 1);
+by (rtac LIM_not_zeroE 1 THEN assume_tac 2);
+by (arith_tac 1);
+qed "LIM_const_eq";
+
+(*------------------------
+     Limit is Unique
+ ------------------------*)
+Goal "[| f -- x --> L; f -- x --> M |] ==> L = M";
+by (dtac LIM_minus 1);
+by (dtac LIM_add 1 THEN assume_tac 1);
+by (auto_tac (claset() addSDs [LIM_const_eq RS sym],  simpset()));
+qed "LIM_unique";
+
+(*-------------
+    LIM_mult_zero
+ -------------*)
+Goalw [LIM_def] "[| f -- x --> #0; g -- x --> #0 |] \
+\         ==> (%x. f(x)*g(x)) -- x --> #0";
+by (Step_tac 1);
+by (dres_inst_tac [("x","#1")] spec 1);
+by (dres_inst_tac [("x","r")] spec 1);
+by (cut_facts_tac [real_zero_less_one] 1);
+by (asm_full_simp_tac (simpset() addsimps 
+    [abs_mult]) 1);
+by (Clarify_tac 1);
+by (res_inst_tac [("R1.0","s"),("R2.0","sa")] 
+    real_linear_less2 1);
+by (res_inst_tac [("x","s")] exI 1);
+by (res_inst_tac [("x","sa")] exI 2);
+by (res_inst_tac [("x","sa")] exI 3);
+by (Step_tac 1);
+by (REPEAT(dres_inst_tac [("x","xa")] spec 1) 
+    THEN step_tac (claset() addSEs [order_less_trans]) 1);
+by (REPEAT(dres_inst_tac [("x","xa")] spec 2) 
+    THEN step_tac (claset() addSEs [order_less_trans]) 2);
+by (REPEAT(dres_inst_tac [("x","xa")] spec 3) 
+    THEN step_tac (claset() addSEs [order_less_trans]) 3);
+by (ALLGOALS(res_inst_tac [("t","r")] (real_mult_1 RS subst)));
+by (ALLGOALS(rtac abs_mult_less2));
+by Auto_tac;
+qed "LIM_mult_zero";
+
+Goalw [LIM_def] "(%x. x) -- a --> a";
+by Auto_tac;
+qed "LIM_self";
+
+(*--------------------------------------------------------------
+   Limits are equal for functions equal except at limit point
+ --------------------------------------------------------------*)
+Goalw [LIM_def] 
+      "[| ALL x. x ~= a --> (f x = g x) |] \
+\           ==> (f -- a --> l) = (g -- a --> l)";
+by (auto_tac (claset(), simpset() addsimps [real_add_minus_iff]));
+qed "LIM_equal";
+
+Goal "[| (%x. f(x) + -g(x)) -- a --> #0; \
+\        g -- a --> l |] \
+\      ==> f -- a --> l";
+by (dtac LIM_add 1 THEN assume_tac 1);
+by (auto_tac (claset(), simpset() addsimps [real_add_assoc]));
+qed "LIM_trans";
+
+(***-------------------------------------------------------------***)
+(***           End of Purely Standard Proofs                     ***)
+(***-------------------------------------------------------------***)
+(*--------------------------------------------------------------
+       Standard and NS definitions of Limit
+ --------------------------------------------------------------*)
+Goalw [LIM_def,NSLIM_def,inf_close_def] 
+      "f -- x --> L ==> f -- x --NS> L";
+by (asm_full_simp_tac
+    (simpset() addsimps [Infinitesimal_FreeUltrafilterNat_iff]) 1);
+by (Step_tac 1);
+by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+      simpset() addsimps [real_add_minus_iff, starfun, hypreal_minus, 
+                          hypreal_of_real_def, hypreal_add]));
+by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2 THEN Step_tac 1);
+by (dres_inst_tac [("x","u")] spec 1 THEN Clarify_tac 1);
+by (dres_inst_tac [("x","s")] spec 1 THEN Clarify_tac 1);
+by (subgoal_tac "ALL n::nat. (xa n) ~= x & \
+\                    abs ((xa n) + - x) < s --> abs (f (xa n) + - L) < u" 1);
+by (Blast_tac 2);
+by (dtac FreeUltrafilterNat_all 1);
+by (Ultra_tac 1);
+qed "LIM_NSLIM";
+ 
+(*---------------------------------------------------------------------
+    Limit: NS definition ==> standard definition
+ ---------------------------------------------------------------------*)
+
+Goal "ALL s. #0 < s --> (EX xa.  xa ~= x & \
+\        abs (xa + - x) < s  & r <= abs (f xa + -L)) \
+\     ==> ALL n::nat. EX xa.  xa ~= x & \
+\             abs(xa + -x) < inverse(real_of_posnat n) & r <= abs(f xa + -L)";
+by (Step_tac 1);
+by (cut_inst_tac [("n1","n")]
+    (real_of_posnat_gt_zero RS real_inverse_gt_zero) 1);
+by Auto_tac;
+val lemma_LIM = result();
+
+Goal "ALL s. #0 < s --> (EX xa.  xa ~= x & \
+\        abs (xa + - x) < s  & r <= abs (f xa + -L)) \
+\     ==> EX X. ALL n::nat. X n ~= x & \
+\               abs(X n + -x) < inverse(real_of_posnat n) & r <= abs(f (X n) + -L)";
+by (dtac lemma_LIM 1);
+by (dtac choice 1);
+by (Blast_tac 1);
+val lemma_skolemize_LIM2 = result();
+
+Goal "ALL n. X n ~= x & \
+\         abs (X n + - x) < inverse (real_of_posnat  n) & \
+\         r <= abs (f (X n) + - L) ==> \
+\         ALL n. abs (X n + - x) < inverse (real_of_posnat  n)";
+by (Auto_tac );
+val lemma_simp = result();
+ 
+(*-------------------
+    NSLIM => LIM
+ -------------------*)
+
+Goalw [LIM_def,NSLIM_def,inf_close_def] 
+      "f -- x --NS> L ==> f -- x --> L";
+by (asm_full_simp_tac
+    (simpset() addsimps [Infinitesimal_FreeUltrafilterNat_iff]) 1);
+by (EVERY1[Step_tac, rtac ccontr, Asm_full_simp_tac]);
+by (fold_tac [real_le_def]);
+by (dtac lemma_skolemize_LIM2 1);
+by (Step_tac 1);
+by (dres_inst_tac [("x","Abs_hypreal(hyprel^^{X})")] spec 1);
+by (asm_full_simp_tac
+    (simpset() addsimps [starfun, hypreal_minus, 
+                         hypreal_of_real_def,hypreal_add]) 1);
+by (Step_tac 1);
+by (dtac (lemma_simp RS real_seq_to_hypreal_Infinitesimal) 1);
+by (asm_full_simp_tac
+    (simpset() addsimps 
+       [Infinitesimal_FreeUltrafilterNat_iff,hypreal_of_real_def,
+        hypreal_minus, hypreal_add]) 1);
+by (Blast_tac 1); 
+by (rotate_tac 2 1);
+by (dres_inst_tac [("x","r")] spec 1);
+by (Clarify_tac 1);
+by (dtac FreeUltrafilterNat_all 1);
+by (Ultra_tac 1);
+qed "NSLIM_LIM";
+
+
+(**** Key result ****)
+Goal "(f -- x --> L) = (f -- x --NS> L)";
+by (blast_tac (claset() addIs [LIM_NSLIM,NSLIM_LIM]) 1);
+qed "LIM_NSLIM_iff";
+
+(*-------------------------------------------------------------------*)
+(*   Proving properties of limits using nonstandard definition and   *)
+(*   hence, the properties hold for standard limits as well          *)
+(*-------------------------------------------------------------------*)
+(*------------------------------------------------
+      NSLIM_mult and hence (trivially) LIM_mult
+ ------------------------------------------------*)
+
+Goalw [NSLIM_def]
+     "[| f -- x --NS> l; g -- x --NS> m |] \
+\     ==> (%x. f(x) * g(x)) -- x --NS> (l * m)";
+by (auto_tac (claset() addSIs [inf_close_mult_HFinite],  simpset()));
+qed "NSLIM_mult";
+
+Goal "[| f -- x --> l; g -- x --> m |] \
+\     ==> (%x. f(x) * g(x)) -- x --> (l * m)";
+by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_mult]) 1);
+qed "LIM_mult2";
+
+(*----------------------------------------------
+      NSLIM_add and hence (trivially) LIM_add
+      Note the much shorter proof
+ ----------------------------------------------*)
+Goalw [NSLIM_def]
+     "[| f -- x --NS> l; g -- x --NS> m |] \
+\     ==> (%x. f(x) + g(x)) -- x --NS> (l + m)";
+by (auto_tac (claset() addSIs [inf_close_add], simpset()));
+qed "NSLIM_add";
+
+Goal "[| f -- x --> l; g -- x --> m |] \
+\     ==> (%x. f(x) + g(x)) -- x --> (l + m)";
+by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_add]) 1);
+qed "LIM_add2";
+
+(*----------------------------------------------
+     NSLIM_const
+ ----------------------------------------------*)
+Goalw [NSLIM_def] "(%x. k) -- x --NS> k";
+by Auto_tac;
+qed "NSLIM_const";
+
+Addsimps [NSLIM_const];
+
+Goal "(%x. k) -- x --> k";
+by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);
+qed "LIM_const2";
+
+(*----------------------------------------------
+     NSLIM_minus
+ ----------------------------------------------*)
+Goalw [NSLIM_def] 
+      "f -- a --NS> L ==> (%x. -f(x)) -- a --NS> -L";
+by Auto_tac;  
+qed "NSLIM_minus";
+
+Goal "f -- a --> L ==> (%x. -f(x)) -- a --> -L";
+by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_minus]) 1);
+qed "LIM_minus2";
+
+(*----------------------------------------------
+     NSLIM_add_minus
+ ----------------------------------------------*)
+Goal "[| f -- x --NS> l; g -- x --NS> m |] \
+\     ==> (%x. f(x) + -g(x)) -- x --NS> (l + -m)";
+by (blast_tac (claset() addDs [NSLIM_add,NSLIM_minus]) 1);
+qed "NSLIM_add_minus";
+
+Goal "[| f -- x --> l; g -- x --> m |] \
+\     ==> (%x. f(x) + -g(x)) -- x --> (l + -m)";
+by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
+    NSLIM_add_minus]) 1);
+qed "LIM_add_minus2";
+
+(*-----------------------------
+    NSLIM_inverse
+ -----------------------------*)
+Goalw [NSLIM_def] 
+     "[| f -- a --NS> L;  L ~= #0 |] \
+\     ==> (%x. inverse(f(x))) -- a --NS> (inverse L)";
+by (Clarify_tac 1);
+by (dtac spec 1);
+by (auto_tac (claset(), 
+              simpset() addsimps [hypreal_of_real_inf_close_inverse]));  
+qed "NSLIM_inverse";
+
+Goal "[| f -- a --> L; \
+\        L ~= #0 |] ==> (%x. inverse(f(x))) -- a --> (inverse L)";
+by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_inverse]) 1);
+qed "LIM_inverse";
+
+(*------------------------------
+    NSLIM_zero
+ ------------------------------*)
+Goal "f -- a --NS> l ==> (%x. f(x) + -l) -- a --NS> #0";
+by (res_inst_tac [("z1","l")] (rename_numerals (real_add_minus RS subst)) 1);
+by (rtac NSLIM_add_minus 1 THEN Auto_tac);
+qed "NSLIM_zero";
+
+Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> #0";
+by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_zero]) 1);
+qed "LIM_zero2";
+
+Goal "(%x. f(x) - l) -- x --NS> #0 ==> f -- x --NS> l";
+by (dres_inst_tac [("g","%x. l"),("m","l")] NSLIM_add 1);
+by (auto_tac (claset(),simpset() addsimps [real_diff_def, real_add_assoc]));
+qed "NSLIM_zero_cancel";
+
+Goal "(%x. f(x) - l) -- x --> #0 ==> f -- x --> l";
+by (dres_inst_tac [("g","%x. l"),("m","l")] LIM_add 1);
+by (auto_tac (claset(),simpset() addsimps [real_diff_def, real_add_assoc]));
+qed "LIM_zero_cancel";
+
+
+(*--------------------------
+   NSLIM_not_zero
+ --------------------------*)
+Goalw [NSLIM_def] "k ~= #0 ==> ~ ((%x. k) -- x --NS> #0)";
+by Auto_tac;
+by (res_inst_tac [("x","hypreal_of_real x + ehr")] exI 1);
+by (auto_tac (claset() addIs [Infinitesimal_add_inf_close_self RS inf_close_sym],
+              simpset() addsimps [rename_numerals hypreal_epsilon_not_zero]));
+qed "NSLIM_not_zero";
+
+(* [| k ~= #0; (%x. k) -- x --NS> #0 |] ==> R *)
+bind_thm("NSLIM_not_zeroE", NSLIM_not_zero RS notE);
+
+Goal "k ~= #0 ==> ~ ((%x. k) -- x --> #0)";
+by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_not_zero]) 1);
+qed "LIM_not_zero2";
+
+(*-------------------------------------
+   NSLIM of constant function
+ -------------------------------------*)
+Goal "(%x. k) -- x --NS> L ==> k = L";
+by (rtac ccontr 1);
+by (dtac NSLIM_zero 1);
+by (rtac NSLIM_not_zeroE 1 THEN assume_tac 2);
+by (arith_tac 1);
+qed "NSLIM_const_eq";
+
+Goal "(%x. k) -- x --> L ==> k = L";
+by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
+    NSLIM_const_eq]) 1);
+qed "LIM_const_eq2";
+
+(*------------------------
+     NS Limit is Unique
+ ------------------------*)
+(* can actually be proved more easily by unfolding def! *)
+Goal "[| f -- x --NS> L; f -- x --NS> M |] ==> L = M";
+by (dtac NSLIM_minus 1);
+by (dtac NSLIM_add 1 THEN assume_tac 1);
+by (auto_tac (claset() addSDs [NSLIM_const_eq RS sym], simpset()));
+qed "NSLIM_unique";
+
+Goal "[| f -- x --> L; f -- x --> M |] ==> L = M";
+by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_unique]) 1);
+qed "LIM_unique2";
+
+(*--------------------
+    NSLIM_mult_zero
+ --------------------*)
+Goal "[| f -- x --NS> #0; g -- x --NS> #0 |] \
+\         ==> (%x. f(x)*g(x)) -- x --NS> #0";
+by (dtac NSLIM_mult 1 THEN Auto_tac);
+qed "NSLIM_mult_zero";
+
+(* we can use the corresponding thm LIM_mult2 *)
+(* for standard definition of limit           *)
+
+Goal "[| f -- x --> #0; g -- x --> #0 |] \
+\     ==> (%x. f(x)*g(x)) -- x --> #0";
+by (dtac LIM_mult2 1 THEN Auto_tac);
+qed "LIM_mult_zero2";
+
+(*----------------------------
+    NSLIM_self
+ ----------------------------*)
+Goalw [NSLIM_def] "(%x. x) -- a --NS> a";
+by (auto_tac (claset() addIs [starfun_Idfun_inf_close],simpset()));
+qed "NSLIM_self";
+
+Goal "(%x. x) -- a --> a";
+by (simp_tac (simpset() addsimps [LIM_NSLIM_iff,NSLIM_self]) 1);
+qed "LIM_self2";
+
+(*-----------------------------------------------------------------------------
+   Derivatives and Continuity - NS and Standard properties
+ -----------------------------------------------------------------------------*)
+(*---------------
+    Continuity 
+ ---------------*)
+
+Goalw [isNSCont_def] 
+      "[| isNSCont f a; y @= hypreal_of_real a |] \
+\           ==> (*f* f) y @= hypreal_of_real (f a)";
+by (Blast_tac 1);
+qed "isNSContD";
+
+Goalw [isNSCont_def,NSLIM_def] 
+      "isNSCont f a ==> f -- a --NS> (f a) ";
+by (Blast_tac 1);
+qed "isNSCont_NSLIM";
+
+Goalw [isNSCont_def,NSLIM_def] 
+      "f -- a --NS> (f a) ==> isNSCont f a";
+by Auto_tac;
+by (res_inst_tac [("Q","y = hypreal_of_real a")] 
+    (excluded_middle RS disjE) 1);
+by Auto_tac;
+qed "NSLIM_isNSCont";
+
+(*-----------------------------------------------------
+    NS continuity can be defined using NS Limit in
+    similar fashion to standard def of continuity
+ -----------------------------------------------------*)
+Goal "(isNSCont f a) = (f -- a --NS> (f a))";
+by (blast_tac (claset() addIs [isNSCont_NSLIM,NSLIM_isNSCont]) 1);
+qed "isNSCont_NSLIM_iff";
+
+(*----------------------------------------------
+  Hence, NS continuity can be given
+  in terms of standard limit
+ ---------------------------------------------*)
+Goal "(isNSCont f a) = (f -- a --> (f a))";
+by (asm_full_simp_tac (simpset() addsimps 
+    [LIM_NSLIM_iff,isNSCont_NSLIM_iff]) 1);
+qed "isNSCont_LIM_iff";
+
+(*-----------------------------------------------
+  Moreover, it's trivial now that NS continuity 
+  is equivalent to standard continuity
+ -----------------------------------------------*)
+Goalw [isCont_def] "(isNSCont f a) = (isCont f a)";
+by (rtac isNSCont_LIM_iff 1);
+qed "isNSCont_isCont_iff";
+
+(*----------------------------------------
+  Standard continuity ==> NS continuity 
+ ----------------------------------------*)
+Goal "isCont f a ==> isNSCont f a";
+by (etac (isNSCont_isCont_iff RS iffD2) 1);
+qed "isCont_isNSCont";
+
+(*----------------------------------------
+  NS continuity ==> Standard continuity 
+ ----------------------------------------*)
+Goal "isNSCont f a ==> isCont f a";
+by (etac (isNSCont_isCont_iff RS iffD1) 1);
+qed "isNSCont_isCont";
+
+(*--------------------------------------------------------------------------
+                 Alternative definition of continuity
+ --------------------------------------------------------------------------*)
+(* Prove equivalence between NS limits - *)
+(* seems easier than using standard def  *)
+Goalw [NSLIM_def] "(f -- a --NS> L) = ((%h. f(a + h)) -- #0 --NS> L)";
+by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_zero]));
+by (dres_inst_tac [("x","hypreal_of_real a + x")] spec 1);
+by (dres_inst_tac [("x","-hypreal_of_real a + x")] spec 2);
+by (Step_tac 1);
+by (Asm_full_simp_tac 1);
+by (rtac ((mem_infmal_iff RS iffD2) RS 
+    (Infinitesimal_add_inf_close_self RS inf_close_sym)) 1);
+by (rtac (inf_close_minus_iff2 RS iffD1) 4);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 3);
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 2);
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 4);
+by (auto_tac (claset(),
+       simpset() addsimps [starfun, hypreal_of_real_def, hypreal_minus,
+              hypreal_add, real_add_assoc, inf_close_refl, hypreal_zero_def]));
+qed "NSLIM_h_iff";
+
+Goal "(f -- a --NS> f a) = ((%h. f(a + h)) -- #0 --NS> f a)";
+by (rtac NSLIM_h_iff 1);
+qed "NSLIM_isCont_iff";
+
+Goal "(f -- a --> f a) = ((%h. f(a + h)) -- #0 --> f(a))";
+by (simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_isCont_iff]) 1);
+qed "LIM_isCont_iff";
+
+Goalw [isCont_def] "(isCont f x) = ((%h. f(x + h)) -- #0 --> f(x))";
+by (simp_tac (simpset() addsimps [LIM_isCont_iff]) 1);
+qed "isCont_iff";
+
+(*--------------------------------------------------------------------------
+   Immediate application of nonstandard criterion for continuity can offer 
+   very simple proofs of some standard property of continuous functions
+ --------------------------------------------------------------------------*)
+(*------------------------
+     sum continuous
+ ------------------------*)
+Goal "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) + g(x)) a";
+by (auto_tac (claset() addIs [inf_close_add],
+              simpset() addsimps [isNSCont_isCont_iff RS sym, isNSCont_def]));
+qed "isCont_add";
+
+(*------------------------
+     mult continuous
+ ------------------------*)
+Goal "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) * g(x)) a";
+by (auto_tac (claset() addSIs [starfun_mult_HFinite_inf_close],
+              simpset() delsimps [starfun_mult RS sym]
+			addsimps [isNSCont_isCont_iff RS sym, isNSCont_def]));
+qed "isCont_mult";
+
+(*-------------------------------------------
+     composition of continuous functions
+     Note very short straightforard proof!
+ ------------------------------------------*)
+Goal "[| isCont f a; isCont g (f a) |] \
+\     ==> isCont (g o f) a";
+by (auto_tac (claset(),simpset() addsimps [isNSCont_isCont_iff RS sym,
+              isNSCont_def,starfun_o RS sym]));
+qed "isCont_o";
+
+Goal "[| isCont f a; isCont g (f a) |] \
+\     ==> isCont (%x. g (f x)) a";
+by (auto_tac (claset() addDs [isCont_o],simpset() addsimps [o_def]));
+qed "isCont_o2";
+
+Goalw [isNSCont_def] "isNSCont f a ==> isNSCont (%x. - f x) a";
+by Auto_tac; 
+qed "isNSCont_minus";
+
+Goal "isCont f a ==> isCont (%x. - f x) a";
+by (auto_tac (claset(),simpset() addsimps [isNSCont_isCont_iff RS sym,
+              isNSCont_minus]));
+qed "isCont_minus";
+
+Goalw [isCont_def]  
+      "[| isCont f x; f x ~= #0 |] ==> isCont (%x. inverse (f x)) x";
+by (blast_tac (claset() addIs [LIM_inverse]) 1);
+qed "isCont_inverse";
+
+Goal "[| isNSCont f x; f x ~= #0 |] ==> isNSCont (%x. inverse (f x)) x";
+by (auto_tac (claset() addIs [isCont_inverse],simpset() addsimps 
+    [isNSCont_isCont_iff]));
+qed "isNSCont_inverse";
+
+Goalw [real_diff_def] 
+      "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) - g(x)) a";
+by (auto_tac (claset() addIs [isCont_add,isCont_minus],simpset()));
+qed "isCont_diff";
+
+Goalw [isCont_def]  "isCont (%x. k) a";
+by (Simp_tac 1);
+qed "isCont_const";
+Addsimps [isCont_const];
+
+Goalw [isNSCont_def]  "isNSCont (%x. k) a";
+by (Simp_tac 1);
+qed "isNSCont_const";
+Addsimps [isNSCont_const];
+
+Goalw [isNSCont_def]  "isNSCont abs a";
+by (auto_tac (claset() addIs [inf_close_hrabs],
+              simpset() addsimps [hypreal_of_real_hrabs RS sym,
+                                  starfun_rabs_hrabs]));
+qed "isNSCont_rabs";
+Addsimps [isNSCont_rabs];
+
+Goal "isCont abs a";
+by (auto_tac (claset(), simpset() addsimps [isNSCont_isCont_iff RS sym]));
+qed "isCont_rabs";
+Addsimps [isCont_rabs];
+
+(****************************************************************
+(%* Leave as commented until I add topology theory or remove? *%)
+(%*------------------------------------------------------------
+  Elementary topology proof for a characterisation of 
+  continuity now: a function f is continuous if and only 
+  if the inverse image, {x. f(x) : A}, of any open set A 
+  is always an open set
+ ------------------------------------------------------------*%)
+Goal "[| isNSopen A; ALL x. isNSCont f x |] \
+\              ==> isNSopen {x. f x : A}";
+by (auto_tac (claset(),simpset() addsimps [isNSopen_iff1]));
+by (dtac (mem_monad_inf_close RS inf_close_sym) 1);
+by (dres_inst_tac [("x","a")] spec 1);
+by (dtac isNSContD 1 THEN assume_tac 1);
+by (dtac bspec 1 THEN assume_tac 1);
+by (dres_inst_tac [("x","( *f* f) x")] inf_close_mem_monad2 1);
+by (blast_tac (claset() addIs [starfun_mem_starset]) 1);
+qed "isNSCont_isNSopen";
+
+Goalw [isNSCont_def]
+          "ALL A. isNSopen A --> isNSopen {x. f x : A} \
+\              ==> isNSCont f x";
+by (auto_tac (claset() addSIs [(mem_infmal_iff RS iffD1) RS 
+     (inf_close_minus_iff RS iffD2)],simpset() addsimps 
+      [Infinitesimal_def,SReal_iff]));
+by (dres_inst_tac [("x","{z. abs(z + -f(x)) < ya}")] spec 1);
+by (etac (isNSopen_open_interval RSN (2,impE)) 1);
+by (auto_tac (claset(),simpset() addsimps [isNSopen_def,isNSnbhd_def]));
+by (dres_inst_tac [("x","x")] spec 1);
+by (auto_tac (claset() addDs [inf_close_sym RS inf_close_mem_monad],
+    simpset() addsimps [hypreal_of_real_zero RS sym,STAR_starfun_rabs_add_minus]));
+qed "isNSopen_isNSCont";
+
+Goal "(ALL x. isNSCont f x) = \
+\     (ALL A. isNSopen A --> isNSopen {x. f(x) : A})";
+by (blast_tac (claset() addIs [isNSCont_isNSopen,
+    isNSopen_isNSCont]) 1);
+qed "isNSCont_isNSopen_iff";
+
+(%*------- Standard version of same theorem --------*%)
+Goal "(ALL x. isCont f x) = \
+\         (ALL A. isopen A --> isopen {x. f(x) : A})";
+by (auto_tac (claset() addSIs [isNSCont_isNSopen_iff],
+              simpset() addsimps [isNSopen_isopen_iff RS sym,
+              isNSCont_isCont_iff RS sym]));
+qed "isCont_isopen_iff";
+*******************************************************************)
+
+(*-----------------------------------------------------------------
+                        Uniform continuity
+ ------------------------------------------------------------------*)
+Goalw [isNSUCont_def] 
+      "[| isNSUCont f; x @= y|] ==> (*f* f) x @= (*f* f) y";
+by (Blast_tac 1);
+qed "isNSUContD";
+
+Goalw [isUCont_def,isCont_def,LIM_def]
+     "isUCont f ==> EX x. isCont f x";
+by (Force_tac 1);
+qed "isUCont_isCont";
+
+Goalw [isNSUCont_def,isUCont_def,inf_close_def] 
+     "isUCont f ==> isNSUCont f";
+by (asm_full_simp_tac (simpset() addsimps 
+    [Infinitesimal_FreeUltrafilterNat_iff]) 1);
+by (Step_tac 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [starfun,
+    hypreal_minus, hypreal_add]));
+by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2 THEN Step_tac 1);
+by (dres_inst_tac [("x","u")] spec 1 THEN Clarify_tac 1);
+by (dres_inst_tac [("x","s")] spec 1 THEN Clarify_tac 1);
+by (subgoal_tac "ALL n::nat. abs ((xa n) + - (xb n)) < s --> abs (f (xa n) + - f (xb n)) < u" 1);
+by (Blast_tac 2);
+by (thin_tac "ALL x y. abs (x + - y) < s --> abs (f x + - f y) < u" 1);
+by (dtac FreeUltrafilterNat_all 1);
+by (Ultra_tac 1);
+qed "isUCont_isNSUCont";
+
+Goal "ALL s. #0 < s --> (EX z y. abs (z + - y) < s & r <= abs (f z + -f y)) \
+\     ==> ALL n::nat. EX z y.  \
+\              abs(z + -y) < inverse(real_of_posnat n) & \
+\              r <= abs(f z + -f y)";
+by (Step_tac 1);
+by (cut_inst_tac [("n1","n")] (real_of_posnat_gt_zero RS real_inverse_gt_zero) 1);
+by Auto_tac;
+val lemma_LIMu = result();
+
+Goal "ALL s. #0 < s --> (EX z y. abs (z + - y) < s  & r <= abs (f z + -f y)) \
+\     ==> EX X Y. ALL n::nat. \
+\              abs(X n + -(Y n)) < inverse(real_of_posnat n) & \
+\              r <= abs(f (X n) + -f (Y n))";
+by (dtac lemma_LIMu 1);
+by (dtac choice 1);
+by (Step_tac 1);
+by (dtac choice 1);
+by (Blast_tac 1);
+val lemma_skolemize_LIM2u = result();
+
+Goal "ALL n. abs (X n + -Y n) < inverse (real_of_posnat  n) & \
+\         r <= abs (f (X n) + - f(Y n)) ==> \
+\         ALL n. abs (X n + - Y n) < inverse (real_of_posnat  n)";
+by (Auto_tac );
+val lemma_simpu = result();
+
+Goalw [isNSUCont_def,isUCont_def,inf_close_def] 
+     "isNSUCont f ==> isUCont f";
+by (asm_full_simp_tac (simpset() addsimps 
+                       [Infinitesimal_FreeUltrafilterNat_iff]) 1);
+by (EVERY1[Step_tac, rtac ccontr, Asm_full_simp_tac]);
+by (fold_tac [real_le_def]);
+by (dtac lemma_skolemize_LIM2u 1);
+by (Step_tac 1);
+by (dres_inst_tac [("x","Abs_hypreal(hyprel^^{X})")] spec 1);
+by (dres_inst_tac [("x","Abs_hypreal(hyprel^^{Y})")] spec 1);
+by (asm_full_simp_tac
+    (simpset() addsimps [starfun, hypreal_minus,hypreal_add]) 1);
+by Auto_tac;
+by (dtac (lemma_simpu RS real_seq_to_hypreal_Infinitesimal2) 1);
+by (asm_full_simp_tac (simpset() addsimps 
+     [Infinitesimal_FreeUltrafilterNat_iff, hypreal_minus,hypreal_add]) 1);
+by (Blast_tac 1);
+by (rotate_tac 2 1);
+by (dres_inst_tac [("x","r")] spec 1);
+by (Clarify_tac 1);
+by (dtac FreeUltrafilterNat_all 1);
+by (Ultra_tac 1);
+qed "isNSUCont_isUCont";
+
+(*------------------------------------------------------------------
+                         Derivatives
+ ------------------------------------------------------------------*)
+Goalw [deriv_def] 
+      "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- #0 --> D)";
+by (Blast_tac 1);        
+qed "DERIV_iff";
+
+Goalw [deriv_def] 
+      "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- #0 --NS> D)";
+by (simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);
+qed "DERIV_NS_iff";
+
+Goalw [deriv_def] 
+      "DERIV f x :> D \
+\      ==> (%h. (f(x + h) + - f(x))/h) -- #0 --> D";
+by (Blast_tac 1);        
+qed "DERIVD";
+
+Goalw [deriv_def] "DERIV f x :> D ==> \
+\          (%h. (f(x + h) + - f(x))/h) -- #0 --NS> D";
+by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);
+qed "NS_DERIVD";
+
+(* Uniqueness *)
+Goalw [deriv_def] 
+      "[| DERIV f x :> D; DERIV f x :> E |] ==> D = E";
+by (blast_tac (claset() addIs [LIM_unique]) 1);
+qed "DERIV_unique";
+
+Goalw [nsderiv_def] 
+     "[| NSDERIV f x :> D; NSDERIV f x :> E |] ==> D = E";
+by (cut_facts_tac [Infinitesimal_epsilon, hypreal_epsilon_not_zero] 1);
+by (auto_tac (claset() addSDs [inst "x" "ehr" bspec] 
+                       addSIs [inj_hypreal_of_real RS injD] 
+                       addDs [inf_close_trans3],
+              simpset()));
+qed "NSDeriv_unique";
+
+(*------------------------------------------------------------------------
+                          Differentiable
+ ------------------------------------------------------------------------*)
+
+Goalw [differentiable_def] 
+      "f differentiable x ==> EX D. DERIV f x :> D";
+by (assume_tac 1);
+qed "differentiableD";
+
+Goalw [differentiable_def] 
+      "DERIV f x :> D ==> f differentiable x";
+by (Blast_tac 1);
+qed "differentiableI";
+
+Goalw [NSdifferentiable_def] 
+      "f NSdifferentiable x ==> EX D. NSDERIV f x :> D";
+by (assume_tac 1);
+qed "NSdifferentiableD";
+
+Goalw [NSdifferentiable_def] 
+      "NSDERIV f x :> D ==> f NSdifferentiable x";
+by (Blast_tac 1);
+qed "NSdifferentiableI";
+
+(*--------------------------------------------------------
+      Alternative definition for differentiability
+ -------------------------------------------------------*)
+
+Goalw [LIM_def] 
+ "((%h. (f(a + h) + - f(a))/h) -- #0 --> D) = \
+\ ((%x. (f(x) + -f(a)) / (x + -a)) -- a --> D)";
+by (Step_tac 1);
+by (ALLGOALS(dtac spec));
+by (Step_tac 1);
+by (Blast_tac 1 THEN Blast_tac 2);
+by (ALLGOALS(res_inst_tac [("x","s")] exI));
+by (Step_tac 1);
+by (dres_inst_tac [("x","x + -a")] spec 1);
+by (dres_inst_tac [("x","x + a")] spec 2);
+by (auto_tac (claset(), simpset() addsimps real_add_ac));
+qed "DERIV_LIM_iff";
+
+Goalw [deriv_def] "(DERIV f x :> D) = \
+\         ((%z. (f(z) + -f(x)) / (z + -x)) -- x --> D)";
+by (simp_tac (simpset() addsimps [DERIV_LIM_iff]) 1);
+qed "DERIV_iff2";
+
+(*--------------------------------------------------------
+  Equivalence of NS and standard defs of differentiation
+ -------------------------------------------------------*)
+(*-------------------------------------------
+   First NSDERIV in terms of NSLIM 
+ -------------------------------------------*)
+
+(*--- first equivalence ---*)
+Goalw [nsderiv_def,NSLIM_def] 
+      "(NSDERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- #0 --NS> D)";
+by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_zero]));
+by (dres_inst_tac [("x","xa")] bspec 1);
+by (rtac ccontr 3);
+by (dres_inst_tac [("x","h")] spec 3);
+by (auto_tac (claset(),
+              simpset() addsimps [mem_infmal_iff, starfun_lambda_cancel]));
+qed "NSDERIV_NSLIM_iff";
+
+(*--- second equivalence ---*)
+Goal "(NSDERIV f x :> D) = \
+\         ((%z. (f(z) + -f(x)) / (z + -x)) -- x --NS> D)";
+by (full_simp_tac (simpset() addsimps 
+     [NSDERIV_NSLIM_iff, DERIV_LIM_iff, LIM_NSLIM_iff RS sym]) 1);
+qed "NSDERIV_NSLIM_iff2";
+
+(* while we're at it! *)
+Goalw [real_diff_def]
+     "(NSDERIV f x :> D) = \
+\     (ALL xa. \
+\       xa ~= hypreal_of_real x & xa @= hypreal_of_real x --> \
+\       (*f* (%z. (f z - f x) / (z - x))) xa @= hypreal_of_real D)";
+by (auto_tac (claset(), simpset() addsimps [NSDERIV_NSLIM_iff2, NSLIM_def]));
+qed "NSDERIV_iff2";
+
+
+Goal "(NSDERIV f x :> D) ==> \
+\    (ALL u. \
+\       u @= hypreal_of_real x --> \
+\       (*f* (%z. f z - f x)) u @= hypreal_of_real D * (u - hypreal_of_real x))";
+by (auto_tac (claset(), simpset() addsimps [NSDERIV_iff2]));
+by (case_tac "u = hypreal_of_real x" 1);
+by (auto_tac (claset(), 
+              simpset() addsimps [hypreal_diff_def, hypreal_of_real_zero]));
+by (dres_inst_tac [("x","u")] spec 1);
+by Auto_tac;
+by (dres_inst_tac [("c","u - hypreal_of_real x"),("b","hypreal_of_real D")]
+     inf_close_mult1 1);
+by (ALLGOALS(dtac (hypreal_not_eq_minus_iff RS iffD1)));
+by (subgoal_tac "(*f* (%z. z - x)) u ~= (0::hypreal)" 2);
+by (rotate_tac ~1 2);
+by (auto_tac (claset(),
+    simpset() addsimps [real_diff_def, hypreal_diff_def, 
+		(inf_close_minus_iff RS iffD1) RS (mem_infmal_iff RS iffD2),  
+			Infinitesimal_subset_HFinite RS subsetD]));
+qed "NSDERIVD5";
+
+Goal "(NSDERIV f x :> D) ==> \
+\     (ALL h: Infinitesimal. \
+\              ((*f* f)(hypreal_of_real x + h) - \
+\                hypreal_of_real (f x))@= (hypreal_of_real D) * h)";
+by (auto_tac (claset(),simpset() addsimps [nsderiv_def]));
+by (case_tac "h = (0::hypreal)" 1);
+by (auto_tac (claset(),simpset() addsimps [hypreal_diff_def]));
+by (dres_inst_tac [("x","h")] bspec 1);
+by (dres_inst_tac [("c","h")] inf_close_mult1 2);
+by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],
+              simpset() addsimps [hypreal_diff_def]));
+qed "NSDERIVD4";
+
+Goal "(NSDERIV f x :> D) ==> \
+\     (ALL h: Infinitesimal - {0}. \
+\              ((*f* f)(hypreal_of_real x + h) - \
+\                hypreal_of_real (f x))@= (hypreal_of_real D) * h)";
+by (auto_tac (claset(),simpset() addsimps [nsderiv_def]));
+by (rtac ccontr 1 THEN dres_inst_tac [("x","h")] bspec 1);
+by (dres_inst_tac [("c","h")] inf_close_mult1 2);
+by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],
+              simpset() addsimps [hypreal_mult_assoc, hypreal_diff_def]));
+qed "NSDERIVD3";
+
+(*--------------------------------------------------------------
+          Now equivalence between NSDERIV and DERIV
+ -------------------------------------------------------------*)
+Goalw [deriv_def] "(NSDERIV f x :> D) = (DERIV f x :> D)";
+by (simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,LIM_NSLIM_iff]) 1);
+qed "NSDERIV_DERIV_iff";
+
+(*---------------------------------------------------
+         Differentiability implies continuity 
+         nice and simple "algebraic" proof
+ --------------------------------------------------*)
+Goalw [nsderiv_def]
+      "NSDERIV f x :> D ==> isNSCont f x";
+by (auto_tac (claset(),simpset() addsimps 
+        [isNSCont_NSLIM_iff,NSLIM_def]));
+by (dtac (inf_close_minus_iff RS iffD1) 1);
+by (dtac (hypreal_not_eq_minus_iff RS iffD1) 1);
+by (dres_inst_tac [("x","-hypreal_of_real x + xa")] bspec 1);
+by (asm_full_simp_tac (simpset() addsimps 
+    [hypreal_add_assoc RS sym]) 2);
+by (auto_tac (claset(),simpset() addsimps 
+    [mem_infmal_iff RS sym,hypreal_add_commute]));
+by (dres_inst_tac [("c","xa + -hypreal_of_real x")] inf_close_mult1 1);
+by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite
+    RS subsetD],simpset() addsimps [hypreal_mult_assoc]));
+by (dres_inst_tac [("x3","D")] (HFinite_hypreal_of_real RSN
+    (2,Infinitesimal_HFinite_mult) RS (mem_infmal_iff RS iffD1)) 1);
+by (blast_tac (claset() addIs [inf_close_trans,
+    hypreal_mult_commute RS subst,
+    (inf_close_minus_iff RS iffD2)]) 1);
+qed "NSDERIV_isNSCont";
+
+(* Now Sandard proof *)
+Goal "DERIV f x :> D ==> isCont f x";
+by (asm_full_simp_tac (simpset() addsimps 
+    [NSDERIV_DERIV_iff RS sym, isNSCont_isCont_iff RS sym,
+     NSDERIV_isNSCont]) 1);
+qed "DERIV_isCont";
+
+(*----------------------------------------------------------------------------
+      Differentiation rules for combinations of functions
+      follow from clear, straightforard, algebraic 
+      manipulations
+ ----------------------------------------------------------------------------*)
+(*-------------------------
+    Constant function
+ ------------------------*)
+
+(* use simple constant nslimit theorem *)
+Goal "(NSDERIV (%x. k) x :> #0)";
+by (simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff]) 1);
+qed "NSDERIV_const";
+Addsimps [NSDERIV_const];
+
+Goal "(DERIV (%x. k) x :> #0)";
+by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1);
+qed "DERIV_const";
+Addsimps [DERIV_const];
+
+(*-----------------------------------------------------
+    Sum of functions- proved easily
+ ----------------------------------------------------*)
+
+
+Goal "[| NSDERIV f x :> Da;  NSDERIV g x :> Db |] \
+\     ==> NSDERIV (%x. f x + g x) x :> Da + Db";
+by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,
+           NSLIM_def]) 1 THEN REPEAT(Step_tac 1));
+by (auto_tac (claset(),
+       simpset() addsimps [hypreal_add_divide_distrib]));
+by (dres_inst_tac [("b","hypreal_of_real Da"),
+                   ("d","hypreal_of_real Db")] inf_close_add 1);
+by (auto_tac (claset(), simpset() addsimps hypreal_add_ac));
+qed "NSDERIV_add";
+
+(* Standard theorem *)
+Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \
+\     ==> DERIV (%x. f x + g x) x :> Da + Db";
+by (asm_full_simp_tac (simpset() addsimps [NSDERIV_add,
+                                     NSDERIV_DERIV_iff RS sym]) 1);
+qed "DERIV_add";
+
+(*-----------------------------------------------------
+  Product of functions - Proof is trivial but tedious
+  and long due to rearrangement of terms  
+ ----------------------------------------------------*)
+
+Goal "((a::hypreal)*b) + -(c*d) = (b*(a + -c)) + (c*(b + -d))";
+by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib2]) 1);
+val lemma_nsderiv1 = result();
+
+Goal "[| (x + y) / z = hypreal_of_real D + yb; z ~= 0; \
+\        z : Infinitesimal; yb : Infinitesimal |] \
+\     ==> x + y @= #0";
+by (forw_inst_tac [("c1","z")] (hypreal_mult_right_cancel RS iffD2) 1 
+    THEN assume_tac 1);
+by (thin_tac "(x + y) / z = hypreal_of_real D + yb" 1);
+by (auto_tac (claset() addSIs [Infinitesimal_HFinite_mult2, HFinite_add],
+              simpset() addsimps [hypreal_mult_assoc, mem_infmal_iff RS sym]));
+by (etac (Infinitesimal_subset_HFinite RS subsetD) 1);
+val lemma_nsderiv2 = result();
+
+
+Goal "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \
+\     ==> NSDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))";
+by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff, NSLIM_def]) 1 
+    THEN REPEAT(Step_tac 1));
+by (auto_tac (claset(),
+       simpset() addsimps [starfun_lambda_cancel, hypreal_of_real_zero,
+              lemma_nsderiv1]));
+by (simp_tac (simpset() addsimps [hypreal_add_divide_distrib]) 1); 
+by (REPEAT(dtac (bex_Infinitesimal_iff2 RS iffD2) 1));
+by (auto_tac (claset(),
+        simpset() delsimps [hypreal_times_divide1_eq]
+		  addsimps [hypreal_times_divide1_eq RS sym]));
+by (dres_inst_tac [("D","Db")] lemma_nsderiv2 1);
+by (dtac (inf_close_minus_iff RS iffD2 RS (bex_Infinitesimal_iff2 RS iffD2)) 4);
+by (auto_tac (claset() addSIs [inf_close_add_mono1],
+      simpset() addsimps [hypreal_add_mult_distrib, hypreal_add_mult_distrib2, 
+			  hypreal_mult_commute, hypreal_add_assoc]));
+by (res_inst_tac [("w1","hypreal_of_real Db * hypreal_of_real (f x)")]
+    (hypreal_add_commute RS subst) 1);
+by (auto_tac (claset() addSIs [Infinitesimal_add_inf_close_self2 RS inf_close_sym,
+			       Infinitesimal_add, Infinitesimal_mult,
+			       Infinitesimal_hypreal_of_real_mult,
+			       Infinitesimal_hypreal_of_real_mult2],
+	      simpset() addsimps [hypreal_add_assoc RS sym]));
+qed "NSDERIV_mult";
+
+Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \
+\     ==> DERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))";
+by (asm_full_simp_tac (simpset() addsimps [NSDERIV_mult,
+                                           NSDERIV_DERIV_iff RS sym]) 1);
+qed "DERIV_mult";
+
+(*----------------------------
+   Multiplying by a constant
+ ---------------------------*)
+Goal "NSDERIV f x :> D \
+\     ==> NSDERIV (%x. c * f x) x :> c*D";
+by (asm_full_simp_tac 
+    (simpset() addsimps [real_times_divide1_eq RS sym, NSDERIV_NSLIM_iff,
+                         real_minus_mult_eq2, real_add_mult_distrib2 RS sym] 
+             delsimps [real_times_divide1_eq, real_minus_mult_eq2 RS sym]) 1);
+by (etac (NSLIM_const RS NSLIM_mult) 1);
+qed "NSDERIV_cmult";
+
+(* let's do the standard proof though theorem *)
+(* LIM_mult2 follows from a NS proof          *)
+
+Goalw [deriv_def] 
+      "DERIV f x :> D \
+\      ==> DERIV (%x. c * f x) x :> c*D";
+by (asm_full_simp_tac 
+    (simpset() addsimps [real_times_divide1_eq RS sym, NSDERIV_NSLIM_iff,
+                         real_minus_mult_eq2, real_add_mult_distrib2 RS sym] 
+             delsimps [real_times_divide1_eq, real_minus_mult_eq2 RS sym]) 1);
+by (etac (LIM_const RS LIM_mult2) 1);
+qed "DERIV_cmult";
+
+(*--------------------------------
+   Negation of function
+ -------------------------------*)
+Goal "NSDERIV f x :> D ==> NSDERIV (%x. -(f x)) x :> -D";
+by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff]) 1);
+by (res_inst_tac [("t","f x")] (real_minus_minus RS subst) 1);
+by (asm_simp_tac (simpset() addsimps [real_minus_add_distrib RS sym,
+                                      real_minus_mult_eq1 RS sym] 
+                   delsimps [real_minus_add_distrib, real_minus_minus]) 1);
+by (etac NSLIM_minus 1);
+qed "NSDERIV_minus";
+
+Goal "DERIV f x :> D \
+\     ==> DERIV (%x. -(f x)) x :> -D";
+by (asm_full_simp_tac (simpset() addsimps 
+    [NSDERIV_minus,NSDERIV_DERIV_iff RS sym]) 1);
+qed "DERIV_minus";
+
+(*-------------------------------
+   Subtraction
+ ------------------------------*)
+Goal "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \
+\     ==> NSDERIV (%x. f x + -g x) x :> Da + -Db";
+by (blast_tac (claset() addDs [NSDERIV_add,NSDERIV_minus]) 1);
+qed "NSDERIV_add_minus";
+
+Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \
+\     ==> DERIV (%x. f x + -g x) x :> Da + -Db";
+by (blast_tac (claset() addDs [DERIV_add,DERIV_minus]) 1);
+qed "DERIV_add_minus";
+
+Goalw [real_diff_def]
+     "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \
+\     ==> NSDERIV (%x. f x - g x) x :> Da - Db";
+by (blast_tac (claset() addIs [NSDERIV_add_minus]) 1);
+qed "NSDERIV_diff";
+
+Goalw [real_diff_def]
+     "[| DERIV f x :> Da; DERIV g x :> Db |] \
+\      ==> DERIV (%x. f x - g x) x :> Da - Db";
+by (blast_tac (claset() addIs [DERIV_add_minus]) 1);
+qed "DERIV_diff";
+
+(*---------------------------------------------------------------
+                     (NS) Increment
+ ---------------------------------------------------------------*)
+Goalw [increment_def] 
+      "f NSdifferentiable x ==> \
+\     increment f x h = (*f* f) (hypreal_of_real(x) + h) + \
+\     -hypreal_of_real (f x)";
+by (Blast_tac 1);
+qed "incrementI";
+
+Goal "NSDERIV f x :> D ==> \
+\    increment f x h = (*f* f) (hypreal_of_real(x) + h) + \
+\    -hypreal_of_real (f x)";
+by (etac (NSdifferentiableI RS incrementI) 1);
+qed "incrementI2";
+
+(* The Increment theorem -- Keisler p. 65 *)
+Goal "[| NSDERIV f x :> D; h: Infinitesimal; h ~= #0 |] \
+\     ==> EX e: Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h";
+by (forw_inst_tac [("h","h")] incrementI2 1 THEN rewtac nsderiv_def);
+by (dtac bspec 1 THEN Auto_tac);
+by (dtac (bex_Infinitesimal_iff2 RS iffD2) 1 THEN Step_tac 1);
+by (forw_inst_tac [("b1","hypreal_of_real(D) + y")] 
+    (rename_numerals (hypreal_mult_right_cancel RS iffD2)) 1);
+by (thin_tac "((*f* f) (hypreal_of_real(x) + h) + \
+\   - hypreal_of_real (f x)) / h = hypreal_of_real(D) + y" 2);
+by (assume_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_times_divide1_eq RS sym]
+             delsimps [hypreal_times_divide1_eq]) 1);
+by (auto_tac (claset(),
+              simpset() addsimps [hypreal_add_mult_distrib]));
+qed "increment_thm";
+
+Goal "[| NSDERIV f x :> D; h @= #0; h ~= #0 |] \
+\     ==> EX e: Infinitesimal. increment f x h = \
+\             hypreal_of_real(D)*h + e*h";
+by (blast_tac (claset() addSDs [mem_infmal_iff RS iffD2] 
+                        addSIs [increment_thm]) 1);
+qed "increment_thm2";
+
+Goal "[| NSDERIV f x :> D; h @= #0; h ~= #0 |] \
+\     ==> increment f x h @= #0";
+by (dtac increment_thm2 1 THEN auto_tac (claset() addSIs 
+    [Infinitesimal_HFinite_mult2,HFinite_add],simpset() addsimps 
+    [hypreal_add_mult_distrib RS sym,mem_infmal_iff RS sym]));
+by (etac (Infinitesimal_subset_HFinite RS subsetD) 1);
+qed "increment_inf_close_zero";
+
+(*---------------------------------------------------------------
+   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 *)
+Goalw [nsderiv_def] 
+      "[| NSDERIV g x :> D; \
+\              (*f* g) (hypreal_of_real(x) + xa) = hypreal_of_real(g x);\
+\              xa : Infinitesimal;\
+\              xa ~= #0 \
+\           |] ==> D = #0";
+by (dtac bspec 1);
+by Auto_tac;
+qed "NSDERIV_zero";
+
+(* can be proved differently using NSLIM_isCont_iff *)
+Goalw [nsderiv_def] 
+     "[| NSDERIV f x :> D;  h: Infinitesimal;  h ~= #0 |]  \
+\     ==> (*f* f) (hypreal_of_real(x) + h) + -hypreal_of_real(f x) @= #0";    
+by (asm_full_simp_tac (simpset() addsimps 
+    [mem_infmal_iff RS sym]) 1);
+by (rtac Infinitesimal_ratio 1);
+by (rtac inf_close_hypreal_of_real_HFinite 3);
+by Auto_tac;
+qed "NSDERIV_inf_close";
+
+(*--------------------------------------------------------------- 
+   from one version of differentiability 
+ 
+                f(x) - f(a)
+              --------------- @= Db
+                  x - a
+ ---------------------------------------------------------------*)
+Goal "[| NSDERIV f (g x) :> Da; \
+\        (*f* g) (hypreal_of_real(x) + xa) ~= hypreal_of_real (g x); \
+\        (*f* g) (hypreal_of_real(x) + xa) @= hypreal_of_real (g x) \
+\     |] ==> ((*f* f) ((*f* g) (hypreal_of_real(x) + xa)) \
+\                  + - hypreal_of_real (f (g x))) \
+\             / ((*f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real (g x)) \
+\            @= hypreal_of_real(Da)";
+by (auto_tac (claset(),
+       simpset() addsimps [NSDERIV_NSLIM_iff2, NSLIM_def]));
+qed "NSDERIVD1";
+
+(*-------------------------------------------------------------- 
+   from other version of differentiability 
+
+                f(x + h) - f(x)
+               ----------------- @= Db
+                       h
+ --------------------------------------------------------------*)
+Goal "[| NSDERIV g x :> Db; xa: Infinitesimal; xa ~= #0 |] \
+\     ==> ((*f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real(g x)) / xa \
+\         @= hypreal_of_real(Db)";
+by (auto_tac (claset(),
+    simpset() addsimps [NSDERIV_NSLIM_iff, NSLIM_def, 
+		hypreal_of_real_zero, mem_infmal_iff, starfun_lambda_cancel]));
+qed "NSDERIVD2";
+
+Goal "(z::hypreal) ~= 0 ==> x*y = (x*inverse(z))*(z*y)";
+by Auto_tac;  
+qed "lemma_chain";
+
+(*------------------------------------------------------
+  This proof uses both definitions of differentiability.
+ ------------------------------------------------------*)
+Goal "[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |] \
+\     ==> NSDERIV (f o g) x :> Da * Db";
+by (asm_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,
+    NSLIM_def,hypreal_of_real_zero,mem_infmal_iff RS sym]) 1 THEN Step_tac 1);
+by (forw_inst_tac [("f","g")] NSDERIV_inf_close 1);
+by (auto_tac (claset(),
+              simpset() addsimps [starfun_lambda_cancel2, starfun_o RS sym]));
+by (case_tac "(*f* g) (hypreal_of_real(x) + xa) = hypreal_of_real (g x)" 1);
+by (dres_inst_tac [("g","g")] NSDERIV_zero 1);
+by (auto_tac (claset(),
+    simpset() addsimps [hypreal_divide_def, hypreal_of_real_zero]));
+by (res_inst_tac [("z1","(*f* g) (hypreal_of_real(x) + xa) + -hypreal_of_real (g x)"),
+    ("y1","inverse xa")] (lemma_chain RS ssubst) 1);
+by (etac (hypreal_not_eq_minus_iff RS iffD1) 1);
+by (rtac inf_close_mult_hypreal_of_real 1);
+by (fold_tac [hypreal_divide_def]);
+by (blast_tac (claset() addIs [NSDERIVD1,
+    inf_close_minus_iff RS iffD2]) 1);
+by (blast_tac (claset() addIs [NSDERIVD2]) 1);
+qed "NSDERIV_chain";
+
+(* standard version *)
+Goal "[| DERIV f (g x) :> Da; \
+\                 DERIV g x :> Db \
+\              |] ==> DERIV (f o g) x :> Da * Db";
+by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym,
+    NSDERIV_chain]) 1);
+qed "DERIV_chain";
+
+Goal "[| DERIV f (g x) :> Da; DERIV g x :> Db |] \
+\     ==> DERIV (%x. f (g x)) x :> Da * Db";
+by (auto_tac (claset() addDs [DERIV_chain], simpset() addsimps [o_def]));
+qed "DERIV_chain2";
+
+(*------------------------------------------------------------------
+           Differentiation of natural number powers
+ ------------------------------------------------------------------*)
+Goal "NSDERIV (%x. x) x :> #1";
+by (auto_tac (claset(),
+     simpset() addsimps [NSDERIV_NSLIM_iff,
+          NSLIM_def ,starfun_Id, hypreal_of_real_zero,
+           hypreal_of_real_one]));
+qed "NSDERIV_Id";
+Addsimps [NSDERIV_Id];
+
+(*derivative of the identity function*)
+Goal "DERIV (%x. x) x :> #1";
+by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1);
+qed "DERIV_Id";
+Addsimps [DERIV_Id];
+
+bind_thm ("isCont_Id", DERIV_Id RS DERIV_isCont);
+
+(*derivative of linear multiplication*)
+Goal "DERIV (op * c) x :> c";
+by (cut_inst_tac [("c","c"),("x","x")] (DERIV_Id RS DERIV_cmult) 1);
+by (Asm_full_simp_tac 1);
+qed "DERIV_cmult_Id";
+Addsimps [DERIV_cmult_Id];
+
+Goal "NSDERIV (op * c) x :> c";
+by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff]) 1);
+qed "NSDERIV_cmult_Id";
+Addsimps [NSDERIV_cmult_Id];
+
+Goal "DERIV (%x. x ^ n) x :> real_of_nat n * (x ^ (n - 1))";
+by (induct_tac "n" 1);
+by (dtac (DERIV_Id RS DERIV_mult) 2);
+by (auto_tac (claset(),simpset() addsimps 
+    [real_add_mult_distrib]));
+by (case_tac "0 < n" 1);
+by (dres_inst_tac [("x","x")] realpow_minus_mult 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [real_mult_assoc,real_add_commute]));
+qed "DERIV_pow";
+
+(* NS version *)
+Goal "NSDERIV (%x. x ^ n) x :> real_of_nat n * (x ^ (n - 1))";
+by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,DERIV_pow]) 1);
+qed "NSDERIV_pow";
+
+(*---------------------------------------------------------------
+                    Power of -1 
+ ---------------------------------------------------------------*)
+
+(*Can't get rid of x ~= #0 because it isn't continuous at zero*)
+Goalw [nsderiv_def]
+     "x ~= #0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ 2))";
+by (rtac ballI 1 THEN Asm_full_simp_tac 1 THEN Step_tac 1);
+by (forward_tac [Infinitesimal_add_not_zero] 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 2); 
+by (auto_tac (claset(),
+     simpset() addsimps [starfun_inverse_inverse, realpow_two] 
+               delsimps [hypreal_minus_mult_eq1 RS sym,
+                         hypreal_minus_mult_eq2 RS sym]));
+by (asm_full_simp_tac
+     (simpset() addsimps [hypreal_inverse_add,
+          hypreal_inverse_distrib RS sym, hypreal_minus_inverse RS sym] 
+          @ hypreal_add_ac @ hypreal_mult_ac 
+       delsimps [hypreal_minus_mult_eq1 RS sym,
+                 hypreal_minus_mult_eq2 RS sym] ) 1);
+by (asm_simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym,
+                                      hypreal_add_mult_distrib2] 
+         delsimps [hypreal_minus_mult_eq1 RS sym, 
+                   hypreal_minus_mult_eq2 RS sym]) 1);
+by (res_inst_tac [("y"," inverse(- hypreal_of_real x * hypreal_of_real x)")] 
+                 inf_close_trans 1);
+by (rtac inverse_add_Infinitesimal_inf_close2 1);
+by (auto_tac (claset() addSDs [hypreal_of_real_HFinite_diff_Infinitesimal], 
+         simpset() addsimps [hypreal_minus_inverse RS sym,
+                             HFinite_minus_iff]));
+by (rtac Infinitesimal_HFinite_mult2 1); 
+by Auto_tac;  
+qed "NSDERIV_inverse";
+
+
+Goal "x ~= #0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ 2))";
+by (asm_simp_tac (simpset() addsimps [NSDERIV_inverse,
+         NSDERIV_DERIV_iff RS sym] delsimps [realpow_Suc]) 1);
+qed "DERIV_inverse";
+
+(*--------------------------------------------------------------
+        Derivative of inverse 
+ -------------------------------------------------------------*)
+Goal "[| DERIV f x :> d; f(x) ~= #0 |] \
+\     ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))";
+by (rtac (real_mult_commute RS subst) 1);
+by (asm_simp_tac (simpset() addsimps [real_minus_mult_eq1,
+    realpow_inverse] delsimps [realpow_Suc, 
+    real_minus_mult_eq1 RS sym]) 1);
+by (fold_goals_tac [o_def]);
+by (blast_tac (claset() addSIs [DERIV_chain,DERIV_inverse]) 1);
+qed "DERIV_inverse_fun";
+
+Goal "[| NSDERIV f x :> d; f(x) ~= #0 |] \
+\     ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))";
+by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,
+            DERIV_inverse_fun] delsimps [realpow_Suc]) 1);
+qed "NSDERIV_inverse_fun";
+
+(*--------------------------------------------------------------
+        Derivative of quotient 
+ -------------------------------------------------------------*)
+Goal "[| DERIV f x :> d; DERIV g x :> e; g(x) ~= #0 |] \
+\      ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) + -(e*f(x))) / (g(x) ^ 2)";
+by (dres_inst_tac [("f","g")] DERIV_inverse_fun 1);
+by (dtac DERIV_mult 2);
+by (REPEAT(assume_tac 1));
+by (asm_full_simp_tac
+    (simpset() addsimps [real_divide_def, real_add_mult_distrib2,
+                         realpow_inverse,real_minus_mult_eq1] @ real_mult_ac 
+       delsimps [realpow_Suc, real_minus_mult_eq1 RS sym,
+                 real_minus_mult_eq2 RS sym]) 1);
+qed "DERIV_quotient";
+
+Goal "[| NSDERIV f x :> d; DERIV g x :> e; g(x) ~= #0 |] \
+\      ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) \
+\                           + -(e*f(x))) / (g(x) ^ 2)";
+by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,
+            DERIV_quotient] delsimps [realpow_Suc]) 1);
+qed "NSDERIV_quotient";
+ 
+(* ------------------------------------------------------------------------ *)
+(* Caratheodory formulation of derivative at a point: standard proof        *)
+(* ------------------------------------------------------------------------ *)
+
+Goal "(DERIV f x :> l) = \
+\     (EX g. (ALL z. f z - f x = g z * (z - x)) & isCont g x & g x = l)";
+by (Step_tac 1);
+by (res_inst_tac 
+    [("x","%z. if  z = x then l else (f(z) - f(x)) / (z - x)")] exI 1);
+by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,
+    ARITH_PROVE "z ~= x ==> z - x ~= (#0::real)"]));
+by (auto_tac (claset(),simpset() addsimps [isCont_iff,DERIV_iff]));
+by (ALLGOALS(rtac (LIM_equal RS iffD1)));
+by (auto_tac (claset(),simpset() addsimps [real_diff_def,real_mult_assoc]));
+qed "CARAT_DERIV";
+
+Goal "NSDERIV f x :> l ==> \
+\     EX g. (ALL z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l";
+by (auto_tac (claset(),simpset() addsimps [NSDERIV_DERIV_iff,
+    isNSCont_isCont_iff,CARAT_DERIV]));
+qed "CARAT_NSDERIV";
+
+(* How about a NS proof? *)
+Goal "(ALL z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l \
+\     ==> NSDERIV f x :> l";
+by (auto_tac (claset(), 
+              simpset() delsimprocs real_cancel_factor
+                        addsimps [NSDERIV_iff2]));
+by (auto_tac (claset(),
+              simpset() addsimps [hypreal_mult_assoc]));
+by (asm_full_simp_tac (simpset() addsimps [hypreal_eq_minus_iff3 RS sym,
+                                           hypreal_diff_def]) 1);
+by (asm_full_simp_tac (simpset() addsimps [isNSCont_def]) 1);
+qed "CARAT_DERIVD";
+ 
+
+
+(*--------------------------------------------------------------------------*)
+(* Lemmas about nested intervals and proof by bisection (cf.Harrison)       *)
+(* All considerably tidied by lcp                                           *)
+(*--------------------------------------------------------------------------*)
+
+Goal "(ALL n. (f::nat=>real) n <= f (Suc n)) --> f m <= f(m + no)";
+by (induct_tac "no" 1);
+by (auto_tac (claset() addIs [order_trans], simpset()));
+qed_spec_mp "lemma_f_mono_add";
+
+Goal "[| ALL n. f(n) <= f(Suc n); \
+\        ALL n. g(Suc n) <= g(n); \
+\        ALL n. f(n) <= g(n) |] \
+\     ==> Bseq f";
+by (res_inst_tac [("k","f 0"),("K","g 0")] BseqI2 1 THEN rtac allI 1);
+by (induct_tac "n" 1);
+by (auto_tac (claset() addIs [order_trans], simpset()));
+by (res_inst_tac [("y","g(Suc na)")] order_trans 1);
+by (induct_tac "na" 2);
+by (auto_tac (claset() addIs [order_trans], simpset()));
+qed "f_inc_g_dec_Beq_f";
+
+Goal "[| ALL n. f(n) <= f(Suc n); \
+\        ALL n. g(Suc n) <= g(n); \
+\        ALL n. f(n) <= g(n) |] \
+\     ==> Bseq g";
+by (stac (Bseq_minus_iff RS sym) 1);
+by (res_inst_tac [("g","%x. -(f x)")] f_inc_g_dec_Beq_f 1); 
+by Auto_tac;  
+qed "f_inc_g_dec_Beq_g";
+
+Goal "[| ALL n. f n <= f (Suc n);  convergent f |] ==> f n <= lim f";
+by (rtac real_leI 1);
+by (auto_tac (claset(), 
+      simpset() addsimps [convergent_LIMSEQ_iff, LIMSEQ_iff, monoseq_Suc]));
+by (dtac real_less_sum_gt_zero 1);
+by (dres_inst_tac [("x","f n + - lim f")] spec 1);
+by Safe_tac;
+by (dres_inst_tac [("P","%na. no<=na --> ?Q na"),("x","no + n")] spec 2);
+by Auto_tac;
+by (subgoal_tac "lim f <= f(no + n)" 1);
+by (induct_tac "no" 2);
+by (auto_tac (claset() addIs [order_trans],
+              simpset() addsimps [real_diff_def, real_abs_def]));
+by (dres_inst_tac [("x","f(no + n)"),("no1","no")] 
+    (lemma_f_mono_add RSN (2,order_less_le_trans)) 1);
+by (auto_tac (claset(), simpset() addsimps [add_commute]));
+qed "f_inc_imp_le_lim";
+
+Goal "convergent g ==> lim (%x. - g x) = - (lim g)";
+by (rtac (LIMSEQ_minus RS limI) 1); 
+by (asm_full_simp_tac (simpset() addsimps [convergent_LIMSEQ_iff]) 1); 
+qed "lim_uminus";
+
+Goal "[| ALL n. g(Suc n) <= g(n);  convergent g |] ==> lim g <= g n";
+by (subgoal_tac "- (g n) <= - (lim g)" 1);
+by (cut_inst_tac [("f", "%x. - (g x)")] f_inc_imp_le_lim 2);
+by (auto_tac (claset(), 
+              simpset() addsimps [lim_uminus, convergent_minus_iff RS sym]));  
+qed "g_dec_imp_lim_le";
+
+Goal "[| ALL n. f(n) <= f(Suc n); \
+\        ALL n. g(Suc n) <= g(n); \
+\        ALL n. f(n) <= g(n) |] \
+\     ==> EX l m. l <= m &  ((ALL n. f(n) <= l) & f ----> l) & \
+\                           ((ALL n. m <= g(n)) & g ----> m)";
+by (subgoal_tac "monoseq f & monoseq g" 1);
+by (force_tac (claset(), simpset() addsimps [LIMSEQ_iff,monoseq_Suc]) 2);
+by (subgoal_tac "Bseq f & Bseq g" 1);
+by (blast_tac (claset() addIs [f_inc_g_dec_Beq_f, f_inc_g_dec_Beq_g]) 2); 
+by (auto_tac (claset() addSDs [Bseq_monoseq_convergent],
+              simpset() addsimps [convergent_LIMSEQ_iff]));
+by (res_inst_tac [("x","lim f")] exI 1);
+by (res_inst_tac [("x","lim g")] exI 1);
+by (auto_tac (claset() addIs [LIMSEQ_le], simpset()));
+by (auto_tac (claset(), 
+              simpset() addsimps [f_inc_imp_le_lim, g_dec_imp_lim_le, 
+                                  convergent_LIMSEQ_iff]));  
+qed "lemma_nest";
+
+Goal "[| ALL n. f(n) <= f(Suc n); \
+\        ALL n. g(Suc n) <= g(n); \
+\        ALL n. f(n) <= g(n); \
+\        (%n. f(n) - g(n)) ----> #0 |] \
+\     ==> EX l. ((ALL n. f(n) <= l) & f ----> l) & \
+\               ((ALL n. l <= g(n)) & g ----> l)";
+by (dtac lemma_nest 1 THEN Auto_tac);
+by (subgoal_tac "l = m" 1);
+by (dres_inst_tac [("X","f")] LIMSEQ_diff 2);
+by (auto_tac (claset() addIs [LIMSEQ_unique], simpset()));
+qed "lemma_nest_unique";
+
+
+Goal "a <= b ==> \
+\  ALL n. fst (Bolzano_bisect P a b n) <= snd (Bolzano_bisect P a b n)";
+by (rtac allI 1);
+by (induct_tac "n" 1);
+by (auto_tac (claset(), simpset() addsimps [Let_def, split_def]));  
+qed "Bolzano_bisect_le";
+
+Goal "a <= b ==> \
+\  ALL n. fst(Bolzano_bisect P a b n) <= fst (Bolzano_bisect P a b (Suc n))";
+by (rtac allI 1);
+by (induct_tac "n" 1);
+by (auto_tac (claset(), 
+              simpset() addsimps [Bolzano_bisect_le, Let_def, split_def]));  
+qed "Bolzano_bisect_fst_le_Suc";
+
+Goal "a <= b ==> \
+\  ALL n. snd(Bolzano_bisect P a b (Suc n)) <= snd (Bolzano_bisect P a b n)";
+by (rtac allI 1);
+by (induct_tac "n" 1);
+by (auto_tac (claset(), 
+              simpset() addsimps [Bolzano_bisect_le, Let_def, split_def]));  
+qed "Bolzano_bisect_Suc_le_snd";
+
+Goal "((x::real) = y / (#2 * z)) = (#2 * x = y/z)";
+by Auto_tac;  
+by (dres_inst_tac [("f","%u. (#1/#2)*u")] arg_cong 1); 
+by Auto_tac;  
+qed "eq_divide_2_times_iff";
+
+Goal "a <= b ==> \
+\     snd(Bolzano_bisect P a b n) - fst(Bolzano_bisect P a b n) = \
+\     (b-a) / (#2 ^ n)";
+by (induct_tac "n" 1);
+by (auto_tac (claset(), 
+      simpset() addsimps [eq_divide_2_times_iff, real_add_divide_distrib, 
+                          Let_def, split_def]));
+by (auto_tac (claset(), 
+              simpset() addsimps (real_add_ac@[Bolzano_bisect_le, real_diff_def])));
+qed "Bolzano_bisect_diff";
+
+val Bolzano_nest_unique =
+    [Bolzano_bisect_fst_le_Suc, Bolzano_bisect_Suc_le_snd, Bolzano_bisect_le] 
+    MRS lemma_nest_unique;
+
+(*P_prem is a looping simprule, so it works better if it isn't an assumption*)
+val P_prem::notP_prem::rest =
+Goal "[| !!a b c. [| P(a,b); P(b,c); a <= b; b <= c|] ==> P(a,c); \
+\        ~ P(a,b);  a <= b |] ==> \
+\     ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))";
+by (cut_facts_tac rest 1);
+by (induct_tac "n" 1);
+by (auto_tac (claset(), 
+              simpset() delsimps [surjective_pairing RS sym]
+			addsimps [notP_prem, Let_def, split_def]));  
+by (swap_res_tac [P_prem] 1);
+by (assume_tac 1); 
+by (auto_tac (claset(), simpset() addsimps [Bolzano_bisect_le]));  
+qed "not_P_Bolzano_bisect";
+
+(*Now we re-package P_prem as a formula*)
+Goal "[| ALL a b c. P(a,b) & P(b,c) & a <= b & b <= c --> P(a,c); \
+\        ~ P(a,b);  a <= b |] ==> \
+\     ALL n. ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))";
+by (blast_tac (claset() addSEs [not_P_Bolzano_bisect RSN (2,rev_notE)]) 1); 
+qed "not_P_Bolzano_bisect'";
+
+
+Goal "[| ALL a b c. P(a,b) & P(b,c) & a <= b & b <= c --> P(a,c); \
+\        ALL x. EX d::real. #0 < d & \
+\               (ALL a b. a <= x & x <= b & (b - a) < d --> P(a,b)); \
+\        a <= b |]  \
+\     ==> P(a,b)";
+by (rtac (inst "P1" "P" Bolzano_nest_unique RS exE) 1);
+by (REPEAT (assume_tac 1));
+by (rtac LIMSEQ_minus_cancel 1);
+by (asm_simp_tac (simpset() addsimps [Bolzano_bisect_diff,
+                                      LIMSEQ_divide_realpow_zero]) 1); 
+by (rtac ccontr 1);
+by (dtac not_P_Bolzano_bisect' 1); 
+by (REPEAT (assume_tac 1));
+by (rename_tac "l" 1);
+by (dres_inst_tac [("x","l")] spec 1 THEN Clarify_tac 1);
+by (rewtac LIMSEQ_def);
+by (dres_inst_tac [("P", "%r. #0<r --> ?Q r"), ("x","d/#2")] spec 1);
+by (dres_inst_tac [("P", "%r. #0<r --> ?Q r"), ("x","d/#2")] spec 1);
+by (dtac real_less_half_sum 1);
+by Safe_tac;
+(*linear arithmetic bug if we just use Asm_simp_tac*)
+by (ALLGOALS Asm_full_simp_tac);
+by (dres_inst_tac [("x","fst(Bolzano_bisect P a b (no + noa))")] spec 1);
+by (dres_inst_tac [("x","snd(Bolzano_bisect P a b (no + noa))")] spec 1);
+by Safe_tac;
+by (ALLGOALS Asm_simp_tac);
+by (res_inst_tac [("y","abs(fst(Bolzano_bisect P a b(no + noa)) - l) + \
+\                       abs(snd(Bolzano_bisect P a b(no + noa)) - l)")] 
+    order_le_less_trans 1);
+by (asm_simp_tac (simpset() addsimps [real_abs_def]) 1);  
+by (rtac (real_sum_of_halves RS subst) 1);
+by (rtac real_add_less_mono 1);
+by (ALLGOALS 
+    (asm_full_simp_tac (simpset() addsimps [symmetric real_diff_def])));
+qed "lemma_BOLZANO";
+
+
+Goal "((ALL a b c. (a <= b & b <= c & P(a,b) & P(b,c)) --> P(a,c)) & \
+\      (ALL x. EX d::real. #0 < d & \
+\               (ALL a b. a <= x & x <= b & (b - a) < d --> P(a,b)))) \
+\     --> (ALL a b. a <= b --> P(a,b))";
+by (Clarify_tac 1);
+by (blast_tac (claset() addIs [lemma_BOLZANO]) 1); 
+qed "lemma_BOLZANO2";
+
+
+(*----------------------------------------------------------------------------*)
+(* Intermediate Value Theorem (prove contrapositive by bisection)             *)
+(*----------------------------------------------------------------------------*)
+
+Goal "[| f(a) <= y & y <= f(b); \
+\        a <= b; \
+\        (ALL x. a <= x & x <= b --> isCont f x) |] \
+\     ==> EX x. a <= x & x <= b & f(x) = y";
+by (rtac contrapos_pp 1);
+by (assume_tac 1);
+by (cut_inst_tac
+    [("P","%(u,v). a <= u & u <= v & v <= b --> ~(f(u) <= y & y <= f(v))")] 
+    lemma_BOLZANO2 1);
+by (Step_tac 1);
+by (ALLGOALS(Asm_full_simp_tac));
+by (Blast_tac 2);
+by (asm_full_simp_tac (simpset() addsimps [isCont_iff,LIM_def]) 1);
+by (rtac ccontr 1);
+by (subgoal_tac "a <= x & x <= b" 1);
+by (Asm_full_simp_tac 2);
+by (dres_inst_tac [("P", "%d. #0<d --> ?P d"),("x","#1")] spec 2);
+by (Step_tac 2);
+by (Asm_full_simp_tac 2);
+by (Asm_full_simp_tac 2);
+by (REPEAT(blast_tac (claset() addIs [order_trans]) 2));
+by (REPEAT(dres_inst_tac [("x","x")] spec 1));
+by (Asm_full_simp_tac 1);
+by (dres_inst_tac [("P", "%r. ?P r --> (EX s. #0<s & ?Q r s)"),
+                   ("x","abs(y - f x)")] spec 1);
+by (Step_tac 1);
+by (asm_full_simp_tac (simpset() addsimps []) 1);
+by (dres_inst_tac [("x","s")] spec 1);
+by (Clarify_tac 1);
+by (cut_inst_tac [("R1.0","f x"),("R2.0","y")] real_linear 1);
+by Safe_tac;
+by (dres_inst_tac [("x","ba - x")] spec 1);
+by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [abs_iff])));
+by (dres_inst_tac [("x","aa - x")] spec 1);
+by (case_tac "x <= aa" 1);
+by (ALLGOALS Asm_full_simp_tac);
+by (dres_inst_tac [("z","x"),("w","aa")] real_le_anti_sym 1);
+by (assume_tac 1 THEN Asm_full_simp_tac 1);
+qed "IVT";
+
+
+Goal "[| f(b) <= y & y <= f(a); \
+\        a <= b; \
+\        (ALL x. a <= x & x <= b --> isCont f x) \
+\     |] ==> EX x. a <= x & x <= b & f(x) = y";
+by (subgoal_tac "- f a <= -y & -y <= - f b" 1);
+by (thin_tac "f b <= y & y <= f a" 1);
+by (dres_inst_tac [("f","%x. - f x")] IVT 1);
+by (auto_tac (claset() addIs [isCont_minus],simpset()));
+qed "IVT2";
+
+
+(*HOL style here: object-level formulations*)
+Goal "(f(a) <= y & y <= f(b) & a <= b & \
+\     (ALL x. a <= x & x <= b --> isCont f x)) \
+\     --> (EX x. a <= x & x <= b & f(x) = y)";
+by (blast_tac (claset() addIs [IVT]) 1);
+qed "IVT_objl";
+
+Goal "(f(b) <= y & y <= f(a) & a <= b & \
+\     (ALL x. a <= x & x <= b --> isCont f x)) \
+\     --> (EX x. a <= x & x <= b & f(x) = y)";
+by (blast_tac (claset() addIs [IVT2]) 1);
+qed "IVT2_objl";
+
+(*---------------------------------------------------------------------------*)
+(* By bisection, function continuous on closed interval is bounded above     *)
+(*---------------------------------------------------------------------------*)
+
+Goal "abs (real_of_nat x) = real_of_nat x";
+by (auto_tac (claset() addIs [abs_eqI1],simpset()));
+qed "abs_real_of_nat_cancel";
+Addsimps [abs_real_of_nat_cancel];
+
+Goal "~ abs(x) + 1r < x";
+by (rtac real_leD 1);
+by (auto_tac (claset() addIs [abs_ge_self RS order_trans],simpset()));
+qed "abs_add_one_not_less_self";
+Addsimps [abs_add_one_not_less_self];
+
+
+Goal "[| a <= b; ALL x. a <= x & x <= b --> isCont f x |]\
+\     ==> EX M. ALL x. a <= x & x <= b --> f(x) <= M";
+by (cut_inst_tac [("P","%(u,v). a <= u & u <= v & v <= b --> \
+\                         (EX M. ALL x. u <= x & x <= v --> f x <= M)")] 
+    lemma_BOLZANO2 1);
+by (Step_tac 1);
+by (ALLGOALS(Asm_full_simp_tac));
+by (cut_inst_tac [("x","M"),("y","Ma")] 
+    (CLAIM "x <= y | y <= (x::real)") 1);
+by (Step_tac 1);
+by (res_inst_tac [("x","Ma")] exI 1);
+by (Step_tac 1);
+by (cut_inst_tac [("x","xb"),("y","xa")] 
+    (CLAIM "x <= y | y <= (x::real)") 1);
+by (Step_tac 1);
+by (rtac order_trans 1 THEN assume_tac 2);
+by (Asm_full_simp_tac 1);
+by (Asm_full_simp_tac 1);
+by (res_inst_tac [("x","M")] exI 1);
+by (Step_tac 1);
+by (cut_inst_tac [("x","xb"),("y","xa")] 
+    (CLAIM "x <= y | y <= (x::real)") 1);
+by (Step_tac 1);
+by (Asm_full_simp_tac 1);
+by (rtac order_trans 1 THEN assume_tac 2);
+by (Asm_full_simp_tac 1);
+by (case_tac "a <= x & x <= b" 1);
+by (res_inst_tac [("x","#1")] exI 2);
+by (auto_tac (claset(),
+              simpset() addsimps [LIM_def,isCont_iff]));
+by (dres_inst_tac [("x","x")] spec 1 THEN Auto_tac);
+by (thin_tac "ALL M. EX x. a <= x & x <= b & ~ f x <= M" 1);
+by (dres_inst_tac [("x","#1")] spec 1);
+by Auto_tac;  
+by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
+by (res_inst_tac [("x","abs(f x) + #1")] exI 1 THEN Step_tac 1);
+by (dres_inst_tac [("x","xa - x")] spec 1 THEN Auto_tac);
+by (res_inst_tac [("y","abs(f xa)")] order_trans 3);
+by (res_inst_tac [("y","abs(f x) + abs(f(xa) - f(x))")] order_trans 4);
+by (auto_tac (claset() addIs [abs_triangle_ineq RSN (2, order_trans)],
+              simpset() addsimps [real_diff_def,abs_ge_self]));
+by (auto_tac (claset(),
+              simpset() addsimps [real_abs_def] addsplits [split_if_asm]));
+qed "isCont_bounded";
+
+(*----------------------------------------------------------------------------*)
+(* Refine the above to existence of least upper bound                         *)
+(*----------------------------------------------------------------------------*)
+
+Goal "((EX x. x : S) & (EX y. isUb UNIV S (y::real))) --> \
+\     (EX t. isLub UNIV S t)";
+by (blast_tac (claset() addIs [reals_complete]) 1);
+qed "lemma_reals_complete";
+
+Goal "[| a <= b; ALL x. a <= x & x <= b --> isCont f x |] \
+\        ==> EX M. (ALL x. a <= x & x <= b --> f(x) <= M) & \
+\                  (ALL N. N < M --> (EX x. a <= x & x <= b & N < f(x)))";
+by (cut_inst_tac [("S","Collect (%y. EX x. a <= x & x <= b & y = f x)")]
+    lemma_reals_complete 1);
+by Auto_tac;
+by (dtac isCont_bounded 1 THEN assume_tac 1);
+by (auto_tac (claset(),simpset() addsimps [isUb_def,leastP_def,
+    isLub_def,setge_def,setle_def]));
+by (rtac exI 1 THEN Auto_tac);
+by (REPEAT(dtac spec 1) THEN Auto_tac);
+by (dres_inst_tac [("x","x")] spec 1);
+by (auto_tac (claset() addSIs [real_leI],simpset()));
+qed "isCont_has_Ub";
+
+(*----------------------------------------------------------------------------*)
+(* Now show that it attains its upper bound                                   *)
+(*----------------------------------------------------------------------------*)
+
+Goal "[| a <= b; ALL x. a <= x & x <= b --> isCont f x |] \
+\        ==> EX M. (ALL x. a <= x & x <= b --> f(x) <= M) & \
+\                  (EX x. a <= x & x <= b & f(x) = M)";
+by (ftac isCont_has_Ub 1 THEN assume_tac 1);
+by (Clarify_tac 1);
+by (res_inst_tac [("x","M")] exI 1);
+by (Asm_full_simp_tac 1); 
+by (rtac ccontr 1);
+by (subgoal_tac "ALL x. a <= x & x <= b --> f x < M" 1 THEN Step_tac 1);
+by (rtac ccontr 2 THEN dtac real_leI 2);
+by (dres_inst_tac [("z","M")] real_le_anti_sym 2);
+by (REPEAT(Blast_tac 2));
+by (subgoal_tac "ALL x. a <= x & x <= b --> isCont (%x. inverse(M - f x)) x" 1);
+by (Step_tac 1);
+by (EVERY[rtac isCont_inverse 2, rtac isCont_diff 2, rtac notI 4]);
+by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [real_diff_eq_eq])));
+by (Blast_tac 2);
+by (subgoal_tac 
+    "EX k. ALL x. a <= x & x <= b --> (%x. inverse(M - (f x))) x <= k" 1);
+by (rtac isCont_bounded 2);
+by (Step_tac 1);
+by (subgoal_tac "ALL x. a <= x & x <= b --> #0 < inverse(M - f(x))" 1);
+by (Asm_full_simp_tac 1); 
+by (Step_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [real_less_diff_eq]) 2);
+by (subgoal_tac 
+    "ALL x. a <= x & x <= b --> (%x. inverse(M - (f x))) x < (k + #1)" 1);
+by (Step_tac 1);
+by (res_inst_tac [("y","k")] order_le_less_trans 2);
+by (asm_full_simp_tac (simpset() addsimps [real_zero_less_one]) 3);
+by (Asm_full_simp_tac 2); 
+by (subgoal_tac "ALL x. a <= x & x <= b --> \
+\                inverse(k + #1) < inverse((%x. inverse(M - (f x))) x)" 1);
+by (Step_tac 1);
+by (rtac real_inverse_less_swap 2);
+by (ALLGOALS Asm_full_simp_tac);
+by (dres_inst_tac [("P", "%N. N<M --> ?Q N"),
+                   ("x","M - inverse(k + #1)")] spec 1);
+by (Step_tac 1 THEN dtac real_leI 1);
+by (dtac (real_le_diff_eq RS iffD1) 1);
+by (REPEAT(dres_inst_tac [("x","a")] spec 1));
+by (Asm_full_simp_tac 1);
+by (asm_full_simp_tac 
+    (simpset() addsimps [real_inverse_eq_divide, pos_real_divide_le_eq]) 1); 
+by (cut_inst_tac [("x","k"),("y","M-f a")] real_0_less_mult_iff 1);
+by (Asm_full_simp_tac 1); 
+(*last one*)
+by (REPEAT(dres_inst_tac [("x","x")] spec 1));
+by (Asm_full_simp_tac 1); 
+qed "isCont_eq_Ub";
+
+
+(*----------------------------------------------------------------------------*)
+(* Same theorem for lower bound                                               *)
+(*----------------------------------------------------------------------------*)
+
+Goal "[| a <= b; ALL x. a <= x & x <= b --> isCont f x |] \
+\        ==> EX M. (ALL x. a <= x & x <= b --> M <= f(x)) & \
+\                  (EX x. a <= x & x <= b & f(x) = M)";
+by (subgoal_tac "ALL x. a <= x & x <= b --> isCont (%x. -(f x)) x" 1);
+by (blast_tac (claset() addIs [isCont_minus]) 2);
+by (dres_inst_tac [("f","(%x. -(f x))")] isCont_eq_Ub 1);
+by (Step_tac 1);
+by Auto_tac;
+qed "isCont_eq_Lb";
+
+
+(* ------------------------------------------------------------------------- *)
+(* Another version.                                                          *)
+(* ------------------------------------------------------------------------- *)
+
+Goal "[|a <= b; ALL x. a <= x & x <= b --> isCont f x |] \
+\     ==> EX L M. (ALL x. a <= x & x <= b --> L <= f(x) & f(x) <= M) & \
+\         (ALL y. L <= y & y <= M --> (EX x. a <= x & x <= b & (f(x) = y)))";
+by (ftac isCont_eq_Lb 1);
+by (ftac isCont_eq_Ub 2);
+by (REPEAT(assume_tac 1));
+by (Step_tac 1);
+by (res_inst_tac [("x","f x")] exI 1);
+by (res_inst_tac [("x","f xa")] exI 1);
+by (Asm_full_simp_tac 1);
+by (Step_tac 1);
+by (cut_inst_tac [("x","x"),("y","xa")] (CLAIM "x <= y | y <= (x::real)") 1);
+by (Step_tac 1);
+by (cut_inst_tac [("f","f"),("a","x"),("b","xa"),("y","y")] IVT_objl 1);
+by (cut_inst_tac [("f","f"),("a","xa"),("b","x"),("y","y")] IVT2_objl 2);
+by (Step_tac 1);
+by (res_inst_tac [("x","xb")] exI 2);
+by (res_inst_tac [("x","xb")] exI 4);
+by (ALLGOALS(Asm_full_simp_tac));
+qed "isCont_Lb_Ub";
+
+(*----------------------------------------------------------------------------*)
+(* If f'(x) > 0 then x is locally strictly increasing at the right            *)
+(*----------------------------------------------------------------------------*)
+
+Goalw [deriv_def,LIM_def] 
+    "[| DERIV f x :> l;  #0 < l |] ==> \
+\      EX d. #0 < d & (ALL h. #0 < h & h < d --> f(x) < f(x + h))";
+by (dtac spec 1 THEN Auto_tac);
+by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
+by (subgoal_tac "#0 < l*h" 1);
+by (asm_full_simp_tac (simpset() addsimps [real_0_less_mult_iff]) 2); 
+by (dres_inst_tac [("x","h")] spec 1);
+by (asm_full_simp_tac
+    (simpset() addsimps [real_abs_def, real_inverse_eq_divide, 
+                 pos_real_le_divide_eq, pos_real_less_divide_eq]
+              addsplits [split_if_asm]) 1); 
+qed "DERIV_left_inc";
+
+Goalw [deriv_def,LIM_def] 
+    "[| DERIV f x :> l;  l < #0 |] ==> \
+\      EX d. #0 < d & (ALL h. #0 < h & h < d --> f(x) < f(x - h))";
+by (dres_inst_tac [("x","-l")] spec 1 THEN Auto_tac);
+by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
+by (subgoal_tac "l*h < #0" 1);
+by (asm_full_simp_tac (simpset() addsimps [real_mult_less_0_iff]) 2); 
+by (dres_inst_tac [("x","-h")] spec 1);
+by (asm_full_simp_tac
+    (simpset() addsimps [real_abs_def, real_inverse_eq_divide, 
+                         pos_real_less_divide_eq,
+                         symmetric real_diff_def]
+               addsplits [split_if_asm]) 1);
+by (subgoal_tac "#0 < (f (x - h) - f x)/h" 1);
+by (arith_tac 2);
+by (asm_full_simp_tac
+    (simpset() addsimps [pos_real_less_divide_eq]) 1); 
+qed "DERIV_left_dec";
+
+
+Goal "[| DERIV f x :> l; \
+\        EX d. #0 < d & (ALL y. abs(x - y) < d --> f(y) <= f(x)) |] \
+\     ==> l = #0";
+by (res_inst_tac [("R1.0","l"),("R2.0","#0")] real_linear_less2 1);
+by (Step_tac 1);
+by (dtac DERIV_left_dec 1);
+by (dtac DERIV_left_inc 3);
+by (Step_tac 1);
+by (dres_inst_tac [("d1.0","d"),("d2.0","da")] real_lbound_gt_zero 1);
+by (dres_inst_tac [("d1.0","d"),("d2.0","da")] real_lbound_gt_zero 3);
+by (Step_tac 1);
+by (dres_inst_tac [("x","x - e")] spec 1);
+by (dres_inst_tac [("x","x + e")] spec 2);
+by (auto_tac (claset(), simpset() addsimps [real_abs_def]));
+qed "DERIV_local_max";
+
+(*----------------------------------------------------------------------------*)
+(* Similar theorem for a local minimum                                        *)
+(*----------------------------------------------------------------------------*)
+
+Goal "[| DERIV f x :> l; \
+\        EX d::real. #0 < d & (ALL y. abs(x - y) < d --> f(x) <= f(y)) |] \
+\     ==> l = #0";
+by (dtac (DERIV_minus RS DERIV_local_max) 1); 
+by Auto_tac;  
+qed "DERIV_local_min";
+
+(*----------------------------------------------------------------------------*)
+(* In particular if a function is locally flat                                *)
+(*----------------------------------------------------------------------------*)
+
+Goal "[| DERIV f x :> l; \
+\        EX d. #0 < d & (ALL y. abs(x - y) < d --> f(x) = f(y)) |] \
+\     ==> l = #0";
+by (auto_tac (claset() addSDs [DERIV_local_max],simpset()));
+qed "DERIV_local_const";
+
+(*----------------------------------------------------------------------------*)
+(* Lemma about introducing open ball in open interval                         *)
+(*----------------------------------------------------------------------------*)
+
+Goal "[| a < x;  x < b |] ==> \
+\       EX d::real. #0 < d &  (ALL y. abs(x - y) < d --> a < y & y < b)";
+by (simp_tac (simpset() addsimps [abs_interval_iff]) 1);
+by (cut_inst_tac [("x","x - a"),("y","b - x")] 
+    (CLAIM "x <= y | y <= (x::real)") 1);
+by (Step_tac 1);
+by (res_inst_tac [("x","x - a")] exI 1);
+by (res_inst_tac [("x","b - x")] exI 2);
+by Auto_tac;
+by (auto_tac (claset(),simpset() addsimps [real_less_diff_eq]));
+qed "lemma_interval_lt";
+
+Goal "[| a < x;  x < b |] ==> \
+\       EX d::real. #0 < d &  (ALL y. abs(x - y) < d --> a <= y & y <= b)";
+by (dtac lemma_interval_lt 1);
+by Auto_tac;
+by (auto_tac (claset() addSIs [exI] ,simpset()));
+qed "lemma_interval";
+
+(*-----------------------------------------------------------------------
+            Rolle's Theorem
+   If f is defined and continuous on the finite closed interval [a,b]
+   and differentiable a least on the open interval (a,b), and f(a) = f(b),
+   then x0 : (a,b) such that f'(x0) = #0
+ ----------------------------------------------------------------------*)
+
+Goal "[| a < b; f(a) = f(b); \
+\        ALL x. a <= x & x <= b --> isCont f x; \
+\        ALL x. a < x & x < b --> f differentiable x \
+\     |] ==> EX z. a < z & z < b & DERIV f z :> #0";
+by (ftac (order_less_imp_le RS isCont_eq_Ub) 1);
+by (EVERY1[assume_tac,Step_tac]);
+by (ftac (order_less_imp_le RS isCont_eq_Lb) 1);
+by (EVERY1[assume_tac,Step_tac]);
+by (case_tac "a < x & x < b" 1 THEN etac conjE 1);
+by (Asm_full_simp_tac 2);
+by (forw_inst_tac [("a","a"),("x","x")] lemma_interval 1);
+by (EVERY1[assume_tac,etac exE]);
+by (res_inst_tac [("x","x")] exI 1 THEN Asm_full_simp_tac 1);
+by (subgoal_tac "(EX l. DERIV f x :> l) & \
+\        (EX d. #0 < d & (ALL y. abs(x - y) < d --> f(y) <= f(x)))" 1);
+by (Clarify_tac 1 THEN rtac conjI 2);
+by (blast_tac (claset() addIs [differentiableD]) 2);
+by (Blast_tac 2);
+by (ftac DERIV_local_max 1);
+by (EVERY1[Blast_tac,Blast_tac]);
+by (case_tac "a < xa & xa < b" 1 THEN etac conjE 1);
+by (Asm_full_simp_tac 2);
+by (forw_inst_tac [("a","a"),("x","xa")] lemma_interval 1);
+by (EVERY1[assume_tac,etac exE]);
+by (res_inst_tac [("x","xa")] exI 1 THEN Asm_full_simp_tac 1);
+by (subgoal_tac "(EX l. DERIV f xa :> l) & \
+\        (EX d. #0 < d & (ALL y. abs(xa - y) < d --> f(xa) <= f(y)))" 1);
+by (Clarify_tac 1 THEN rtac conjI 2);
+by (blast_tac (claset() addIs [differentiableD]) 2);
+by (Blast_tac 2);
+by (ftac DERIV_local_min 1);
+by (EVERY1[Blast_tac,Blast_tac]);
+by (subgoal_tac "ALL x. a <= x & x <= b --> f(x) = f(b)" 1);
+by (Clarify_tac 2);
+by (rtac real_le_anti_sym 2);
+by (subgoal_tac "f b = f x" 2);
+by (Asm_full_simp_tac 2);
+by (res_inst_tac [("x1","a"),("y1","x")] (order_le_imp_less_or_eq RS disjE) 2);
+by (assume_tac 2);
+by (dres_inst_tac [("z","x"),("w","b")] real_le_anti_sym 2);
+by (subgoal_tac "f b = f xa" 5);
+by (Asm_full_simp_tac 5);
+by (res_inst_tac [("x1","a"),("y1","xa")] (order_le_imp_less_or_eq RS disjE) 5);
+by (assume_tac 5);
+by (dres_inst_tac [("z","xa"),("w","b")] real_le_anti_sym 5);
+by (REPEAT(Asm_full_simp_tac 2));
+by (dtac real_dense 1 THEN etac exE 1);
+by (res_inst_tac [("x","r")] exI 1 THEN Asm_full_simp_tac 1);
+by (etac conjE 1);
+by (forw_inst_tac [("a","a"),("x","r")] lemma_interval 1);
+by (EVERY1[assume_tac, etac exE]);
+by (subgoal_tac "(EX l. DERIV f r :> l) & \
+\        (EX d. #0 < d & (ALL y. abs(r - y) < d --> f(r) = f(y)))" 1);
+by (Clarify_tac 1 THEN rtac conjI 2);
+by (blast_tac (claset() addIs [differentiableD]) 2);
+by (EVERY1[ftac DERIV_local_const, Blast_tac, Blast_tac]);
+by (res_inst_tac [("x","d")] exI 1);
+by (EVERY1[rtac conjI, Blast_tac, rtac allI, rtac impI]);
+by (res_inst_tac [("s","f b")] trans 1);
+by (blast_tac (claset() addSDs [order_less_imp_le]) 1);
+by (rtac sym 1 THEN Blast_tac 1);
+qed "Rolle";
+
+(*----------------------------------------------------------------------------*)
+(* Mean value theorem                                                         *)
+(*----------------------------------------------------------------------------*)
+
+Goal "f a - (f b - f a)/(b - a) * a = \
+\     f b - (f b - f a)/(b - a) * (b::real)";
+by (case_tac "a = b" 1);
+by (asm_full_simp_tac (simpset() addsimps [REAL_DIVIDE_ZERO]) 1); 
+by (res_inst_tac [("c1","b - a")] (real_mult_left_cancel RS iffD1) 1);
+by (arith_tac 1);
+by (auto_tac (claset(),
+              simpset() addsimps [real_diff_mult_distrib2]));
+by (auto_tac (claset(),
+           simpset() addsimps [real_diff_mult_distrib]));
+qed "lemma_MVT";
+
+Goal "[| a < b; \
+\        ALL x. a <= x & x <= b --> isCont f x; \
+\        ALL x. a < x & x < b --> f differentiable x |] \
+\     ==>  EX l z. a < z & z < b & DERIV f z :> l & \
+\                  (f(b) - f(a) = (b - a) * l)";
+by (dres_inst_tac [("f","%x. f(x) - (((f(b) - f(a)) / (b - a)) * x)")]
+    Rolle 1);
+by (rtac lemma_MVT 1);
+by (Step_tac 1);
+by (rtac isCont_diff 1 THEN Blast_tac 1);
+by (rtac (isCont_const RS isCont_mult) 1);
+by (rtac isCont_Id 1); 
+by (dres_inst_tac [("P", "%x. ?Pre x --> f differentiable x"), 
+                   ("x","x")] spec 1);
+by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 1);
+by (Step_tac 1);
+by (res_inst_tac [("x","xa - ((f(b) - f(a)) / (b - a))")] exI 1);
+by (rtac DERIV_diff 1 THEN assume_tac 1);
+(*derivative of a linear function is the constant...*)
+by (subgoal_tac "(%x. (f b - f a) * x / (b - a)) = \
+\                op * ((f b - f a) / (b - a))" 1);
+by (rtac ext 2 THEN Simp_tac 2);
+by (Asm_full_simp_tac 1); 
+(*final case*)
+by (res_inst_tac [("x","((f(b) - f(a)) / (b - a))")] exI 1);
+by (res_inst_tac [("x","z")] exI 1);
+by (Step_tac 1);
+by (Asm_full_simp_tac 2); 
+by (subgoal_tac "DERIV (%x. ((f(b) - f(a)) / (b - a)) * x) z :> \
+\                           ((f(b) - f(a)) / (b - a))" 1);
+by (rtac DERIV_cmult_Id 2); 
+by (dtac DERIV_add 1 THEN assume_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [real_add_assoc, real_diff_def]) 1);
+qed "MVT";
+
+(*----------------------------------------------------------------------------*)
+(* Theorem that function is constant if its derivative is 0 over an interval. *)
+(*----------------------------------------------------------------------------*)
+
+Goal "[| a < b; \
+\        ALL x. a <= x & x <= b --> isCont f x; \
+\        ALL x. a < x & x < b --> DERIV f x :> #0 |] \
+\       ==> (f b = f a)";
+by (dtac MVT 1 THEN assume_tac 1);
+by (blast_tac (claset() addIs [differentiableI]) 1);
+by (auto_tac (claset() addSDs [DERIV_unique],simpset() 
+    addsimps [real_diff_eq_eq]));
+qed "DERIV_isconst_end";
+
+Goal "[| a < b; \
+\        ALL x. a <= x & x <= b --> isCont f x; \
+\        ALL x. a < x & x < b --> DERIV f x :> #0 |] \
+\       ==> ALL x. a <= x & x <= b --> f x = f a";
+by (Step_tac 1);
+by (dres_inst_tac [("x","a")] order_le_imp_less_or_eq 1);
+by (Step_tac 1);
+by (dres_inst_tac [("b","x")] DERIV_isconst_end 1);
+by Auto_tac;
+qed "DERIV_isconst1";
+
+Goal "[| a < b; \
+\        ALL x. a <= x & x <= b --> isCont f x; \
+\        ALL x. a < x & x < b --> DERIV f x :> #0; \
+\        a <= x; x <= b |] \
+\       ==> f x = f a";
+by (blast_tac (claset() addDs [DERIV_isconst1]) 1);
+qed "DERIV_isconst2";
+
+Goal "ALL x. DERIV f x :> #0 ==> f(x) = f(y)";
+by (res_inst_tac [("R1.0","x"),("R2.0","y")] real_linear_less2 1);
+by (rtac sym 1);
+by (auto_tac (claset() addIs [DERIV_isCont,DERIV_isconst_end],simpset()));
+qed "DERIV_isconst_all";
+
+Goal "[|a ~= b; ALL x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b - a) * k";
+by (res_inst_tac [("R1.0","a"),("R2.0","b")] real_linear_less2 1);
+by Auto_tac;
+by (ALLGOALS(dres_inst_tac [("f","f")] MVT));
+by (auto_tac (claset() addDs [DERIV_isCont,DERIV_unique],simpset() addsimps 
+    [differentiable_def]));
+by (auto_tac (claset() addDs [DERIV_unique],
+       simpset() addsimps [real_add_mult_distrib, real_diff_def,
+                           real_minus_mult_eq1 RS sym]));
+qed "DERIV_const_ratio_const";
+
+Goal "[|a ~= b; ALL x. DERIV f x :> k |] ==> (f(b) - f(a))/(b - a) = k";
+by (res_inst_tac [("c1","b - a")] (real_mult_right_cancel RS iffD1) 1);
+by (auto_tac (claset() addSDs [DERIV_const_ratio_const], 
+              simpset() addsimps [real_mult_assoc]));
+qed "DERIV_const_ratio_const2";
+
+Goal "((a + b) /#2 - a) = (b - a)/(#2::real)";
+by Auto_tac;  
+qed "real_average_minus_first";
+Addsimps [real_average_minus_first];
+
+Goal "((b + a)/#2 - a) = (b - a)/(#2::real)";
+by Auto_tac;  
+qed "real_average_minus_second";
+Addsimps [real_average_minus_second];
+
+
+(* Gallileo's "trick": average velocity = av. of end velocities *)
+Goal "[|a ~= (b::real); ALL x. DERIV v x :> k|] \
+\     ==> v((a + b)/#2) = (v a + v b)/#2";
+by (res_inst_tac [("R1.0","a"),("R2.0","b")] real_linear_less2 1);
+by Auto_tac;
+by (ftac DERIV_const_ratio_const2 1 THEN assume_tac 1);
+by (ftac DERIV_const_ratio_const2 2 THEN assume_tac 2);
+by (dtac real_less_half_sum 1);
+by (dtac real_gt_half_sum 2); 
+by (ftac (real_not_refl2 RS DERIV_const_ratio_const2) 1 THEN assume_tac 1);
+by (dtac ((real_not_refl2 RS not_sym) RS DERIV_const_ratio_const2) 2
+    THEN assume_tac 2);
+by (ALLGOALS (dres_inst_tac [("f","%u. (b-a)*u")] arg_cong)); 
+by (auto_tac (claset(), simpset() addsimps [real_inverse_eq_divide])); 
+by (asm_full_simp_tac (simpset() addsimps [real_add_commute, eq_commute]) 1);  
+qed "DERIV_const_average";
+
+
+(* ------------------------------------------------------------------------ *)
+(* Dull lemma that an continuous injection on an interval must have a strict*)
+(* maximum at an end point, not in the middle.                              *)
+(* ------------------------------------------------------------------------ *)
+
+Goal "[|#0 < d; ALL z. abs(z - x) <= d --> g(f z) = z; \
+\       ALL z. abs(z - x) <= d --> isCont f z |]  \
+\     ==> ~(ALL z. abs(z - x) <= d --> f(z) <= f(x))";
+by (rtac notI 1);
+by (rotate_tac 3 1);
+by (forw_inst_tac [("x","x - d")] spec 1);
+by (forw_inst_tac [("x","x + d")] spec 1);
+by (Step_tac 1);
+by (cut_inst_tac [("x","f(x - d)"),("y","f(x + d)")] 
+    (ARITH_PROVE "x <= y | y <= (x::real)") 4);
+by (etac disjE 4);
+by (REPEAT(arith_tac 1));
+by (cut_inst_tac [("f","f"),("a","x - d"),("b","x"),("y","f(x + d)")]
+    IVT_objl 1);
+by (Step_tac 1);
+by (arith_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);
+by (dres_inst_tac [("f","g")] arg_cong 1);
+by (rotate_tac 2 1);
+by (forw_inst_tac [("x","xa")] spec 1);
+by (dres_inst_tac [("x","x + d")] spec 1);
+by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);
+(* 2nd case: similar *)
+by (cut_inst_tac [("f","f"),("a","x"),("b","x + d"),("y","f(x - d)")]
+    IVT2_objl 1);
+by (Step_tac 1);
+by (arith_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);
+by (dres_inst_tac [("f","g")] arg_cong 1);
+by (rotate_tac 2 1);
+by (forw_inst_tac [("x","xa")] spec 1);
+by (dres_inst_tac [("x","x - d")] spec 1);
+by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);
+qed "lemma_isCont_inj";
+
+(* ------------------------------------------------------------------------ *)
+(* Similar version for lower bound                                          *)
+(* ------------------------------------------------------------------------ *)
+
+Goal "[|#0 < d; ALL z. abs(z - x) <= d --> g(f z) = z; \
+\       ALL z. abs(z - x) <= d --> isCont f z |]  \
+\     ==> ~(ALL z. abs(z - x) <= d --> f(x) <= f(z))";
+by (auto_tac (claset() addSDs [(asm_full_simplify (simpset()) 
+    (read_instantiate [("f","%x. - f x"),("g","%y. g(-y)"),("x","x"),("d","d")]
+     lemma_isCont_inj))],simpset() addsimps [isCont_minus]));
+qed "lemma_isCont_inj2";
+
+(* ------------------------------------------------------------------------ *)
+(* Show there's an interval surrounding f(x) in f[[x - d, x + d]]           *)
+(* Also from John's theory                                                  *)
+(* ------------------------------------------------------------------------ *)
+
+Addsimps [zero_eq_numeral_0,one_eq_numeral_1];
+
+val lemma_le = ARITH_PROVE "#0 <= (d::real) ==> -d <= d";
+
+(* FIXME: awful proof - needs improvement *)
+Goal "[| #0 < d; ALL z. abs(z - x) <= d --> g(f z) = z; \
+\        ALL z. abs(z - x) <= d --> isCont f z |] \
+\      ==> EX e. #0 < e & \
+\                 (ALL y. \
+\                     abs(y - f(x)) <= e --> \
+\                     (EX z. abs(z - x) <= d & (f z = y)))";
+by (ftac order_less_imp_le 1);
+by (dtac (lemma_le RS (asm_full_simplify (simpset()) (read_instantiate 
+    [("f","f"),("a","x - d"),("b","x + d")] isCont_Lb_Ub))) 1);
+by (Step_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);
+by (subgoal_tac "L <= f x & f x <= M" 1);
+by (dres_inst_tac [("P", "%v. ?P v --> ?Q v & ?R v"), ("x","x")] spec 2);
+by (Asm_full_simp_tac 2);
+by (subgoal_tac "L < f x & f x < M" 1);
+by (Step_tac 1);
+by (dres_inst_tac [("x","L")] (ARITH_PROVE "x < y ==> #0 < y - (x::real)") 1);
+by (dres_inst_tac [("x","f x")] (ARITH_PROVE "x < y ==> #0 < y - (x::real)") 1);
+by (dres_inst_tac [("d1.0","f x - L"),("d2.0","M - f x")] 
+    (rename_numerals real_lbound_gt_zero) 1);
+by (Step_tac 1);
+by (res_inst_tac [("x","e")] exI 1);
+by (Step_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);
+by (dres_inst_tac [("P","%v. ?PP v --> (EX xa. ?Q v xa)"),("x","y")] spec 1);
+by (Step_tac 1 THEN REPEAT(arith_tac 1));
+by (res_inst_tac [("x","xa")] exI 1);
+by (arith_tac 1);
+by (ALLGOALS(etac (ARITH_PROVE "[|x <= y; x ~= y |] ==> x < (y::real)")));
+by (ALLGOALS(rotate_tac 3));
+by (dtac lemma_isCont_inj2 1);
+by (assume_tac 2);
+by (dtac lemma_isCont_inj 3);
+by (assume_tac 4);
+by (TRYALL(assume_tac));
+by (Step_tac 1);
+by (ALLGOALS(dres_inst_tac [("x","z")] spec));
+by (ALLGOALS(arith_tac));
+qed "isCont_inj_range";
+
+
+(* ------------------------------------------------------------------------ *)
+(* Continuity of inverse function                                           *)
+(* ------------------------------------------------------------------------ *)
+
+Goal "[| #0 < d; ALL z. abs(z - x) <= d --> g(f(z)) = z; \
+\        ALL z. abs(z - x) <= d --> isCont f z |] \
+\     ==> isCont g (f x)";
+by (simp_tac (simpset() addsimps [isCont_iff,LIM_def]) 1);
+by (Step_tac 1);
+by (dres_inst_tac [("d1.0","r")] (rename_numerals real_lbound_gt_zero) 1);
+by (assume_tac 1 THEN Step_tac 1);
+by (subgoal_tac "ALL z. abs(z - x) <= e --> (g(f z) = z)" 1);
+by (Force_tac 2);
+by (subgoal_tac "ALL z. abs(z - x) <= e --> isCont f z" 1);
+by (Force_tac 2);
+by (dres_inst_tac [("d","e")] isCont_inj_range 1);
+by (assume_tac 2 THEN assume_tac 1);
+by (Step_tac 1);
+by (res_inst_tac [("x","ea")] exI 1);
+by Auto_tac;
+by (rotate_tac 4 1);
+by (dres_inst_tac [("x","f(x) + xa")] spec 1);
+by Auto_tac;
+by (dtac sym 1 THEN Auto_tac);
+by (arith_tac 1);
+qed "isCont_inverse";
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/Lim.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,79 @@
+(*  Title       : Lim.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Theory of limits, continuity and 
+                  differentiation of real=>real functions
+*)
+
+Lim = SEQ + RealAbs + 
+
+(*-----------------------------------------------------------------------
+    Limits, continuity and differentiation: standard and NS definitions
+ -----------------------------------------------------------------------*)
+
+constdefs
+  LIM :: [real=>real,real,real] => bool
+				("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60)
+  "f -- a --> L ==
+     ALL r. #0 < r --> 
+	     (EX s. #0 < s & (ALL x. (x ~= a & (abs(x + -a) < s)
+			  --> abs(f x + -L) < r)))"
+
+  NSLIM :: [real=>real,real,real] => bool
+			      ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60)
+  "f -- a --NS> L == (ALL x. (x ~= hypreal_of_real a & 
+		      x @= hypreal_of_real a -->
+		      (*f* f) x @= hypreal_of_real L))"   
+
+  isCont :: [real=>real,real] => bool
+  "isCont f a == (f -- a --> (f a))"        
+
+  (* NS definition dispenses with limit notions *)
+  isNSCont :: [real=>real,real] => bool
+  "isNSCont f a == (ALL y. y @= hypreal_of_real a --> 
+			   (*f* f) y @= hypreal_of_real (f a))"
+
+  (* differentiation: D is derivative of function f at x *)
+  deriv:: [real=>real,real,real] => bool
+			    ("(DERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60)
+  "DERIV f x :> D == ((%h. (f(x + h) + -f(x))/h) -- #0 --> D)"
+
+  nsderiv :: [real=>real,real,real] => bool
+			    ("(NSDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60)
+  "NSDERIV f x :> D == (ALL h: Infinitesimal - {0}. 
+			((*f* f)(hypreal_of_real x + h) + 
+			 - hypreal_of_real (f x))/h @= hypreal_of_real D)"
+
+  differentiable :: [real=>real,real] => bool   (infixl 60)
+  "f differentiable x == (EX D. DERIV f x :> D)"
+
+  NSdifferentiable :: [real=>real,real] => bool   (infixl 60)
+  "f NSdifferentiable x == (EX D. NSDERIV f x :> D)"
+
+  increment :: [real=>real,real,hypreal] => hypreal
+  "increment f x h == (@inc. f NSdifferentiable x & 
+		       inc = (*f* f)(hypreal_of_real x + h) + -hypreal_of_real (f x))"
+
+  isUCont :: (real=>real) => bool
+  "isUCont f ==  (ALL r. #0 < r --> 
+		      (EX s. #0 < s & (ALL x y. abs(x + -y) < s
+			    --> abs(f x + -f y) < r)))"
+
+  isNSUCont :: (real=>real) => bool
+  "isNSUCont f == (ALL x y. x @= y --> (*f* f) x @= (*f* f) y)"
+
+
+(*Used in the proof of the Bolzano theorem*)
+consts
+  Bolzano_bisect :: "[real*real=>bool, real, real, nat] => (real*real)"
+
+primrec
+  "Bolzano_bisect P a b 0 = (a,b)"
+  "Bolzano_bisect P a b (Suc n) =
+      (let (x,y) = Bolzano_bisect P a b n
+       in if P(x, (x+y)/#2) then ((x+y)/#2, y)
+                            else (x, (x+y)/#2) )"
+  
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/NSA.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,2215 @@
+(*  Title       : NSA.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Infinite numbers, Infinitesimals,
+                  infinitely close relation  etc.
+*) 
+
+fun CLAIM_SIMP str thms =
+               prove_goal (the_context()) str
+               (fn prems => [cut_facts_tac prems 1,
+                   auto_tac (claset(),simpset() addsimps thms)]);
+fun CLAIM str = CLAIM_SIMP str [];
+
+(*--------------------------------------------------------------------
+     Closure laws for members of (embedded) set standard real SReal
+ --------------------------------------------------------------------*)
+
+Goalw [SReal_def] "[| (x::hypreal): SReal; y: SReal |] ==> x + y: SReal";
+by (Step_tac 1);
+by (res_inst_tac [("x","r + ra")] exI 1);
+by (Simp_tac 1);
+qed "SReal_add";
+
+Goalw [SReal_def] "[| (x::hypreal): SReal; y: SReal |] ==> x * y: SReal";
+by (Step_tac 1);
+by (res_inst_tac [("x","r * ra")] exI 1);
+by (simp_tac (simpset() addsimps [hypreal_of_real_mult]) 1);
+qed "SReal_mult";
+
+Goalw [SReal_def] "(x::hypreal): SReal ==> inverse x : SReal";
+by (blast_tac (claset() addIs [hypreal_of_real_inverse RS sym]) 1); 
+qed "SReal_inverse";
+
+Goal "[| (x::hypreal): SReal;  y: SReal |] ==> x/y: SReal";
+by (asm_simp_tac (simpset() addsimps [SReal_mult,SReal_inverse,
+                                      hypreal_divide_def]) 1); 
+qed "SReal_divide";
+
+Goalw [SReal_def] "(x::hypreal): SReal ==> -x : SReal";
+by (blast_tac (claset() addIs [hypreal_of_real_minus RS sym]) 1); 
+qed "SReal_minus";
+
+Goal "(-x : SReal) = ((x::hypreal): SReal)";
+by Auto_tac;  
+by (etac SReal_minus 2); 
+by (dtac SReal_minus 1); 
+by Auto_tac;  
+qed "SReal_minus_iff";
+Addsimps [SReal_minus_iff]; 
+
+Goal "[| (x::hypreal) + y : SReal; y: SReal |] ==> x: SReal";
+by (dres_inst_tac [("x","y")] SReal_minus 1);
+by (dtac SReal_add 1);
+by (assume_tac 1); 
+by Auto_tac;  
+qed "SReal_add_cancel";
+
+Goalw [SReal_def] "(x::hypreal): SReal ==> abs x : SReal";
+by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_hrabs]));
+qed "SReal_hrabs";
+
+Goalw [SReal_def] "hypreal_of_real x: SReal";
+by (Blast_tac 1);
+qed "SReal_hypreal_of_real";
+Addsimps [SReal_hypreal_of_real];
+
+Goalw [hypreal_number_of_def] "(number_of w ::hypreal) : SReal";
+by (rtac SReal_hypreal_of_real 1);
+qed "SReal_number_of";
+Addsimps [SReal_number_of];
+
+Goalw [hypreal_divide_def] "r : SReal ==> r/(number_of w::hypreal) : SReal";
+by (blast_tac (claset() addSIs [SReal_number_of, SReal_mult, 
+                                SReal_inverse]) 1);
+qed "SReal_divide_number_of";
+
+(* Infinitesimal ehr not in SReal *)
+
+Goalw [SReal_def] "ehr ~: SReal";
+by (auto_tac (claset(),
+              simpset() addsimps [hypreal_of_real_not_eq_epsilon RS not_sym]));
+qed "SReal_epsilon_not_mem";
+
+Goalw [SReal_def] "whr ~: SReal";
+by (auto_tac (claset(),
+              simpset() addsimps [hypreal_of_real_not_eq_omega RS not_sym]));
+qed "SReal_omega_not_mem";
+
+Goalw [SReal_def] "{x. hypreal_of_real x : SReal} = (UNIV::real set)";
+by Auto_tac;
+qed "SReal_UNIV_real";
+
+Goalw [SReal_def] "(x: SReal) = (EX y. x = hypreal_of_real y)";
+by Auto_tac;
+qed "SReal_iff";
+
+Goalw [SReal_def] "hypreal_of_real ``(UNIV::real set) = SReal";
+by Auto_tac;
+qed "hypreal_of_real_image";
+
+Goalw [SReal_def] "inv hypreal_of_real ``SReal = (UNIV::real set)";
+by Auto_tac;
+by (rtac (inj_hypreal_of_real RS inv_f_f RS subst) 1);
+by (Blast_tac 1);
+qed "inv_hypreal_of_real_image";
+
+Goalw [SReal_def] 
+      "[| EX x. x: P; P <= SReal |] ==> EX Q. P = hypreal_of_real `` Q";
+by (Best_tac 1); 
+qed "SReal_hypreal_of_real_image";
+
+Goal "[| (x::hypreal): SReal; y: SReal;  x<y |] ==> EX r: SReal. x<r & r<y";
+by (auto_tac (claset(), simpset() addsimps [SReal_iff]));
+by (dtac real_dense 1 THEN Step_tac 1);
+by (res_inst_tac [("x","hypreal_of_real r")] bexI 1);
+by Auto_tac;
+qed "SReal_dense";
+
+(*------------------------------------------------------------------
+                   Completeness of SReal
+ ------------------------------------------------------------------*)
+Goal "P <= SReal ==> ((EX x:P. y < x) = \ 
+\     (EX X. hypreal_of_real X : P & y < hypreal_of_real X))";
+by (blast_tac (claset() addSDs [SReal_iff RS iffD1]) 1);
+by (flexflex_tac );
+qed "SReal_sup_lemma";
+
+Goal "[| P <= SReal; EX x. x: P; EX y : SReal. ALL x: P. x < y |] \
+\     ==> (EX X. X: {w. hypreal_of_real w : P}) & \
+\         (EX Y. ALL X: {w. hypreal_of_real w : P}. X < Y)";
+by (rtac conjI 1);
+by (fast_tac (claset() addSDs [SReal_iff RS iffD1]) 1);
+by (Auto_tac THEN forward_tac [subsetD] 1 THEN assume_tac 1);
+by (dtac (SReal_iff RS iffD1) 1);
+by (Auto_tac THEN res_inst_tac [("x","ya")] exI 1);
+by Auto_tac;
+qed "SReal_sup_lemma2";
+
+(*------------------------------------------------------
+    lifting of ub and property of lub
+ -------------------------------------------------------*)
+Goalw [isUb_def,setle_def] 
+      "(isUb (SReal) (hypreal_of_real `` Q) (hypreal_of_real Y)) = \
+\      (isUb (UNIV :: real set) Q Y)";
+by Auto_tac;
+qed "hypreal_of_real_isUb_iff";
+
+Goalw [isLub_def,leastP_def] 
+     "isLub SReal (hypreal_of_real `` Q) (hypreal_of_real Y) \
+\     ==> isLub (UNIV :: real set) Q Y";
+by (auto_tac (claset() addIs [hypreal_of_real_isUb_iff RS iffD2],
+              simpset() addsimps [hypreal_of_real_isUb_iff, setge_def]));
+qed "hypreal_of_real_isLub1";
+
+Goalw [isLub_def,leastP_def] 
+      "isLub (UNIV :: real set) Q Y \
+\      ==> isLub SReal (hypreal_of_real `` Q) (hypreal_of_real Y)";
+by (auto_tac (claset(),
+              simpset() addsimps [hypreal_of_real_isUb_iff, setge_def]));
+by (forw_inst_tac [("x2","x")] (isUbD2a RS (SReal_iff RS iffD1) RS exE) 1);
+by (assume_tac 2);
+by (dres_inst_tac [("x","xa")] spec 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_isUb_iff]));
+qed "hypreal_of_real_isLub2";
+
+Goal "(isLub SReal (hypreal_of_real `` Q) (hypreal_of_real Y)) = \
+\     (isLub (UNIV :: real set) Q Y)";
+by (blast_tac (claset() addIs [hypreal_of_real_isLub1,
+                               hypreal_of_real_isLub2]) 1);
+qed "hypreal_of_real_isLub_iff";
+
+(* lemmas *)
+Goalw [isUb_def] 
+     "isUb SReal P Y ==> EX Yo. isUb SReal P (hypreal_of_real Yo)";
+by (auto_tac (claset(), simpset() addsimps [SReal_iff]));
+qed "lemma_isUb_hypreal_of_real";
+
+Goalw [isLub_def,leastP_def,isUb_def] 
+     "isLub SReal P Y ==> EX Yo. isLub SReal P (hypreal_of_real Yo)";
+by (auto_tac (claset(), simpset() addsimps [SReal_iff]));
+qed "lemma_isLub_hypreal_of_real";
+
+Goalw [isLub_def,leastP_def,isUb_def] 
+     "EX Yo. isLub SReal P (hypreal_of_real Yo) ==> EX Y. isLub SReal P Y";
+by Auto_tac;
+qed "lemma_isLub_hypreal_of_real2";
+
+Goal "[| P <= SReal;  EX x. x: P;  EX Y. isUb SReal P Y |] \
+\     ==> EX t::hypreal. isLub SReal P t";
+by (forward_tac [SReal_hypreal_of_real_image] 1);
+by (Auto_tac THEN dtac lemma_isUb_hypreal_of_real 1);
+by (auto_tac (claset() addSIs [reals_complete, lemma_isLub_hypreal_of_real2], 
+     simpset() addsimps [hypreal_of_real_isLub_iff,hypreal_of_real_isUb_iff]));
+qed "SReal_complete";
+
+(*--------------------------------------------------------------------
+        Set of finite elements is a subring of the extended reals
+ --------------------------------------------------------------------*)
+Goalw [HFinite_def] "[|x : HFinite; y : HFinite|] ==> (x+y) : HFinite";
+by (blast_tac (claset() addSIs [SReal_add,hrabs_add_less]) 1);
+qed "HFinite_add";
+
+Goalw [HFinite_def] "[|x : HFinite; y : HFinite|] ==> x*y : HFinite";
+by (Asm_full_simp_tac 1); 
+by (blast_tac (claset() addSIs [SReal_mult,hrabs_mult_less]) 1);
+qed "HFinite_mult";
+
+Goalw [HFinite_def] "(-x : HFinite) = (x : HFinite)";
+by (Simp_tac 1);
+qed "HFinite_minus_iff";
+
+Goalw [SReal_def,HFinite_def] "SReal <= HFinite";
+by Auto_tac;
+by (res_inst_tac [("x","#1 + abs(hypreal_of_real r)")] exI 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_hrabs]));
+by (res_inst_tac [("x","#1 + abs r")] exI 1);
+by (Simp_tac 1);
+qed "SReal_subset_HFinite";
+
+Goal "hypreal_of_real x : HFinite";
+by (auto_tac (claset() addIs [(SReal_subset_HFinite RS subsetD)],
+              simpset()));
+qed "HFinite_hypreal_of_real";
+
+Addsimps [HFinite_hypreal_of_real];
+
+Goalw [HFinite_def] "x : HFinite ==> EX t: SReal. abs x < t";
+by Auto_tac;
+qed "HFiniteD";
+
+Goalw [HFinite_def] "x : HFinite ==> abs x : HFinite";
+by (auto_tac (claset(), simpset() addsimps [hrabs_idempotent]));
+qed "HFinite_hrabs";
+
+Goalw [HFinite_def,Bex_def] "x ~: HFinite ==> abs x ~: HFinite";
+by (auto_tac (claset(), simpset() addsimps [hrabs_idempotent]));
+qed "not_HFinite_hrabs";
+
+Goal "number_of w : HFinite";
+by (rtac (SReal_number_of RS (SReal_subset_HFinite RS subsetD)) 1);
+qed "HFinite_number_of";
+Addsimps [HFinite_number_of];
+
+Goal "[|x : HFinite; y <= x; #0 <= y |] ==> y: HFinite";
+by (case_tac "x <= #0" 1);
+by (dres_inst_tac [("y","x")] order_trans 1);
+by (dtac hypreal_le_anti_sym 2);
+by (auto_tac (claset() addSDs [not_hypreal_leE], simpset()));
+by (auto_tac (claset() addSIs [bexI] addIs [order_le_less_trans],
+     simpset() addsimps [hrabs_eqI1,hrabs_eqI2,hrabs_minus_eqI1,HFinite_def]));
+qed "HFinite_bounded"; 
+
+(*------------------------------------------------------------------
+       Set of infinitesimals is a subring of the hyperreals   
+ ------------------------------------------------------------------*)
+Goalw [Infinitesimal_def]
+      "x : Infinitesimal ==> ALL r: SReal. #0 < r --> abs x < r";
+by Auto_tac;
+qed "InfinitesimalD";
+
+Goalw [Infinitesimal_def] "#0 : Infinitesimal";
+by (simp_tac (simpset() addsimps [hrabs_zero]) 1);
+qed "Infinitesimal_zero";
+AddIffs [Infinitesimal_zero];
+
+Goal "x/(#2::hypreal) + x/(#2::hypreal) = x";
+by Auto_tac;  
+qed "hypreal_sum_of_halves";
+
+Goal "#0 < r ==> #0 < r/(#2::hypreal)";
+by Auto_tac;  
+qed "hypreal_half_gt_zero";
+
+Goalw [Infinitesimal_def] 
+     "[| x : Infinitesimal; y : Infinitesimal |] ==> (x+y) : Infinitesimal";
+by Auto_tac;
+by (rtac (hypreal_sum_of_halves RS subst) 1);
+by (dtac hypreal_half_gt_zero 1);
+by (blast_tac (claset() addIs [hrabs_add_less, hrabs_add_less, 
+                               SReal_divide_number_of]) 1); 
+qed "Infinitesimal_add";
+
+Goalw [Infinitesimal_def] "(-x:Infinitesimal) = (x:Infinitesimal)";
+by (Full_simp_tac 1);
+qed "Infinitesimal_minus_iff";
+Addsimps [Infinitesimal_minus_iff];
+
+Goal "[| x : Infinitesimal;  y : Infinitesimal |] ==> x-y : Infinitesimal";
+by (asm_simp_tac
+    (simpset() addsimps [hypreal_diff_def, Infinitesimal_add]) 1);
+qed "Infinitesimal_diff";
+
+Goalw [Infinitesimal_def] 
+     "[| x : Infinitesimal; y : Infinitesimal |] ==> (x * y) : Infinitesimal";
+by Auto_tac;
+by (case_tac "y=#0" 1);
+by (cut_inst_tac [("u","abs x"),("v","#1"),("x","abs y"),("y","r")] 
+    hypreal_mult_less_mono 2);
+by Auto_tac;  
+qed "Infinitesimal_mult";
+
+Goal "[| x : Infinitesimal; y : HFinite |] ==> (x * y) : Infinitesimal";
+by (auto_tac (claset() addSDs [HFiniteD],
+              simpset() addsimps [Infinitesimal_def]));
+by (forward_tac [hrabs_less_gt_zero] 1);
+by (dres_inst_tac [("x","r/t")] bspec 1); 
+by (blast_tac (claset() addIs [SReal_divide]) 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_0_less_divide_iff]) 1); 
+by (case_tac "x=0 | y=0" 1);
+by (cut_inst_tac [("u","abs x"),("v","r/t"),("x","abs y")] 
+    hypreal_mult_less_mono 2);
+by (auto_tac (claset(), simpset() addsimps [hypreal_0_less_divide_iff]));  
+qed "Infinitesimal_HFinite_mult";
+
+Goal "[| x : Infinitesimal; y : HFinite |] ==> (y * x) : Infinitesimal";
+by (auto_tac (claset() addDs [Infinitesimal_HFinite_mult],
+              simpset() addsimps [hypreal_mult_commute]));
+qed "Infinitesimal_HFinite_mult2";
+
+(*** rather long proof ***)
+Goalw [HInfinite_def,Infinitesimal_def] 
+     "x: HInfinite ==> inverse x: Infinitesimal";
+by Auto_tac;
+by (eres_inst_tac [("x","inverse r")] ballE 1);
+by (rtac (hrabs_inverse RS ssubst) 1);
+by (forw_inst_tac [("x1","r"),("z","abs x")] 
+    (hypreal_inverse_gt_0 RS order_less_trans) 1);
+by (assume_tac 1);
+by (dtac ((hypreal_inverse_inverse RS sym) RS subst) 1);
+by (rtac (hypreal_inverse_less_iff RS iffD1) 1);
+by (auto_tac (claset(), simpset() addsimps [SReal_inverse]));  
+qed "HInfinite_inverse_Infinitesimal";
+
+
+
+Goalw [HInfinite_def] "[|x: HInfinite;y: HInfinite|] ==> (x*y): HInfinite";
+by Auto_tac;
+by (eres_inst_tac [("x","#1")] ballE 1);
+by (eres_inst_tac [("x","r")] ballE 1);
+by (case_tac "y=0" 1); 
+by (cut_inst_tac [("x","#1"),("y","abs x"),
+                  ("u","r"),("v","abs y")] hypreal_mult_less_mono 2);
+by (auto_tac (claset(), simpset() addsimps hypreal_mult_ac));  
+qed "HInfinite_mult";
+
+Goalw [HInfinite_def] 
+      "[|x: HInfinite; #0 <= y; #0 <= x|] ==> (x + y): HInfinite";
+by (auto_tac (claset() addSIs [hypreal_add_zero_less_le_mono],
+              simpset() addsimps [hrabs_eqI1, hypreal_add_commute,
+                                  hypreal_le_add_order]));
+qed "HInfinite_add_ge_zero";
+
+Goal "[|x: HInfinite; #0 <= y; #0 <= x|] ==> (y + x): HInfinite";
+by (auto_tac (claset() addSIs [HInfinite_add_ge_zero],
+              simpset() addsimps [hypreal_add_commute]));
+qed "HInfinite_add_ge_zero2";
+
+Goal "[|x: HInfinite; #0 < y; #0 < x|] ==> (x + y): HInfinite";
+by (blast_tac (claset() addIs [HInfinite_add_ge_zero,
+                    order_less_imp_le]) 1);
+qed "HInfinite_add_gt_zero";
+
+Goalw [HInfinite_def] "(-x: HInfinite) = (x: HInfinite)";
+by Auto_tac;
+qed "HInfinite_minus_iff";
+
+Goal "[|x: HInfinite; y <= #0; x <= #0|] ==> (x + y): HInfinite";
+by (dtac (HInfinite_minus_iff RS iffD2) 1);
+by (rtac (HInfinite_minus_iff RS iffD1) 1);
+by (auto_tac (claset() addIs [HInfinite_add_ge_zero],
+              simpset() addsimps [hypreal_minus_zero_le_iff]));
+qed "HInfinite_add_le_zero";
+
+Goal "[|x: HInfinite; y < #0; x < #0|] ==> (x + y): HInfinite";
+by (blast_tac (claset() addIs [HInfinite_add_le_zero,
+                               order_less_imp_le]) 1);
+qed "HInfinite_add_lt_zero";
+
+Goal "[|a: HFinite; b: HFinite; c: HFinite|] \ 
+\     ==> a*a + b*b + c*c : HFinite";
+by (auto_tac (claset() addIs [HFinite_mult,HFinite_add], simpset()));
+qed "HFinite_sum_squares";
+
+Goal "x ~: Infinitesimal ==> x ~= #0";
+by Auto_tac;
+qed "not_Infinitesimal_not_zero";
+
+Goal "x: HFinite - Infinitesimal ==> x ~= #0";
+by Auto_tac;
+qed "not_Infinitesimal_not_zero2";
+
+Goal "x : Infinitesimal ==> abs x : Infinitesimal";
+by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1);
+by Auto_tac;
+qed "Infinitesimal_hrabs";
+
+Goal "x ~: Infinitesimal ==> abs x ~: Infinitesimal";
+by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1);
+by Auto_tac;
+qed "not_Infinitesimal_hrabs";
+
+Goal "abs x : Infinitesimal ==> x : Infinitesimal";
+by (rtac ccontr 1);
+by (blast_tac (claset() addDs [not_Infinitesimal_hrabs]) 1);
+qed "hrabs_Infinitesimal_Infinitesimal";
+
+Goal "x : HFinite - Infinitesimal ==> abs x : HFinite - Infinitesimal";
+by (fast_tac (set_cs addSEs [HFinite_hrabs,not_Infinitesimal_hrabs]) 1);
+qed "HFinite_diff_Infinitesimal_hrabs";
+
+Goalw [Infinitesimal_def] 
+      "[| e : Infinitesimal; abs x < e |] ==> x : Infinitesimal";
+by (auto_tac (claset() addSDs [bspec], simpset()));
+by (dres_inst_tac [("x","e")] (hrabs_ge_self RS order_le_less_trans) 1);
+by (fast_tac (claset() addIs [order_less_trans]) 1);
+qed "hrabs_less_Infinitesimal";
+
+Goal "[| e : Infinitesimal; abs x <= e |] ==> x : Infinitesimal";
+by (blast_tac (claset() addDs [order_le_imp_less_or_eq] 
+                        addIs [hrabs_Infinitesimal_Infinitesimal,
+                               hrabs_less_Infinitesimal]) 1);
+qed "hrabs_le_Infinitesimal";
+
+Goalw [Infinitesimal_def] 
+      "[| e : Infinitesimal; \
+\         e' : Infinitesimal; \
+\         e' < x ; x < e |] ==> x : Infinitesimal";
+by (auto_tac (claset() addSDs [bspec], simpset()));
+by (dres_inst_tac [("x","e")] (hrabs_ge_self RS order_le_less_trans) 1);
+by (dtac (hrabs_interval_iff RS iffD1) 1);
+by (fast_tac(claset() addIs [order_less_trans,hrabs_interval_iff RS iffD2]) 1);
+qed "Infinitesimal_interval";
+
+Goal "[| e : Infinitesimal; e' : Infinitesimal; \
+\        e' <= x ; x <= e |] ==> x : Infinitesimal";
+by (auto_tac (claset() addIs [Infinitesimal_interval],
+    simpset() addsimps [hypreal_le_eq_less_or_eq]));
+qed "Infinitesimal_interval2";
+
+Goalw [Infinitesimal_def] 
+     "[| x ~: Infinitesimal;  y ~: Infinitesimal|] ==> (x*y) ~:Infinitesimal";
+by (Clarify_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [linorder_not_less]) 1); 
+by (eres_inst_tac [("x","r*ra")] ballE 1);
+by (fast_tac (claset() addIs [SReal_mult]) 2);
+by (auto_tac (claset(), simpset() addsimps [hypreal_0_less_mult_iff]));
+by (cut_inst_tac [("x","ra"),("y","abs y"),
+                  ("u","r"),("v","abs x")] hypreal_mult_le_mono 1);
+by Auto_tac;  
+qed "not_Infinitesimal_mult";
+
+Goal "x*y : Infinitesimal ==> x : Infinitesimal | y : Infinitesimal";
+by (rtac ccontr 1);
+by (dtac (de_Morgan_disj RS iffD1) 1);
+by (fast_tac (claset() addDs [not_Infinitesimal_mult]) 1);
+qed "Infinitesimal_mult_disj";
+
+Goal "x: HFinite-Infinitesimal ==> x ~= #0";
+by (Blast_tac 1);
+qed "HFinite_Infinitesimal_not_zero";
+
+Goal "[| x : HFinite - Infinitesimal; \
+\                  y : HFinite - Infinitesimal \
+\               |] ==> (x*y) : HFinite - Infinitesimal";
+by (Clarify_tac 1);
+by (blast_tac (claset() addDs [HFinite_mult,not_Infinitesimal_mult]) 1);
+qed "HFinite_Infinitesimal_diff_mult";
+
+Goalw [Infinitesimal_def,HFinite_def]  
+      "Infinitesimal <= HFinite";
+by Auto_tac;  
+by (res_inst_tac [("x","#1")] bexI 1); 
+by Auto_tac;  
+qed "Infinitesimal_subset_HFinite";
+
+Goal "x: Infinitesimal ==> x * hypreal_of_real r : Infinitesimal";
+by (etac (HFinite_hypreal_of_real RSN 
+          (2,Infinitesimal_HFinite_mult)) 1);
+qed "Infinitesimal_hypreal_of_real_mult";
+
+Goal "x: Infinitesimal ==> hypreal_of_real r * x: Infinitesimal";
+by (etac (HFinite_hypreal_of_real RSN 
+          (2,Infinitesimal_HFinite_mult2)) 1);
+qed "Infinitesimal_hypreal_of_real_mult2";
+
+(*----------------------------------------------------------------------
+                   Infinitely close relation @=
+ ----------------------------------------------------------------------*)
+
+Goalw [Infinitesimal_def,inf_close_def] 
+        "x:Infinitesimal = (x @= #0)";
+by (Simp_tac 1);
+qed "mem_infmal_iff";
+
+Goalw [inf_close_def]" (x @= y) = (x + -y @= #0)";
+by (Simp_tac 1);
+qed "inf_close_minus_iff";
+
+Goalw [inf_close_def]" (x @= y) = (-y + x @= #0)";
+by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
+qed "inf_close_minus_iff2";
+
+Goalw [inf_close_def,Infinitesimal_def]  "x @= x";
+by (Simp_tac 1);
+qed "inf_close_refl";
+AddIffs [inf_close_refl];
+
+Goalw [inf_close_def]  "x @= y ==> y @= x";
+by (rtac (hypreal_minus_distrib1 RS subst) 1);
+by (etac (Infinitesimal_minus_iff RS iffD2) 1);
+qed "inf_close_sym";
+
+Goalw [inf_close_def]  "[| x @= y; y @= z |] ==> x @= z";
+by (dtac Infinitesimal_add 1); 
+by (assume_tac 1); 
+by Auto_tac;  
+qed "inf_close_trans";
+
+Goal "[| r @= x; s @= x |] ==> r @= s";
+by (blast_tac (claset() addIs [inf_close_sym, inf_close_trans]) 1); 
+qed "inf_close_trans2";
+
+Goal "[| x @= r; x @= s|] ==> r @= s";
+by (blast_tac (claset() addIs [inf_close_sym, inf_close_trans]) 1); 
+qed "inf_close_trans3";
+
+Goal "(number_of w @= x) = (x @= number_of w)";
+by (blast_tac (claset() addIs [inf_close_sym]) 1); 
+qed "number_of_inf_close_reorient";
+Addsimps [number_of_inf_close_reorient];
+
+Goal "(x-y : Infinitesimal) = (x @= y)";
+by (auto_tac (claset(),
+              simpset() addsimps [hypreal_diff_def, inf_close_minus_iff RS sym,
+                                  mem_infmal_iff]));
+qed "Infinitesimal_inf_close_minus";
+
+Goalw [monad_def] "(x @= y) = (monad(x)=monad(y))";
+by (auto_tac (claset() addDs [inf_close_sym] 
+                       addSEs [inf_close_trans,equalityCE],
+              simpset()));
+qed "inf_close_monad_iff";
+
+Goal "[| x: Infinitesimal; y: Infinitesimal |] ==> x @= y";
+by (asm_full_simp_tac (simpset() addsimps [mem_infmal_iff]) 1);
+by (blast_tac (claset() addIs [inf_close_trans, inf_close_sym]) 1); 
+qed "Infinitesimal_inf_close";
+
+val prem1::prem2::rest = 
+goalw thy [inf_close_def] "[| a @= b; c @= d |] ==> a+c @= b+d";
+by (rtac (hypreal_minus_add_distrib RS ssubst) 1);
+by (rtac (hypreal_add_assoc RS ssubst) 1);
+by (res_inst_tac [("y1","c")] (hypreal_add_left_commute RS subst) 1);
+by (rtac (hypreal_add_assoc RS subst) 1);
+by (rtac ([prem1,prem2] MRS Infinitesimal_add) 1);
+qed "inf_close_add";
+
+Goal "a @= b ==> -a @= -b";
+by (rtac ((inf_close_minus_iff RS iffD2) RS inf_close_sym) 1);
+by (dtac (inf_close_minus_iff RS iffD1) 1);
+by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
+qed "inf_close_minus";
+
+Goal "-a @= -b ==> a @= b";
+by (auto_tac (claset() addDs [inf_close_minus], simpset()));
+qed "inf_close_minus2";
+
+Goal "(-a @= -b) = (a @= b)";
+by (blast_tac (claset() addIs [inf_close_minus,inf_close_minus2]) 1);
+qed "inf_close_minus_cancel";
+
+Addsimps [inf_close_minus_cancel];
+
+Goal "[| a @= b; c @= d |] ==> a + -c @= b + -d";
+by (blast_tac (claset() addSIs [inf_close_add,inf_close_minus]) 1);
+qed "inf_close_add_minus";
+
+Goalw [inf_close_def] "[| a @= b; c: HFinite|] ==> a*c @= b*c"; 
+by (asm_full_simp_tac (simpset() addsimps [Infinitesimal_HFinite_mult,
+    hypreal_minus_mult_eq1,hypreal_add_mult_distrib RS sym] 
+    delsimps [hypreal_minus_mult_eq1 RS sym]) 1);
+qed "inf_close_mult1";
+
+Goal "[|a @= b; c: HFinite|] ==> c*a @= c*b"; 
+by (asm_simp_tac (simpset() addsimps [inf_close_mult1,hypreal_mult_commute]) 1);
+qed "inf_close_mult2";
+
+Goal "[|u @= v*x; x @= y; v: HFinite|] ==> u @= v*y";
+by (fast_tac (claset() addIs [inf_close_mult2,inf_close_trans]) 1);
+qed "inf_close_mult_subst";
+
+Goal "[| u @= x*v; x @= y; v: HFinite |] ==> u @= y*v";
+by (fast_tac (claset() addIs [inf_close_mult1,inf_close_trans]) 1);
+qed "inf_close_mult_subst2";
+
+Goal "[| u @= x*hypreal_of_real v; x @= y |] ==> u @= y*hypreal_of_real v";
+by (auto_tac (claset() addIs [inf_close_mult_subst2], simpset()));
+qed "inf_close_mult_subst_SReal";
+
+Goalw [inf_close_def]  "a = b ==> a @= b";
+by (Asm_simp_tac 1);
+qed "inf_close_eq_imp";
+
+Goal "x: Infinitesimal ==> -x @= x"; 
+by (fast_tac (HOL_cs addIs [Infinitesimal_minus_iff RS iffD2,
+    mem_infmal_iff RS iffD1,inf_close_trans2]) 1);
+qed "Infinitesimal_minus_inf_close";
+
+Goalw [inf_close_def]  "(EX y: Infinitesimal. x + -z = y) = (x @= z)";
+by (Blast_tac 1);
+qed "bex_Infinitesimal_iff";
+
+Goal "(EX y: Infinitesimal. x = z + y) = (x @= z)";
+by (asm_full_simp_tac (simpset() addsimps [bex_Infinitesimal_iff RS sym]) 1);
+by (Force_tac 1);
+qed "bex_Infinitesimal_iff2";
+
+Goal "[| y: Infinitesimal; x + y = z |] ==> x @= z";
+by (rtac (bex_Infinitesimal_iff RS iffD1) 1);
+by (dtac (Infinitesimal_minus_iff RS iffD2) 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_minus_add_distrib,
+    hypreal_add_assoc RS sym]));
+qed "Infinitesimal_add_inf_close";
+
+Goal "y: Infinitesimal ==> x @= x + y";
+by (rtac (bex_Infinitesimal_iff RS iffD1) 1);
+by (dtac (Infinitesimal_minus_iff RS iffD2) 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_minus_add_distrib,
+    hypreal_add_assoc RS sym]));
+qed "Infinitesimal_add_inf_close_self";
+
+Goal "y: Infinitesimal ==> x @= y + x";
+by (auto_tac (claset() addDs [Infinitesimal_add_inf_close_self],
+    simpset() addsimps [hypreal_add_commute]));
+qed "Infinitesimal_add_inf_close_self2";
+
+Goal "y: Infinitesimal ==> x @= x + -y";
+by (blast_tac (claset() addSIs [Infinitesimal_add_inf_close_self,
+                                Infinitesimal_minus_iff RS iffD2]) 1);
+qed "Infinitesimal_add_minus_inf_close_self";
+
+Goal "[| y: Infinitesimal; x+y @= z|] ==> x @= z";
+by (dres_inst_tac [("x","x")] (Infinitesimal_add_inf_close_self RS inf_close_sym) 1);
+by (etac (inf_close_trans3 RS inf_close_sym) 1);
+by (assume_tac 1);
+qed "Infinitesimal_add_cancel";
+
+Goal "[| y: Infinitesimal; x @= z + y|] ==> x @= z";
+by (dres_inst_tac [("x","z")] (Infinitesimal_add_inf_close_self2  RS inf_close_sym) 1);
+by (etac (inf_close_trans3 RS inf_close_sym) 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
+by (etac inf_close_sym 1);
+qed "Infinitesimal_add_right_cancel";
+
+Goal "d + b  @= d + c ==> b @= c";
+by (dtac (inf_close_minus_iff RS iffD1) 1);
+by (asm_full_simp_tac (simpset() addsimps 
+    [hypreal_minus_add_distrib,inf_close_minus_iff RS sym] 
+    @ hypreal_add_ac) 1);
+qed "inf_close_add_left_cancel";
+
+Goal "b + d @= c + d ==> b @= c";
+by (rtac inf_close_add_left_cancel 1);
+by (asm_full_simp_tac (simpset() addsimps 
+    [hypreal_add_commute]) 1);
+qed "inf_close_add_right_cancel";
+
+Goal "b @= c ==> d + b @= d + c";
+by (rtac (inf_close_minus_iff RS iffD2) 1);
+by (asm_full_simp_tac (simpset() addsimps 
+    [hypreal_minus_add_distrib,inf_close_minus_iff RS sym] 
+    @ hypreal_add_ac) 1);
+qed "inf_close_add_mono1";
+
+Goal "b @= c ==> b + a @= c + a";
+by (asm_simp_tac (simpset() addsimps 
+    [hypreal_add_commute,inf_close_add_mono1]) 1);
+qed "inf_close_add_mono2";
+
+Goal "(a + b @= a + c) = (b @= c)";
+by (fast_tac (claset() addEs [inf_close_add_left_cancel,
+    inf_close_add_mono1]) 1);
+qed "inf_close_add_left_iff";
+
+Addsimps [inf_close_add_left_iff];
+
+Goal "(b + a @= c + a) = (b @= c)";
+by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
+qed "inf_close_add_right_iff";
+
+Addsimps [inf_close_add_right_iff];
+
+Goal "[| x: HFinite; x @= y |] ==> y: HFinite";
+by (dtac (bex_Infinitesimal_iff2 RS iffD2) 1);
+by (Step_tac 1);
+by (dtac (Infinitesimal_subset_HFinite RS subsetD 
+          RS (HFinite_minus_iff RS iffD2)) 1);
+by (dtac HFinite_add 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc]));
+qed "inf_close_HFinite";
+
+Goal "x @= hypreal_of_real D ==> x: HFinite";
+by (rtac (inf_close_sym RSN (2,inf_close_HFinite)) 1);
+by Auto_tac;
+qed "inf_close_hypreal_of_real_HFinite";
+
+Goal "[|a @= b; c @= d; b: HFinite; d: HFinite|] ==> a*c @= b*d";
+by (rtac inf_close_trans 1); 
+by (rtac inf_close_mult2 2); 
+by (rtac inf_close_mult1 1);
+by (blast_tac (claset() addIs [inf_close_HFinite, inf_close_sym]) 2);  
+by Auto_tac;  
+qed "inf_close_mult_HFinite";
+
+Goal "[|a @= hypreal_of_real b; c @= hypreal_of_real d |] \
+\     ==> a*c @= hypreal_of_real b*hypreal_of_real d";
+by (blast_tac (claset() addSIs [inf_close_mult_HFinite,
+            inf_close_hypreal_of_real_HFinite,HFinite_hypreal_of_real]) 1);
+qed "inf_close_mult_hypreal_of_real";
+
+Goal "[| a: SReal; a ~= #0; a*x @= #0 |] ==> x @= #0";
+by (dtac (SReal_inverse RS (SReal_subset_HFinite RS subsetD)) 1);
+by (auto_tac (claset() addDs [inf_close_mult2],
+    simpset() addsimps [hypreal_mult_assoc RS sym]));
+qed "inf_close_SReal_mult_cancel_zero";
+
+(* REM comments: newly added *)
+Goal "[| a: SReal; x @= #0 |] ==> x*a @= #0";
+by (auto_tac (claset() addDs [(SReal_subset_HFinite RS subsetD),
+              inf_close_mult1], simpset()));
+qed "inf_close_mult_SReal1";
+
+Goal "[| a: SReal; x @= #0 |] ==> a*x @= #0";
+by (auto_tac (claset() addDs [(SReal_subset_HFinite RS subsetD),
+              inf_close_mult2], simpset()));
+qed "inf_close_mult_SReal2";
+
+Goal "[|a : SReal; a ~= #0 |] ==> (a*x @= #0) = (x @= #0)";
+by (blast_tac (claset() addIs [inf_close_SReal_mult_cancel_zero,
+    inf_close_mult_SReal2]) 1);
+qed "inf_close_mult_SReal_zero_cancel_iff";
+Addsimps [inf_close_mult_SReal_zero_cancel_iff];
+
+Goal "[| a: SReal; a ~= #0; a* w @= a*z |] ==> w @= z";
+by (dtac (SReal_inverse RS (SReal_subset_HFinite RS subsetD)) 1);
+by (auto_tac (claset() addDs [inf_close_mult2],
+    simpset() addsimps [hypreal_mult_assoc RS sym]));
+qed "inf_close_SReal_mult_cancel";
+
+Goal "[| a: SReal; a ~= #0|] ==> (a* w @= a*z) = (w @= z)";
+by (auto_tac (claset() addSIs [inf_close_mult2,SReal_subset_HFinite RS subsetD] 
+    addIs [inf_close_SReal_mult_cancel], simpset()));
+qed "inf_close_SReal_mult_cancel_iff1";
+Addsimps [inf_close_SReal_mult_cancel_iff1];
+
+Goal "[| z <= f; f @= g; g <= z |] ==> f @= z";
+by (asm_full_simp_tac (simpset() addsimps [bex_Infinitesimal_iff2 RS sym]) 1); 
+by Auto_tac;  
+by (res_inst_tac [("x","g+y-z")] bexI 1);
+by (Simp_tac 1); 
+by (rtac Infinitesimal_interval2 1); 
+by (rtac Infinitesimal_zero 2); 
+by Auto_tac;  
+qed "inf_close_le_bound";
+
+(*-----------------------------------------------------------------
+    Zero is the only infinitesimal that is also a real
+ -----------------------------------------------------------------*)
+
+Goalw [Infinitesimal_def] 
+     "[| x: SReal; y: Infinitesimal; #0 < x |] ==> y < x";
+by (rtac (hrabs_ge_self RS order_le_less_trans) 1);
+by Auto_tac;  
+qed "Infinitesimal_less_SReal";
+
+Goal "y: Infinitesimal ==> ALL r: SReal. #0 < r --> y < r";
+by (blast_tac (claset() addIs [Infinitesimal_less_SReal]) 1);
+qed "Infinitesimal_less_SReal2";
+
+Goalw [Infinitesimal_def] 
+     "[| #0 < y;  y: SReal|] ==> y ~: Infinitesimal";
+by (auto_tac (claset(), simpset() addsimps [hrabs_def]));
+qed "SReal_not_Infinitesimal";
+
+Goal "[| y < #0;  y : SReal |] ==> y ~: Infinitesimal";
+by (stac (Infinitesimal_minus_iff RS sym) 1); 
+by (rtac SReal_not_Infinitesimal 1); 
+by Auto_tac;  
+qed "SReal_minus_not_Infinitesimal";
+
+Goal "SReal Int Infinitesimal = {#0}";
+by Auto_tac;
+by (cut_inst_tac [("x","x"),("y","#0")] hypreal_linear 1);
+by (blast_tac (claset() addDs [SReal_not_Infinitesimal,
+                               SReal_minus_not_Infinitesimal]) 1);
+qed "SReal_Int_Infinitesimal_zero";
+
+Goal "[| x: SReal; x: Infinitesimal|] ==> x = #0";
+by (cut_facts_tac [SReal_Int_Infinitesimal_zero] 1);
+by (Blast_tac 1);
+qed "SReal_Infinitesimal_zero";
+
+Goal "[| x : SReal; x ~= #0 |] ==> x : HFinite - Infinitesimal";
+by (auto_tac (claset() addDs [SReal_Infinitesimal_zero,
+                              SReal_subset_HFinite RS subsetD], 
+              simpset()));
+qed "SReal_HFinite_diff_Infinitesimal";
+
+Goal "hypreal_of_real x ~= #0 ==> hypreal_of_real x : HFinite - Infinitesimal";
+by (rtac SReal_HFinite_diff_Infinitesimal 1);
+by Auto_tac;
+qed "hypreal_of_real_HFinite_diff_Infinitesimal";
+
+Goal "(hypreal_of_real x : Infinitesimal) = (x=#0)";
+by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_zero]));  
+by (rtac ccontr 1); 
+by (rtac (hypreal_of_real_HFinite_diff_Infinitesimal RS DiffD2) 1); 
+by Auto_tac;  
+qed "hypreal_of_real_Infinitesimal_iff_0";
+AddIffs [hypreal_of_real_Infinitesimal_iff_0];
+
+Goal "number_of w ~= (#0::hypreal) ==> number_of w ~: Infinitesimal";
+by (fast_tac (claset() addDs [SReal_number_of RS SReal_Infinitesimal_zero]) 1);
+qed "number_of_not_Infinitesimal";
+Addsimps [number_of_not_Infinitesimal];
+
+Goal "[| y: SReal; x @= y; y~= #0 |] ==> x ~= #0";
+by (cut_inst_tac [("x","y")] hypreal_trichotomy 1);
+by (Asm_full_simp_tac 1); 
+by (blast_tac (claset() addDs 
+		[inf_close_sym RS (mem_infmal_iff RS iffD2),
+		 SReal_not_Infinitesimal, SReal_minus_not_Infinitesimal]) 1);
+qed "inf_close_SReal_not_zero";
+
+Goal "[| x @= y; y : HFinite - Infinitesimal |] \
+\     ==> x : HFinite - Infinitesimal";
+by (auto_tac (claset() addIs [inf_close_sym RSN (2,inf_close_HFinite)],
+              simpset() addsimps [mem_infmal_iff]));
+by (dtac inf_close_trans3 1 THEN assume_tac 1);
+by (blast_tac (claset() addDs [inf_close_sym]) 1);
+qed "HFinite_diff_Infinitesimal_inf_close";
+
+(*The premise y~=0 is essential; otherwise x/y =0 and we lose the 
+  HFinite premise.*)
+Goal "[| y ~= #0;  y: Infinitesimal;  x/y : HFinite |] ==> x : Infinitesimal";
+by (dtac Infinitesimal_HFinite_mult2 1);
+by (assume_tac 1);
+by (asm_full_simp_tac 
+    (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 1);
+qed "Infinitesimal_ratio";
+
+(*------------------------------------------------------------------
+       Standard Part Theorem: Every finite x: R* is infinitely 
+       close to a unique real number (i.e a member of SReal)
+ ------------------------------------------------------------------*)
+(*------------------------------------------------------------------
+         Uniqueness: Two infinitely close reals are equal
+ ------------------------------------------------------------------*)
+
+Goal "[|x: SReal; y: SReal|] ==> (x @= y) = (x = y)"; 
+by Auto_tac;
+by (rewrite_goals_tac [inf_close_def]);
+by (dres_inst_tac [("x","y")] SReal_minus 1);
+by (dtac SReal_add 1 THEN assume_tac 1);
+by (dtac SReal_Infinitesimal_zero 1 THEN assume_tac 1);
+by (dtac sym 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_eq_minus_iff RS sym]) 1);
+qed "SReal_inf_close_iff";
+
+Goal "(number_of v @= number_of w) = (number_of v = (number_of w :: hypreal))";
+by (rtac SReal_inf_close_iff 1); 
+by Auto_tac;  
+qed "number_of_inf_close_iff";
+Addsimps [number_of_inf_close_iff];
+
+Goal "(hypreal_of_real k @= hypreal_of_real m) = (k = m)";
+by Auto_tac;  
+by (rtac (inj_hypreal_of_real RS injD) 1); 
+by (rtac (SReal_inf_close_iff RS iffD1) 1); 
+by Auto_tac;  
+qed "hypreal_of_real_inf_close_iff";
+Addsimps [hypreal_of_real_inf_close_iff];
+
+Goal "(hypreal_of_real k @= number_of w) = (k = number_of w)";
+by (stac (hypreal_of_real_inf_close_iff RS sym) 1); 
+by Auto_tac;  
+qed "hypreal_of_real_inf_close_number_of_iff";
+Addsimps [hypreal_of_real_inf_close_number_of_iff];
+
+Goal "[| r: SReal; s: SReal; r @= x; s @= x|] ==> r = s";
+by (blast_tac (claset() addIs [(SReal_inf_close_iff RS iffD1),
+               inf_close_trans2]) 1);
+qed "inf_close_unique_real";
+
+(*------------------------------------------------------------------
+       Existence of unique real infinitely close
+ ------------------------------------------------------------------*)
+(* lemma about lubs *)
+Goal "!!(x::hypreal). [| isLub R S x; isLub R S y |] ==> x = y";
+by (forward_tac [isLub_isUb] 1);
+by (forw_inst_tac [("x","y")] isLub_isUb 1);
+by (blast_tac (claset() addSIs [hypreal_le_anti_sym]
+                addSDs [isLub_le_isUb]) 1);
+qed "hypreal_isLub_unique";
+
+Goal "x: HFinite ==> EX u. isUb SReal {s. s: SReal & s < x} u";
+by (dtac HFiniteD 1 THEN Step_tac 1);
+by (rtac exI 1 THEN rtac isUbI 1 THEN assume_tac 2);
+by (auto_tac (claset() addIs [order_less_imp_le,setleI,isUbI,
+    order_less_trans], simpset() addsimps [hrabs_interval_iff]));
+qed "lemma_st_part_ub";
+
+Goal "x: HFinite ==> EX y. y: {s. s: SReal & s < x}";
+by (dtac HFiniteD 1 THEN Step_tac 1);
+by (dtac SReal_minus 1);
+by (res_inst_tac [("x","-t")] exI 1); 
+by (auto_tac (claset(), simpset() addsimps [hrabs_interval_iff]));
+qed "lemma_st_part_nonempty";
+
+Goal "{s. s: SReal & s < x} <= SReal";
+by Auto_tac;
+qed "lemma_st_part_subset";
+
+Goal "x: HFinite ==> EX t. isLub SReal {s. s: SReal & s < x} t";
+by (blast_tac (claset() addSIs [SReal_complete,lemma_st_part_ub,
+    lemma_st_part_nonempty, lemma_st_part_subset]) 1);
+qed "lemma_st_part_lub";
+
+Goal "((t::hypreal) + r <= t) = (r <= #0)";
+by (Step_tac 1);
+by (dres_inst_tac [("x","-t")] hypreal_add_left_le_mono1 1);
+by (dres_inst_tac [("x","t")] hypreal_add_left_le_mono1 2);
+by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc RS sym]));
+qed "lemma_hypreal_le_left_cancel";
+
+Goal "[| x: HFinite;  isLub SReal {s. s: SReal & s < x} t; \
+\        r: SReal;  #0 < r |] ==> x <= t + r";
+by (forward_tac [isLubD1a] 1);
+by (rtac ccontr 1 THEN dtac (linorder_not_le RS iffD2) 1);
+by (dres_inst_tac [("x","t")] SReal_add 1 THEN assume_tac 1);
+by (dres_inst_tac [("y","t + r")] (isLubD1 RS setleD) 1);
+by Auto_tac;  
+qed "lemma_st_part_le1";
+
+Goalw [setle_def] "!!x::hypreal. [| S *<= x; x < y |] ==> S *<= y";
+by (auto_tac (claset() addSDs [bspec,order_le_less_trans]
+                       addIs [order_less_imp_le],
+              simpset()));
+qed "hypreal_setle_less_trans";
+
+Goalw [isUb_def] 
+     "!!x::hypreal. [| isUb R S x; x < y; y: R |] ==> isUb R S y";
+by (blast_tac (claset() addIs [hypreal_setle_less_trans]) 1);
+qed "hypreal_gt_isUb";
+
+Goal "[| x: HFinite; x < y; y: SReal |] \
+\              ==> isUb SReal {s. s: SReal & s < x} y";
+by (auto_tac (claset() addDs [order_less_trans]
+    addIs [order_less_imp_le] addSIs [isUbI,setleI], simpset()));
+qed "lemma_st_part_gt_ub";
+
+Goal "t <= t + -r ==> r <= (#0::hypreal)";
+by (dres_inst_tac [("x","-t")] hypreal_add_left_le_mono1 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc RS sym]));
+qed "lemma_minus_le_zero";
+
+Goal "[| x: HFinite; \
+\        isLub SReal {s. s: SReal & s < x} t; \
+\        r: SReal; #0 < r |] \
+\     ==> t + -r <= x";
+by (forward_tac [isLubD1a] 1);
+by (rtac ccontr 1 THEN dtac not_hypreal_leE 1);
+by (dtac SReal_minus 1 THEN dres_inst_tac [("x","t")] 
+    SReal_add 1 THEN assume_tac 1);
+by (dtac lemma_st_part_gt_ub 1 THEN REPEAT(assume_tac 1));
+by (dtac isLub_le_isUb 1 THEN assume_tac 1);
+by (dtac lemma_minus_le_zero 1);
+by (auto_tac (claset() addDs [order_less_le_trans],  simpset()));
+qed "lemma_st_part_le2";
+
+Goal "((x::hypreal) <= t + r) = (x + -t <= r)";
+by Auto_tac;  
+qed "lemma_hypreal_le_swap";
+
+Goal "[| x: HFinite; \
+\        isLub SReal {s. s: SReal & s < x} t; \
+\        r: SReal; #0 < r |] \
+\     ==> x + -t <= r";
+by (blast_tac (claset() addSIs [lemma_hypreal_le_swap RS iffD1,
+                                lemma_st_part_le1]) 1);
+qed "lemma_st_part1a";
+
+Goal "(t + -r <= x) = (-(x + -t) <= (r::hypreal))";
+by Auto_tac;  
+qed "lemma_hypreal_le_swap2";
+
+Goal "[| x: HFinite; \
+\        isLub SReal {s. s: SReal & s < x} t; \
+\        r: SReal;  #0 < r |] \
+\     ==> -(x + -t) <= r";
+by (blast_tac (claset() addSIs [lemma_hypreal_le_swap2 RS iffD1,
+                                lemma_st_part_le2]) 1);
+qed "lemma_st_part2a";
+
+Goal "(x::hypreal): SReal ==> isUb SReal {s. s: SReal & s < x} x";
+by (auto_tac (claset() addIs [isUbI, setleI,order_less_imp_le], simpset()));
+qed "lemma_SReal_ub";
+
+Goal "(x::hypreal): SReal ==> isLub SReal {s. s: SReal & s < x} x";
+by (auto_tac (claset() addSIs [isLubI2,lemma_SReal_ub,setgeI], simpset()));
+by (forward_tac [isUbD2a] 1);
+by (res_inst_tac [("x","x"),("y","y")] hypreal_linear_less2 1);
+by (auto_tac (claset() addSIs [order_less_imp_le], simpset()));
+by (EVERY1[dtac SReal_dense, assume_tac, assume_tac, Step_tac]);
+by (dres_inst_tac [("y","r")] isUbD 1);
+by (auto_tac (claset() addDs [order_less_le_trans], simpset()));
+qed "lemma_SReal_lub";
+
+Goal "[| x: HFinite; \
+\        isLub SReal {s. s: SReal & s < x} t; \
+\        r: SReal; #0 < r |] \
+\     ==> x + -t ~= r";
+by Auto_tac;
+by (forward_tac [isLubD1a RS SReal_minus] 1);
+by (dtac SReal_add_cancel 1 THEN assume_tac 1);
+by (dres_inst_tac [("x","x")] lemma_SReal_lub 1);
+by (dtac hypreal_isLub_unique 1 THEN assume_tac 1);
+by Auto_tac;
+qed "lemma_st_part_not_eq1";
+
+Goal "[| x: HFinite; \
+\        isLub SReal {s. s: SReal & s < x} t; \
+\        r: SReal; #0 < r |] \
+\     ==> -(x + -t) ~= r";
+by (auto_tac (claset(), simpset() addsimps [hypreal_minus_add_distrib]));
+by (forward_tac [isLubD1a] 1);
+by (dtac SReal_add_cancel 1 THEN assume_tac 1);
+by (dres_inst_tac [("x","-x")] SReal_minus 1);
+by (Asm_full_simp_tac 1);
+by (dres_inst_tac [("x","x")] lemma_SReal_lub 1);
+by (dtac hypreal_isLub_unique 1 THEN assume_tac 1);
+by Auto_tac;
+qed "lemma_st_part_not_eq2";
+
+Goal "[| x: HFinite; \
+\        isLub SReal {s. s: SReal & s < x} t; \
+\        r: SReal; #0 < r |] \
+\     ==> abs (x + -t) < r";
+by (forward_tac [lemma_st_part1a] 1);
+by (forward_tac [lemma_st_part2a] 4);
+by Auto_tac;
+by (REPEAT(dtac order_le_imp_less_or_eq 1));
+by (auto_tac (claset() addDs [lemma_st_part_not_eq1,
+    lemma_st_part_not_eq2], simpset() addsimps [hrabs_interval_iff2]));
+qed "lemma_st_part_major";
+
+Goal "[| x: HFinite; \
+\        isLub SReal {s. s: SReal & s < x} t |] \
+\     ==> ALL r: SReal. #0 < r --> abs (x + -t) < r";
+by (blast_tac (claset() addSDs [lemma_st_part_major]) 1);
+qed "lemma_st_part_major2";
+
+(*----------------------------------------------
+  Existence of real and Standard Part Theorem
+ ----------------------------------------------*)
+Goal "x: HFinite ==> \
+\     EX t: SReal. ALL r: SReal. #0 < r --> abs (x + -t) < r";
+by (forward_tac [lemma_st_part_lub] 1 THEN Step_tac 1);
+by (forward_tac [isLubD1a] 1);
+by (blast_tac (claset() addDs [lemma_st_part_major2]) 1);
+qed "lemma_st_part_Ex";
+
+Goalw [inf_close_def,Infinitesimal_def] 
+     "x:HFinite ==> EX t: SReal. x @= t";
+by (dtac lemma_st_part_Ex 1);
+by Auto_tac;  
+qed "st_part_Ex";
+
+(*--------------------------------
+  Unique real infinitely close
+ -------------------------------*)
+Goal "x:HFinite ==> EX! t. t: SReal & x @= t";
+by (dtac st_part_Ex 1 THEN Step_tac 1);
+by (dtac inf_close_sym 2 THEN dtac inf_close_sym 2 
+    THEN dtac inf_close_sym 2);
+by (auto_tac (claset() addSIs [inf_close_unique_real], simpset()));
+qed "st_part_Ex1";
+
+(*------------------------------------------------------------------
+       Finite and Infinite --- more theorems
+ ------------------------------------------------------------------*)
+
+Goalw [HFinite_def,HInfinite_def] "HFinite Int HInfinite = {}";
+by (auto_tac (claset() addIs [hypreal_less_irrefl] 
+              addDs [order_less_trans,bspec],
+              simpset()));
+qed "HFinite_Int_HInfinite_empty";
+Addsimps [HFinite_Int_HInfinite_empty];
+
+Goal "x: HFinite ==> x ~: HInfinite";
+by (EVERY1[Step_tac, dtac IntI, assume_tac]);
+by Auto_tac;
+qed "HFinite_not_HInfinite";
+
+val [prem] = goalw thy [HInfinite_def] "x~: HFinite ==> x: HInfinite";
+by (cut_facts_tac [prem] 1);
+by (Clarify_tac 1);
+by (auto_tac (claset() addSDs [spec],
+              simpset() addsimps [HFinite_def,Bex_def]));
+by (dtac hypreal_leI 1);
+by (dtac order_le_imp_less_or_eq 1);
+by (auto_tac (claset() addSDs [SReal_subset_HFinite RS subsetD],
+    simpset() addsimps [prem,hrabs_idempotent,prem RS not_HFinite_hrabs]));
+qed "not_HFinite_HInfinite";
+
+Goal "x : HInfinite | x : HFinite";
+by (blast_tac (claset() addIs [not_HFinite_HInfinite]) 1);
+qed "HInfinite_HFinite_disj";
+
+Goal "(x : HInfinite) = (x ~: HFinite)";
+by (blast_tac (claset() addDs [HFinite_not_HInfinite,
+               not_HFinite_HInfinite]) 1);
+qed "HInfinite_HFinite_iff";
+
+Goal "(x : HFinite) = (x ~: HInfinite)";
+by (simp_tac (simpset() addsimps [HInfinite_HFinite_iff]) 1);
+qed "HFinite_HInfinite_iff";
+
+(*------------------------------------------------------------------
+       Finite, Infinite and Infinitesimal --- more theorems
+ ------------------------------------------------------------------*)
+
+Goal "x ~: Infinitesimal ==> x : HInfinite | x : HFinite - Infinitesimal";
+by (fast_tac (claset() addIs [not_HFinite_HInfinite]) 1);
+qed "HInfinite_diff_HFinite_Infinitesimal_disj";
+
+Goal "[| x : HFinite; x ~: Infinitesimal |] ==> inverse x : HFinite";
+by (cut_inst_tac [("x","inverse x")] HInfinite_HFinite_disj 1);
+by (auto_tac (claset() addSDs [HInfinite_inverse_Infinitesimal], simpset()));
+qed "HFinite_inverse";
+
+Goal "x : HFinite - Infinitesimal ==> inverse x : HFinite";
+by (blast_tac (claset() addIs [HFinite_inverse]) 1);
+qed "HFinite_inverse2";
+
+(* stronger statement possible in fact *)
+Goal "x ~: Infinitesimal ==> inverse(x) : HFinite";
+by (dtac HInfinite_diff_HFinite_Infinitesimal_disj 1);
+by (blast_tac (claset() addIs [HFinite_inverse,
+                 HInfinite_inverse_Infinitesimal,
+                 Infinitesimal_subset_HFinite RS subsetD]) 1);
+qed "Infinitesimal_inverse_HFinite";
+
+Goal "x : HFinite - Infinitesimal ==> inverse x : HFinite - Infinitesimal";
+by (auto_tac (claset() addIs [Infinitesimal_inverse_HFinite], simpset()));
+by (dtac Infinitesimal_HFinite_mult2 1);
+by (assume_tac 1);
+by (asm_full_simp_tac
+   (simpset() addsimps [not_Infinitesimal_not_zero, hypreal_mult_inverse]) 1); 
+qed "HFinite_not_Infinitesimal_inverse";
+
+Goal "[| x @= y; y :  HFinite - Infinitesimal |] \
+\     ==> inverse x @= inverse y";
+by (forward_tac [HFinite_diff_Infinitesimal_inf_close] 1);
+by (assume_tac 1);
+by (forward_tac [not_Infinitesimal_not_zero2] 1);
+by (forw_inst_tac [("x","x")] not_Infinitesimal_not_zero2 1);
+by (REPEAT(dtac HFinite_inverse2 1));
+by (dtac inf_close_mult2 1 THEN assume_tac 1);
+by Auto_tac;
+by (dres_inst_tac [("c","inverse x")] inf_close_mult1 1 
+    THEN assume_tac 1);
+by (auto_tac (claset() addIs [inf_close_sym],
+    simpset() addsimps [hypreal_mult_assoc]));
+qed "inf_close_inverse";
+
+(*Used for NSLIM_inverse, NSLIMSEQ_inverse*)
+bind_thm ("hypreal_of_real_inf_close_inverse",
+       hypreal_of_real_HFinite_diff_Infinitesimal RSN (2, inf_close_inverse));
+
+Goal "[| x: HFinite - Infinitesimal; \
+\        h : Infinitesimal |] ==> inverse(x + h) @= inverse x";
+by (auto_tac (claset() addIs [inf_close_inverse, inf_close_sym, 
+                              Infinitesimal_add_inf_close_self], 
+              simpset()));
+qed "inverse_add_Infinitesimal_inf_close";
+
+Goal "[| x: HFinite - Infinitesimal; \
+\        h : Infinitesimal |] ==> inverse(h + x) @= inverse x";
+by (rtac (hypreal_add_commute RS subst) 1);
+by (blast_tac (claset() addIs [inverse_add_Infinitesimal_inf_close]) 1);
+qed "inverse_add_Infinitesimal_inf_close2";
+
+Goal "[| x : HFinite - Infinitesimal; \
+\        h : Infinitesimal |] ==> inverse(x + h) + -inverse x @= h"; 
+by (rtac inf_close_trans2 1);
+by (auto_tac (claset() addIs [inverse_add_Infinitesimal_inf_close],
+              simpset() addsimps [mem_infmal_iff,
+                                  inf_close_minus_iff RS sym]));
+qed "inverse_add_Infinitesimal_inf_close_Infinitesimal";
+
+Goal "(x : Infinitesimal) = (x*x : Infinitesimal)";
+by (auto_tac (claset() addIs [Infinitesimal_mult], simpset()));
+by (rtac ccontr 1 THEN forward_tac [Infinitesimal_inverse_HFinite] 1);
+by (forward_tac [not_Infinitesimal_not_zero] 1);
+by (auto_tac (claset() addDs [Infinitesimal_HFinite_mult],
+    simpset() addsimps [hypreal_mult_assoc]));
+qed "Infinitesimal_square_iff";
+Addsimps [Infinitesimal_square_iff RS sym];
+
+Goal "(x*x : HFinite) = (x : HFinite)";
+by (auto_tac (claset() addIs [HFinite_mult], simpset()));
+by (auto_tac (claset() addDs [HInfinite_mult],
+    simpset() addsimps [HFinite_HInfinite_iff]));
+qed "HFinite_square_iff";
+Addsimps [HFinite_square_iff];
+
+Goal "(x*x : HInfinite) = (x : HInfinite)";
+by (auto_tac (claset(), simpset() addsimps 
+    [HInfinite_HFinite_iff]));
+qed "HInfinite_square_iff";
+Addsimps [HInfinite_square_iff];
+
+Goal "[| a: HFinite-Infinitesimal; a* w @= a*z |] ==> w @= z";
+by (Step_tac 1);
+by (ftac HFinite_inverse 1 THEN assume_tac 1);
+by (dtac not_Infinitesimal_not_zero 1);
+by (auto_tac (claset() addDs [inf_close_mult2],
+    simpset() addsimps [hypreal_mult_assoc RS sym]));
+qed "inf_close_HFinite_mult_cancel";
+
+Goal "a: HFinite-Infinitesimal ==> (a * w @= a * z) = (w @= z)";
+by (auto_tac (claset() addIs [inf_close_mult2,
+    inf_close_HFinite_mult_cancel], simpset()));
+qed "inf_close_HFinite_mult_cancel_iff1";
+
+(*------------------------------------------------------------------
+                  more about absolute value (hrabs)
+ ------------------------------------------------------------------*)
+
+Goal "abs x @= x | abs x @= -x";
+by (cut_inst_tac [("x","x")] hrabs_disj 1);
+by Auto_tac;
+qed "inf_close_hrabs_disj";
+
+(*------------------------------------------------------------------
+                  Theorems about monads
+ ------------------------------------------------------------------*)
+
+Goal "monad (abs x) <= monad(x) Un monad(-x)";
+by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1);
+by Auto_tac;
+qed "monad_hrabs_Un_subset";
+
+Goal "e : Infinitesimal ==> monad (x+e) = monad x";
+by (fast_tac (claset() addSIs [Infinitesimal_add_inf_close_self RS inf_close_sym,
+    inf_close_monad_iff RS iffD1]) 1);
+qed "Infinitesimal_monad_eq";
+
+Goalw [monad_def] "(u:monad x) = (-u:monad (-x))";
+by Auto_tac;
+qed "mem_monad_iff";
+
+Goalw [monad_def] "(x:Infinitesimal) = (x:monad #0)";
+by (auto_tac (claset() addIs [inf_close_sym],
+    simpset() addsimps [mem_infmal_iff]));
+qed "Infinitesimal_monad_zero_iff";
+
+Goal "(x:monad #0) = (-x:monad #0)";
+by (simp_tac (simpset() addsimps [Infinitesimal_monad_zero_iff RS sym]) 1);
+qed "monad_zero_minus_iff";
+
+Goal "(x:monad #0) = (abs x:monad #0)";
+by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1);
+by (auto_tac (claset(), simpset() addsimps [monad_zero_minus_iff RS sym]));
+qed "monad_zero_hrabs_iff";
+
+Goalw [monad_def] "x:monad x";
+by (Simp_tac 1);
+qed "mem_monad_self";
+Addsimps [mem_monad_self];
+
+(*------------------------------------------------------------------
+         Proof that x @= y ==> abs x @= abs y
+ ------------------------------------------------------------------*)
+Goal "x @= y ==> {x,y}<=monad x";
+by (Simp_tac 1);
+by (asm_full_simp_tac (simpset() addsimps 
+    [inf_close_monad_iff]) 1);
+qed "inf_close_subset_monad";
+
+Goal "x @= y ==> {x,y}<=monad y";
+by (dtac inf_close_sym 1);
+by (fast_tac (claset() addDs [inf_close_subset_monad]) 1);
+qed "inf_close_subset_monad2";
+
+Goalw [monad_def] "u:monad x ==> x @= u";
+by (Fast_tac 1);
+qed "mem_monad_inf_close";
+
+Goalw [monad_def] "x @= u ==> u:monad x";
+by (Fast_tac 1);
+qed "inf_close_mem_monad";
+
+Goalw [monad_def] "x @= u ==> x:monad u";
+by (blast_tac (claset() addSIs [inf_close_sym]) 1);
+qed "inf_close_mem_monad2";
+
+Goal "[| x @= y;x:monad #0 |] ==> y:monad #0";
+by (dtac mem_monad_inf_close 1);
+by (fast_tac (claset() addIs [inf_close_mem_monad,inf_close_trans]) 1);
+qed "inf_close_mem_monad_zero";
+
+Goal "[| x @= y; x: Infinitesimal |] ==> abs x @= abs y";
+by (dtac (Infinitesimal_monad_zero_iff RS iffD1) 1); 
+by (blast_tac (claset() addIs [inf_close_mem_monad_zero, 
+     monad_zero_hrabs_iff RS iffD1, mem_monad_inf_close, inf_close_trans3]) 1);
+qed "Infinitesimal_inf_close_hrabs";
+
+Goal "[| #0 < x;  x ~:Infinitesimal;  e :Infinitesimal |] ==> e < x";
+by (rtac ccontr 1);
+by (auto_tac (claset()
+               addIs [Infinitesimal_zero RSN (2, Infinitesimal_interval)] 
+               addSDs [hypreal_leI, order_le_imp_less_or_eq],
+              simpset()));
+qed "less_Infinitesimal_less";
+
+Goal "[| #0 < x;  x ~: Infinitesimal|] ==> ALL u: monad x. #0 < u";
+by (Step_tac 1);
+by (dtac (mem_monad_inf_close RS inf_close_sym) 1);
+by (etac (bex_Infinitesimal_iff2 RS iffD2 RS bexE) 1);
+by (dres_inst_tac [("e","-xa")] less_Infinitesimal_less 1);
+by Auto_tac;  
+qed "Ball_mem_monad_gt_zero";
+
+Goal "[| x < #0; x ~: Infinitesimal|] ==> ALL u: monad x. u < #0";
+by (Step_tac 1);
+by (dtac (mem_monad_inf_close RS inf_close_sym) 1);
+by (etac (bex_Infinitesimal_iff RS iffD2 RS bexE) 1);
+by (cut_inst_tac [("x","-x"),("e","xa")] less_Infinitesimal_less 1);
+by Auto_tac;  
+qed "Ball_mem_monad_less_zero";
+(*??GET RID OF QUANTIFIERS ABOVE??*)
+
+Goal "[|#0 < x; x ~: Infinitesimal; x @= y|] ==> #0 < y";
+by (blast_tac (claset() addDs [Ball_mem_monad_gt_zero,
+                               inf_close_subset_monad]) 1);
+qed "lemma_inf_close_gt_zero";
+
+Goal "[|x < #0; x ~: Infinitesimal; x @= y|] ==> y < #0";
+by (blast_tac (claset() addDs [Ball_mem_monad_less_zero,
+    inf_close_subset_monad]) 1);
+qed "lemma_inf_close_less_zero";
+
+Goal "[| x @= y; x < #0; x ~: Infinitesimal |] \
+\     ==> abs x @= abs y";
+by (forward_tac [lemma_inf_close_less_zero] 1);
+by (REPEAT(assume_tac 1));
+by (REPEAT(dtac hrabs_minus_eqI2 1));
+by Auto_tac;
+qed "inf_close_hrabs1";
+
+Goal "[| x @= y; #0 < x; x ~: Infinitesimal |] \
+\     ==> abs x @= abs y";
+by (forward_tac [lemma_inf_close_gt_zero] 1);
+by (REPEAT(assume_tac 1));
+by (REPEAT(dtac hrabs_eqI2 1));
+by Auto_tac;
+qed "inf_close_hrabs2";
+
+Goal "x @= y ==> abs x @= abs y";
+by (res_inst_tac [("Q","x:Infinitesimal")] (excluded_middle RS disjE) 1);
+by (res_inst_tac [("x1","x"),("y1","#0")] (hypreal_linear RS disjE) 1);
+by (auto_tac (claset() addIs [inf_close_hrabs1,inf_close_hrabs2,
+    Infinitesimal_inf_close_hrabs], simpset()));
+qed "inf_close_hrabs";
+
+Goal "abs(x) @= #0 ==> x @= #0";
+by (cut_inst_tac [("x","x")] hrabs_disj 1);
+by (auto_tac (claset() addDs [inf_close_minus], simpset()));
+qed "inf_close_hrabs_zero_cancel";
+
+Goal "e: Infinitesimal ==> abs x @= abs(x+e)";
+by (fast_tac (claset() addIs [inf_close_hrabs,
+       Infinitesimal_add_inf_close_self]) 1);
+qed "inf_close_hrabs_add_Infinitesimal";
+
+Goal "e: Infinitesimal ==> abs x @= abs(x + -e)";
+by (fast_tac (claset() addIs [inf_close_hrabs,
+    Infinitesimal_add_minus_inf_close_self]) 1);
+qed "inf_close_hrabs_add_minus_Infinitesimal";
+
+Goal "[| e: Infinitesimal; e': Infinitesimal; \
+\        abs(x+e) = abs(y+e')|] ==> abs x @= abs y";
+by (dres_inst_tac [("x","x")] inf_close_hrabs_add_Infinitesimal 1);
+by (dres_inst_tac [("x","y")] inf_close_hrabs_add_Infinitesimal 1);
+by (auto_tac (claset() addIs [inf_close_trans2], simpset()));
+qed "hrabs_add_Infinitesimal_cancel";
+
+Goal "[| e: Infinitesimal; e': Infinitesimal; \
+\        abs(x + -e) = abs(y + -e')|] ==> abs x @= abs y";
+by (dres_inst_tac [("x","x")] inf_close_hrabs_add_minus_Infinitesimal 1);
+by (dres_inst_tac [("x","y")] inf_close_hrabs_add_minus_Infinitesimal 1);
+by (auto_tac (claset() addIs [inf_close_trans2], simpset()));
+qed "hrabs_add_minus_Infinitesimal_cancel";
+
+(* interesting slightly counterintuitive theorem: necessary 
+   for proving that an open interval is an NS open set 
+*)
+Goalw [Infinitesimal_def] 
+     "[| x < y;  u: Infinitesimal |] \
+\     ==> hypreal_of_real x + u < hypreal_of_real y";
+by (dtac (hypreal_of_real_less_iff RS iffD2) 1); 
+by (dtac (hypreal_less_minus_iff RS iffD1) 1 THEN Step_tac 1);
+by (rtac (hypreal_less_minus_iff RS iffD2) 1);
+by (dres_inst_tac [("x","hypreal_of_real y + -hypreal_of_real x")] bspec 1);
+by (auto_tac (claset(),
+              simpset() addsimps [hypreal_add_commute, 
+                                  hrabs_interval_iff,
+                                  SReal_add, SReal_minus]));
+qed "Infinitesimal_add_hypreal_of_real_less";
+
+Goal "[| x: Infinitesimal; abs(hypreal_of_real r) < hypreal_of_real y |] \
+\     ==> abs (hypreal_of_real r + x) < hypreal_of_real y";
+by (dres_inst_tac [("x","hypreal_of_real r")] 
+    inf_close_hrabs_add_Infinitesimal 1);
+by (dtac (inf_close_sym RS (bex_Infinitesimal_iff2 RS iffD2)) 1);
+by (auto_tac (claset() addSIs [Infinitesimal_add_hypreal_of_real_less],
+              simpset() addsimps [hypreal_of_real_hrabs]));
+qed "Infinitesimal_add_hrabs_hypreal_of_real_less";
+
+Goal "[| x: Infinitesimal;  abs(hypreal_of_real r) < hypreal_of_real y |] \
+\     ==> abs (x + hypreal_of_real r) < hypreal_of_real y";
+by (rtac (hypreal_add_commute RS subst) 1);
+by (etac Infinitesimal_add_hrabs_hypreal_of_real_less 1);
+by (assume_tac 1);
+qed "Infinitesimal_add_hrabs_hypreal_of_real_less2";
+
+Goalw [hypreal_le_def]
+     "[| u: Infinitesimal; hypreal_of_real x + u <= hypreal_of_real y |] \
+\     ==> hypreal_of_real x <= hypreal_of_real y";
+by (EVERY1 [rtac notI, rtac contrapos_np, assume_tac]);
+by (res_inst_tac [("C","-u")] hypreal_less_add_right_cancel 1);
+by (Asm_full_simp_tac 1); 
+by (dtac (Infinitesimal_minus_iff RS iffD2) 1); 
+by (dtac Infinitesimal_add_hypreal_of_real_less 1); 
+by (assume_tac 1); 
+by Auto_tac;  
+qed "Infinitesimal_add_cancel_hypreal_of_real_le";
+
+Goal "[| u: Infinitesimal;  hypreal_of_real x + u <= hypreal_of_real y |] \
+\     ==> x <= y";
+by (blast_tac (claset() addSIs [hypreal_of_real_le_iff RS iffD1,
+               Infinitesimal_add_cancel_hypreal_of_real_le]) 1);
+qed "Infinitesimal_add_cancel_real_le";
+
+Goal "[| u: Infinitesimal; v: Infinitesimal; \
+\        hypreal_of_real x + u <= hypreal_of_real y + v |] \
+\     ==> hypreal_of_real x <= hypreal_of_real y";
+by (asm_full_simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); 
+by Auto_tac;  
+by (dres_inst_tac [("u","v-u")] Infinitesimal_add_hypreal_of_real_less 1);
+by (auto_tac (claset(), simpset() addsimps [Infinitesimal_diff]));   
+qed "hypreal_of_real_le_add_Infininitesimal_cancel";
+
+Goal "[| u: Infinitesimal; v: Infinitesimal; \
+\        hypreal_of_real x + u <= hypreal_of_real y + v |] \
+\     ==> x <= y";
+by (blast_tac (claset() addSIs [hypreal_of_real_le_iff RS iffD1,
+                          hypreal_of_real_le_add_Infininitesimal_cancel]) 1);
+qed "hypreal_of_real_le_add_Infininitesimal_cancel2";
+
+Goal "[| hypreal_of_real x < e; e: Infinitesimal |] ==> hypreal_of_real x <= #0";
+by (rtac hypreal_leI 1 THEN Step_tac 1);
+by (dtac Infinitesimal_interval 1);
+by (dtac (SReal_hypreal_of_real RS SReal_Infinitesimal_zero) 4);
+by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_zero]));
+qed "hypreal_of_real_less_Infinitesimal_le_zero";
+
+(*used once, in NSDERIV_inverse*)
+Goal "[| h: Infinitesimal; x ~= #0 |] ==> hypreal_of_real x + h ~= #0";
+by Auto_tac;  
+qed "Infinitesimal_add_not_zero";
+
+Goal "x*x + y*y : Infinitesimal ==> x*x : Infinitesimal";
+by (rtac Infinitesimal_interval2 1); 
+by (rtac hypreal_le_square 3); 
+by (rtac hypreal_self_le_add_pos 3); 
+by Auto_tac;  
+qed "Infinitesimal_square_cancel";
+Addsimps [Infinitesimal_square_cancel];
+
+Goal "x*x + y*y : HFinite ==> x*x : HFinite";
+by (rtac HFinite_bounded 1); 
+by (rtac hypreal_self_le_add_pos 2); 
+by (rtac (rename_numerals hypreal_le_square) 2); 
+by (assume_tac 1); 
+qed "HFinite_square_cancel";
+Addsimps [HFinite_square_cancel];
+
+Goal "x*x + y*y : Infinitesimal ==> y*y : Infinitesimal";
+by (rtac Infinitesimal_square_cancel 1);
+by (rtac (hypreal_add_commute RS subst) 1);
+by (Simp_tac 1);
+qed "Infinitesimal_square_cancel2";
+Addsimps [Infinitesimal_square_cancel2];
+
+Goal "x*x + y*y : HFinite ==> y*y : HFinite";
+by (rtac HFinite_square_cancel 1);
+by (rtac (hypreal_add_commute RS subst) 1);
+by (Simp_tac 1);
+qed "HFinite_square_cancel2";
+Addsimps [HFinite_square_cancel2];
+
+Goal "x*x + y*y + z*z : Infinitesimal ==> x*x : Infinitesimal";
+by (blast_tac (claset() addIs [hypreal_self_le_add_pos2,
+    Infinitesimal_interval2, rename_numerals hypreal_le_square]) 1);
+qed "Infinitesimal_sum_square_cancel";
+Addsimps [Infinitesimal_sum_square_cancel];
+
+Goal "x*x + y*y + z*z : HFinite ==> x*x : HFinite";
+by (blast_tac (claset() addIs [hypreal_self_le_add_pos2, HFinite_bounded, 
+                               rename_numerals hypreal_le_square,
+		               HFinite_number_of]) 1);
+qed "HFinite_sum_square_cancel";
+Addsimps [HFinite_sum_square_cancel];
+
+Goal "y*y + x*x + z*z : Infinitesimal ==> x*x : Infinitesimal";
+by (rtac Infinitesimal_sum_square_cancel 1);
+by (asm_full_simp_tac (simpset() addsimps hypreal_add_ac) 1);
+qed "Infinitesimal_sum_square_cancel2";
+Addsimps [Infinitesimal_sum_square_cancel2];
+
+Goal "y*y + x*x + z*z : HFinite ==> x*x : HFinite";
+by (rtac HFinite_sum_square_cancel 1);
+by (asm_full_simp_tac (simpset() addsimps hypreal_add_ac) 1);
+qed "HFinite_sum_square_cancel2";
+Addsimps [HFinite_sum_square_cancel2];
+
+Goal "z*z + y*y + x*x : Infinitesimal ==> x*x : Infinitesimal";
+by (rtac Infinitesimal_sum_square_cancel 1);
+by (asm_full_simp_tac (simpset() addsimps hypreal_add_ac) 1);
+qed "Infinitesimal_sum_square_cancel3";
+Addsimps [Infinitesimal_sum_square_cancel3];
+
+Goal "z*z + y*y + x*x : HFinite ==> x*x : HFinite";
+by (rtac HFinite_sum_square_cancel 1);
+by (asm_full_simp_tac (simpset() addsimps hypreal_add_ac) 1);
+qed "HFinite_sum_square_cancel3";
+Addsimps [HFinite_sum_square_cancel3];
+
+Goal "[| y: monad x; #0 < hypreal_of_real e |] \
+\     ==> abs (y + -x) < hypreal_of_real e";
+by (dtac (mem_monad_inf_close RS inf_close_sym) 1);
+by (dtac (bex_Infinitesimal_iff RS iffD2) 1);
+by (auto_tac (claset() addSDs [InfinitesimalD], simpset()));
+qed "monad_hrabs_less";
+
+Goal "x: monad (hypreal_of_real  a) ==> x: HFinite";
+by (dtac (mem_monad_inf_close RS inf_close_sym) 1);
+by (dtac (bex_Infinitesimal_iff2 RS iffD2) 1);
+by (step_tac (claset() addSDs 
+       [Infinitesimal_subset_HFinite RS subsetD]) 1);
+by (etac (SReal_hypreal_of_real RS (SReal_subset_HFinite 
+         RS subsetD) RS HFinite_add) 1);
+qed "mem_monad_SReal_HFinite";
+
+(*------------------------------------------------------------------
+              Theorems about standard part
+ ------------------------------------------------------------------*)
+
+Goalw [st_def] "x: HFinite ==> st x @= x";
+by (forward_tac [st_part_Ex] 1 THEN Step_tac 1);
+by (rtac someI2 1);
+by (auto_tac (claset() addIs [inf_close_sym], simpset()));
+qed "st_inf_close_self";
+
+Goalw [st_def] "x: HFinite ==> st x: SReal";
+by (forward_tac [st_part_Ex] 1 THEN Step_tac 1);
+by (rtac someI2 1);
+by (auto_tac (claset() addIs [inf_close_sym], simpset()));
+qed "st_SReal";
+
+Goal "x: HFinite ==> st x: HFinite";
+by (etac (st_SReal RS (SReal_subset_HFinite RS subsetD)) 1);
+qed "st_HFinite";
+
+Goalw [st_def] "x: SReal ==> st x = x";
+by (rtac some_equality 1);
+by (fast_tac (claset() addIs [(SReal_subset_HFinite RS subsetD)]) 1);
+by (blast_tac (claset() addDs [SReal_inf_close_iff RS iffD1]) 1);
+qed "st_SReal_eq";
+
+(* should be added to simpset *)
+Goal "st (hypreal_of_real x) = hypreal_of_real x";
+by (rtac (SReal_hypreal_of_real RS st_SReal_eq) 1);
+qed "st_hypreal_of_real";
+
+Goal "[| x: HFinite; y: HFinite; st x = st y |] ==> x @= y";
+by (auto_tac (claset() addSDs [st_inf_close_self] 
+              addSEs [inf_close_trans3], simpset()));
+qed "st_eq_inf_close";
+
+Goal "[| x: HFinite; y: HFinite; x @= y |] ==> st x = st y";
+by (EVERY1 [forward_tac [st_inf_close_self],
+    forw_inst_tac [("x","y")] st_inf_close_self,
+    dtac st_SReal,dtac st_SReal]);
+by (fast_tac (claset() addEs [inf_close_trans,
+    inf_close_trans2,SReal_inf_close_iff RS iffD1]) 1);
+qed "inf_close_st_eq";
+
+Goal "[| x: HFinite; y: HFinite|] \
+\                  ==> (x @= y) = (st x = st y)";
+by (blast_tac (claset() addIs [inf_close_st_eq,
+               st_eq_inf_close]) 1);
+qed "st_eq_inf_close_iff";
+
+Goal "[| x: SReal; e: Infinitesimal |] ==> st(x + e) = x";
+by (forward_tac [st_SReal_eq RS subst] 1);
+by (assume_tac 2);
+by (forward_tac [SReal_subset_HFinite RS subsetD] 1);
+by (forward_tac [Infinitesimal_subset_HFinite RS subsetD] 1);
+by (dtac st_SReal_eq 1);
+by (rtac inf_close_st_eq 1);
+by (auto_tac (claset() addIs  [HFinite_add],
+    simpset() addsimps [Infinitesimal_add_inf_close_self 
+    RS inf_close_sym]));
+qed "st_Infinitesimal_add_SReal";
+
+Goal "[| x: SReal; e: Infinitesimal |] \
+\     ==> st(e + x) = x";
+by (rtac (hypreal_add_commute RS subst) 1);
+by (blast_tac (claset() addSIs [st_Infinitesimal_add_SReal]) 1);
+qed "st_Infinitesimal_add_SReal2";
+
+Goal "x: HFinite ==> \
+\     EX e: Infinitesimal. x = st(x) + e";
+by (blast_tac (claset() addSDs [(st_inf_close_self RS 
+    inf_close_sym),bex_Infinitesimal_iff2 RS iffD2]) 1);
+qed "HFinite_st_Infinitesimal_add";
+
+Goal "[| x: HFinite; y: HFinite |] \
+\     ==> st (x + y) = st(x) + st(y)";
+by (forward_tac [HFinite_st_Infinitesimal_add] 1);
+by (forw_inst_tac [("x","y")] HFinite_st_Infinitesimal_add 1);
+by (Step_tac 1);
+by (subgoal_tac "st (x + y) = st ((st x + e) + (st y + ea))" 1);
+by (dtac sym 2 THEN dtac sym 2);
+by (Asm_full_simp_tac 2);
+by (asm_simp_tac (simpset() addsimps hypreal_add_ac) 1);
+by (REPEAT(dtac st_SReal 1));
+by (dtac SReal_add 1 THEN assume_tac 1);
+by (dtac Infinitesimal_add 1 THEN assume_tac 1);
+by (rtac (hypreal_add_assoc RS subst) 1);
+by (blast_tac (claset() addSIs [st_Infinitesimal_add_SReal2]) 1);
+qed "st_add";
+
+Goal "st (number_of w) = number_of w";
+by (rtac (SReal_number_of RS st_SReal_eq) 1);
+qed "st_number_of";
+Addsimps [st_number_of];
+
+Goal "y: HFinite ==> st(-y) = -st(y)";
+by (forward_tac [HFinite_minus_iff RS iffD2] 1);
+by (rtac hypreal_add_minus_eq_minus 1);
+by (dtac (st_add RS sym) 1 THEN assume_tac 1);
+by Auto_tac;  
+qed "st_minus";
+
+Goalw [hypreal_diff_def]
+     "[| x: HFinite; y: HFinite |] ==> st (x-y) = st(x) - st(y)";
+by (forw_inst_tac [("y1","y")] (st_minus RS sym) 1);
+by (dres_inst_tac [("x1","y")] (HFinite_minus_iff RS iffD2) 1);
+by (asm_simp_tac (simpset() addsimps [st_add]) 1);
+qed "st_diff";
+
+(* lemma *)
+Goal "[| x: HFinite; y: HFinite; \
+\        e: Infinitesimal;       \
+\        ea : Infinitesimal |]   \
+\      ==> e*y + x*ea + e*ea: Infinitesimal";
+by (forw_inst_tac [("x","e"),("y","y")] Infinitesimal_HFinite_mult 1);
+by (forw_inst_tac [("x","ea"),("y","x")] Infinitesimal_HFinite_mult 2);
+by (dtac Infinitesimal_mult 3);
+by (auto_tac (claset() addIs [Infinitesimal_add],
+              simpset() addsimps hypreal_add_ac @ hypreal_mult_ac));
+qed "lemma_st_mult";
+
+Goal "[| x: HFinite; y: HFinite |] \
+\              ==> st (x * y) = st(x) * st(y)";
+by (forward_tac [HFinite_st_Infinitesimal_add] 1);
+by (forw_inst_tac [("x","y")] HFinite_st_Infinitesimal_add 1);
+by (Step_tac 1);
+by (subgoal_tac "st (x * y) = st ((st x + e) * (st y + ea))" 1);
+by (dtac sym 2 THEN dtac sym 2);
+by (Asm_full_simp_tac 2);
+by (thin_tac "x = st x + e" 1);
+by (thin_tac "y = st y + ea" 1);
+by (asm_full_simp_tac (simpset() addsimps 
+    [hypreal_add_mult_distrib,hypreal_add_mult_distrib2]) 1);
+by (REPEAT(dtac st_SReal 1));
+by (full_simp_tac (simpset() addsimps [hypreal_add_assoc]) 1);
+by (rtac st_Infinitesimal_add_SReal 1);
+by (blast_tac (claset() addSIs [SReal_mult]) 1);
+by (REPEAT(dtac (SReal_subset_HFinite RS subsetD) 1));
+by (rtac (hypreal_add_assoc RS subst) 1);
+by (blast_tac (claset() addSIs [lemma_st_mult]) 1);
+qed "st_mult";
+
+Goal "x: Infinitesimal ==> st x = #0";
+by (rtac (st_number_of RS subst) 1);
+by (rtac inf_close_st_eq 1);
+by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],
+              simpset() addsimps [mem_infmal_iff RS sym]));
+qed "st_Infinitesimal";
+
+Goal "st(x) ~= #0 ==> x ~: Infinitesimal";
+by (fast_tac (claset() addIs [st_Infinitesimal]) 1);
+qed "st_not_Infinitesimal";
+
+Goal "[| x: HFinite; st x ~= #0 |] \
+\     ==> st(inverse x) = inverse (st x)";
+by (res_inst_tac [("c1","st x")] (hypreal_mult_left_cancel RS iffD1) 1);
+by (auto_tac (claset(),
+       simpset() addsimps [st_mult RS sym, st_not_Infinitesimal,
+                           HFinite_inverse]));
+by (stac hypreal_mult_inverse 1); 
+by Auto_tac;  
+qed "st_inverse";
+
+Goal "[| x: HFinite; y: HFinite; st y ~= #0 |] \
+\     ==> st(x/y) = (st x) / (st y)";
+by (auto_tac (claset(),
+      simpset() addsimps [hypreal_divide_def, st_mult, st_not_Infinitesimal, 
+                          HFinite_inverse, st_inverse]));
+qed "st_divide";
+Addsimps [st_divide];
+
+Goal "x: HFinite ==> st(st(x)) = st(x)";
+by (blast_tac (claset() addIs [st_HFinite, st_inf_close_self,
+                               inf_close_st_eq]) 1);
+qed "st_idempotent";
+Addsimps [st_idempotent];
+
+(*** lemmas ***)
+Goal "[| x: HFinite; y: HFinite; \
+\        u: Infinitesimal; st x < st y \
+\     |] ==> st x + u < st y";
+by (REPEAT(dtac st_SReal 1));
+by (auto_tac (claset() addSIs [Infinitesimal_add_hypreal_of_real_less],
+              simpset() addsimps [SReal_iff]));
+qed "Infinitesimal_add_st_less";
+
+Goalw [hypreal_le_def]
+     "[| x: HFinite; y: HFinite; \
+\        u: Infinitesimal; st x <= st y + u\
+\     |] ==> st x <= st y";
+by (auto_tac (claset() addDs [Infinitesimal_add_st_less],
+              simpset()));
+qed "Infinitesimal_add_st_le_cancel";
+
+Goal "[| x: HFinite; y: HFinite; x <= y |] \
+\     ==> st(x) <= st(y)";
+by (forward_tac [HFinite_st_Infinitesimal_add] 1);
+by (rotate_tac 1 1);
+by (forward_tac [HFinite_st_Infinitesimal_add] 1);
+by (Step_tac 1);
+by (rtac Infinitesimal_add_st_le_cancel 1);
+by (res_inst_tac [("x","ea"),("y","e")] 
+             Infinitesimal_diff 3);
+by (auto_tac (claset(),
+         simpset() addsimps [hypreal_add_assoc RS sym]));
+qed "st_le";
+
+Goal "[| #0 <= x;  x: HFinite |] ==> #0 <= st x";
+by (rtac (st_number_of RS subst) 1);
+by (auto_tac (claset() addIs [st_le],
+              simpset() delsimps [st_number_of]));
+qed "st_zero_le";
+
+Goal "[| x <= #0;  x: HFinite |] ==> st x <= #0";
+by (rtac (st_number_of RS subst) 1);
+by (auto_tac (claset() addIs [st_le],
+              simpset() delsimps [st_number_of]));
+qed "st_zero_ge";
+
+Goal "x: HFinite ==> abs(st x) = st(abs x)";
+by (case_tac "#0 <= x" 1);
+by (auto_tac (claset() addSDs [not_hypreal_leE, order_less_imp_le],
+              simpset() addsimps [st_zero_le,hrabs_eqI1, hrabs_minus_eqI1,
+                                  st_zero_ge,st_minus]));
+qed "st_hrabs";
+
+(*--------------------------------------------------------------------
+   Alternative definitions for HFinite using Free ultrafilter
+ --------------------------------------------------------------------*)
+
+Goal "[| X: Rep_hypreal x; Y: Rep_hypreal x |] \
+\     ==> {n. X n = Y n} : FreeUltrafilterNat";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by Auto_tac;
+by (Ultra_tac 1);
+qed "FreeUltrafilterNat_Rep_hypreal";
+
+Goal "{n. Yb n < Y n} Int {n. -y = Yb n} <= {n. -y < Y n}";
+by Auto_tac;
+qed "lemma_free1";
+
+Goal "{n. Xa n < Yc n} Int {n. y = Yc n} <= {n. Xa n < y}";
+by Auto_tac;
+qed "lemma_free2";
+
+Goalw [HFinite_def] 
+    "x : HFinite ==> EX X: Rep_hypreal x. \
+\    EX u. {n. abs (X n) < u}:  FreeUltrafilterNat";
+by (auto_tac (claset(), simpset() addsimps 
+    [hrabs_interval_iff]));
+by (auto_tac (claset(), simpset() addsimps 
+    [hypreal_less_def,SReal_iff,hypreal_minus,
+     hypreal_of_real_def]));
+by (dtac FreeUltrafilterNat_Rep_hypreal 1 THEN assume_tac 1);
+by (res_inst_tac [("x","Y")] bexI 1 THEN assume_tac 2);
+by (res_inst_tac [("x","y")] exI 1);
+by (Ultra_tac 1 THEN arith_tac 1);
+qed "HFinite_FreeUltrafilterNat";
+
+Goalw [HFinite_def] 
+     "EX X: Rep_hypreal x. \
+\      EX u. {n. abs (X n) < u}: FreeUltrafilterNat\
+\      ==>  x : HFinite";
+by (auto_tac (claset(), simpset() addsimps 
+    [hrabs_interval_iff]));
+by (res_inst_tac [("x","hypreal_of_real u")] bexI 1);
+by (auto_tac (claset() addSIs [exI], simpset() addsimps 
+    [hypreal_less_def,SReal_iff,hypreal_minus,
+     hypreal_of_real_def]));
+by (ALLGOALS(Ultra_tac THEN' arith_tac));
+qed "FreeUltrafilterNat_HFinite";
+
+Goal "(x : HFinite) = (EX X: Rep_hypreal x. \
+\          EX u. {n. abs (X n) < u}:  FreeUltrafilterNat)";
+by (blast_tac (claset() addSIs [HFinite_FreeUltrafilterNat,
+               FreeUltrafilterNat_HFinite]) 1);
+qed "HFinite_FreeUltrafilterNat_iff";
+
+(*--------------------------------------------------------------------
+   Alternative definitions for HInfinite using Free ultrafilter
+ --------------------------------------------------------------------*)
+Goal "- {n. (u::real) < abs (xa n)} = {n. abs (xa n) <= u}";
+by Auto_tac;
+qed "lemma_Compl_eq";
+
+Goal "- {n. abs (xa n) < (u::real)} = {n. u <= abs (xa n)}";
+by Auto_tac;
+qed "lemma_Compl_eq2";
+
+Goal "{n. abs (xa n) <= (u::real)} Int {n. u <= abs (xa n)} \
+\         = {n. abs(xa n) = u}";
+by Auto_tac;
+qed "lemma_Int_eq1";
+
+Goal "{n. abs (xa n) = u} <= {n. abs (xa n) < u + (#1::real)}";
+by Auto_tac;
+qed "lemma_FreeUltrafilterNat_one";
+
+(*-------------------------------------
+  Exclude this type of sets from free 
+  ultrafilter for Infinite numbers!
+ -------------------------------------*)
+Goal "[| xa: Rep_hypreal x; \
+\                 {n. abs (xa n) = u} : FreeUltrafilterNat \
+\              |] ==> x: HFinite";
+by (rtac FreeUltrafilterNat_HFinite 1);
+by (res_inst_tac [("x","xa")] bexI 1);
+by (res_inst_tac [("x","u + #1")] exI 1);
+by (Ultra_tac 1 THEN assume_tac 1);
+qed "FreeUltrafilterNat_const_Finite";
+
+val [prem] = goal thy "x : HInfinite ==> EX X: Rep_hypreal x. \
+\          ALL u. {n. u < abs (X n)}:  FreeUltrafilterNat";
+by (cut_facts_tac [(prem RS (HInfinite_HFinite_iff RS iffD1))] 1);
+by (cut_inst_tac [("x","x")] Rep_hypreal_nonempty 1);
+by (auto_tac (claset(), simpset() delsimps [Rep_hypreal_nonempty] 
+    addsimps [HFinite_FreeUltrafilterNat_iff,Bex_def]));
+by (REPEAT(dtac spec 1));
+by Auto_tac;
+by (dres_inst_tac [("x","u")] spec 1 THEN 
+    REPEAT(dtac FreeUltrafilterNat_Compl_mem 1));
+by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
+
+
+by (asm_full_simp_tac (simpset() addsimps [lemma_Compl_eq,
+    lemma_Compl_eq2,lemma_Int_eq1]) 1);
+by (auto_tac (claset() addDs [FreeUltrafilterNat_const_Finite],
+    simpset() addsimps [(prem RS (HInfinite_HFinite_iff RS iffD1))]));
+qed "HInfinite_FreeUltrafilterNat";
+
+(* yet more lemmas! *)
+Goal "{n. abs (Xa n) < u} Int {n. X n = Xa n} \
+\         <= {n. abs (X n) < (u::real)}";
+by Auto_tac;
+qed "lemma_Int_HI";
+
+Goal "{n. u < abs (X n)} Int {n. abs (X n) < (u::real)} = {}";
+by (auto_tac (claset() addIs [real_less_asym], simpset()));
+qed "lemma_Int_HIa";
+
+Goal "EX X: Rep_hypreal x. ALL u. \
+\              {n. u < abs (X n)}: FreeUltrafilterNat\
+\              ==>  x : HInfinite";
+by (rtac (HInfinite_HFinite_iff RS iffD2) 1);
+by (Step_tac 1 THEN dtac HFinite_FreeUltrafilterNat 1);
+by Auto_tac;
+by (dres_inst_tac [("x","u")] spec 1);
+by (dtac FreeUltrafilterNat_Rep_hypreal 1 THEN assume_tac 1);
+by (dres_inst_tac [("Y","{n. X n = Xa n}")] FreeUltrafilterNat_Int 1);
+by (dtac (lemma_Int_HI RSN (2,FreeUltrafilterNat_subset)) 2);
+by (dres_inst_tac [("Y","{n. abs (X n) < u}")] FreeUltrafilterNat_Int 2);
+by (auto_tac (claset(), simpset() addsimps [lemma_Int_HIa,
+              FreeUltrafilterNat_empty]));
+qed "FreeUltrafilterNat_HInfinite";
+
+Goal "(x : HInfinite) = (EX X: Rep_hypreal x. \
+\          ALL u. {n. u < abs (X n)}:  FreeUltrafilterNat)";
+by (blast_tac (claset() addSIs [HInfinite_FreeUltrafilterNat,
+               FreeUltrafilterNat_HInfinite]) 1);
+qed "HInfinite_FreeUltrafilterNat_iff";
+
+(*--------------------------------------------------------------------
+   Alternative definitions for Infinitesimal using Free ultrafilter
+ --------------------------------------------------------------------*)
+
+Goal "{n. - u < Yd n} Int {n. xa n = Yd n} <= {n. -u < xa n}";
+by Auto_tac;
+qed "lemma_free4";
+
+Goal "{n. Yb n < u} Int {n. xa n = Yb n} <= {n. xa n < u}";
+by Auto_tac;
+qed "lemma_free5";
+
+Goalw [Infinitesimal_def] 
+          "x : Infinitesimal ==> EX X: Rep_hypreal x. \
+\          ALL u. #0 < u --> {n. abs (X n) < u}:  FreeUltrafilterNat";
+by (auto_tac (claset(), simpset() addsimps [hrabs_interval_iff]));
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (EVERY[Auto_tac, rtac bexI 1, rtac lemma_hyprel_refl 2, Step_tac 1]);
+by (dtac (hypreal_of_real_less_iff RS iffD2) 1);
+by (dres_inst_tac [("x","hypreal_of_real u")] bspec 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_zero]));
+by (auto_tac (claset(), 
+              simpset() addsimps [hypreal_less_def, hypreal_minus,
+                                  hypreal_of_real_def,hypreal_of_real_zero]));
+by (Ultra_tac 1 THEN arith_tac 1);
+qed "Infinitesimal_FreeUltrafilterNat";
+
+Goalw [Infinitesimal_def] 
+     "EX X: Rep_hypreal x. \
+\           ALL u. #0 < u --> {n. abs (X n) < u} : FreeUltrafilterNat \
+\     ==> x : Infinitesimal";
+by (auto_tac (claset(),
+              simpset() addsimps [hrabs_interval_iff,abs_interval_iff]));
+by (auto_tac (claset(),
+              simpset() addsimps [SReal_iff]));
+by (auto_tac (claset() addSIs [exI] 
+                               addIs [FreeUltrafilterNat_subset],
+    simpset() addsimps [hypreal_less_def, hypreal_minus,hypreal_of_real_def]));
+qed "FreeUltrafilterNat_Infinitesimal";
+
+Goal "(x : Infinitesimal) = (EX X: Rep_hypreal x. \
+\          ALL u. #0 < u --> {n. abs (X n) < u}:  FreeUltrafilterNat)";
+by (blast_tac (claset() addSIs [Infinitesimal_FreeUltrafilterNat,
+               FreeUltrafilterNat_Infinitesimal]) 1);
+qed "Infinitesimal_FreeUltrafilterNat_iff";
+
+(*------------------------------------------------------------------------
+         Infinitesimals as smaller than 1/n for all n::nat (> 0)
+ ------------------------------------------------------------------------*)
+Goal "(ALL r. #0 < r --> x < r) = (ALL n. x < inverse(real_of_posnat n))";
+by (auto_tac (claset(),
+        simpset() addsimps [rename_numerals real_of_posnat_gt_zero]));
+by (blast_tac (claset() addSDs [rename_numerals reals_Archimedean] 
+                        addIs [order_less_trans]) 1);
+qed "lemma_Infinitesimal";
+
+Goal "(ALL r: SReal. #0 < r --> x < r) = \
+\     (ALL n. x < inverse(hypreal_of_posnat n))";
+by (Step_tac 1);
+by (dres_inst_tac [("x","inverse (hypreal_of_real(real_of_posnat n))")] 
+    bspec 1);
+by (full_simp_tac (simpset() addsimps [SReal_inverse]) 1); 
+by (rtac (real_of_posnat_gt_zero RS real_inverse_gt_zero RS 
+          (hypreal_of_real_less_iff RS iffD2) RSN(2,impE)) 1);
+by (assume_tac 2);
+by (asm_full_simp_tac (simpset() addsimps 
+       [rename_numerals real_of_posnat_gt_zero, hypreal_of_real_zero,
+        hypreal_of_real_of_posnat]) 1);
+by (auto_tac (claset() addSDs [reals_Archimedean],
+              simpset() addsimps [SReal_iff,hypreal_of_real_zero RS sym]));
+by (dtac (hypreal_of_real_less_iff RS iffD2) 1);
+by (asm_full_simp_tac (simpset() addsimps 
+       [rename_numerals real_of_posnat_gt_zero,hypreal_of_real_of_posnat])1);
+by (blast_tac (claset() addIs [order_less_trans]) 1);
+qed "lemma_Infinitesimal2";
+
+Goalw [Infinitesimal_def] 
+     "Infinitesimal = {x. ALL n. abs x < inverse (hypreal_of_posnat n)}";
+by (auto_tac (claset(), simpset() addsimps [lemma_Infinitesimal2]));
+qed "Infinitesimal_hypreal_of_posnat_iff";
+
+(*-----------------------------------------------------------------------------
+       Actual proof that omega (whr) is an infinite number and 
+       hence that epsilon (ehr) is an infinitesimal number.
+ -----------------------------------------------------------------------------*)
+Goal "{n. n < Suc m} = {n. n < m} Un {n. n = m}";
+by (auto_tac (claset(), simpset() addsimps [less_Suc_eq]));
+qed "Suc_Un_eq";
+
+(*-------------------------------------------
+  Prove that any segment is finite and 
+  hence cannot belong to FreeUltrafilterNat
+ -------------------------------------------*)
+Goal "finite {n::nat. n < m}";
+by (nat_ind_tac "m" 1);
+by (auto_tac (claset(), simpset() addsimps [Suc_Un_eq]));
+qed "finite_nat_segment";
+
+Goal "finite {n::nat. real_of_posnat n < real_of_posnat m}";
+by (auto_tac (claset() addIs [finite_nat_segment], simpset()));
+qed "finite_real_of_posnat_segment";
+
+Goal "finite {n. real_of_posnat n < u}";
+by (cut_inst_tac [("x","u")] reals_Archimedean2 1);
+by (Step_tac 1);
+by (rtac (finite_real_of_posnat_segment RSN (2,finite_subset)) 1);
+by (auto_tac (claset() addDs [order_less_trans], simpset()));
+qed "finite_real_of_posnat_less_real";
+
+Goal "{n. real_of_posnat n <= u} = \
+\     {n. real_of_posnat n < u} Un {n. u = real_of_posnat n}";
+by (auto_tac (claset() addDs [order_le_imp_less_or_eq],
+              simpset() addsimps [order_less_imp_le]));
+qed "lemma_real_le_Un_eq";
+
+Goal "finite {n. real_of_posnat n <= u}";
+by (auto_tac (claset(), simpset() addsimps 
+    [lemma_real_le_Un_eq,lemma_finite_omega_set,
+     finite_real_of_posnat_less_real]));
+qed "finite_real_of_posnat_le_real";
+
+Goal "finite {n. abs(real_of_posnat n) <= u}";
+by (full_simp_tac (simpset() addsimps [rename_numerals 
+    real_of_posnat_gt_zero,abs_eqI2,
+    finite_real_of_posnat_le_real]) 1);
+qed "finite_rabs_real_of_posnat_le_real";
+
+Goal "{n. abs(real_of_posnat n) <= u} ~: FreeUltrafilterNat";
+by (blast_tac (claset() addSIs [FreeUltrafilterNat_finite,
+    finite_rabs_real_of_posnat_le_real]) 1);
+qed "rabs_real_of_posnat_le_real_FreeUltrafilterNat";
+
+Goal "{n. u < real_of_posnat n} : FreeUltrafilterNat";
+by (rtac ccontr 1 THEN dtac FreeUltrafilterNat_Compl_mem 1);
+by (auto_tac (claset(), simpset() addsimps 
+    [CLAIM_SIMP "- {n. u < real_of_posnat  n} = \
+\                {n. real_of_posnat n <= u}" 
+     [real_le_def],finite_real_of_posnat_le_real 
+                   RS FreeUltrafilterNat_finite]));
+qed "FreeUltrafilterNat_nat_gt_real";
+
+(*--------------------------------------------------------------
+ The complement of {n. abs(real_of_posnat n) <= u} = 
+ {n. u < abs (real_of_posnat n)} is in FreeUltrafilterNat 
+ by property of (free) ultrafilters
+ --------------------------------------------------------------*)
+Goal "- {n. abs (real_of_posnat  n) <= u} = \
+\     {n. u < abs (real_of_posnat  n)}";
+by (auto_tac (claset() addSDs [order_le_less_trans],
+              simpset() addsimps [not_real_leE]));
+val lemma = result();
+
+Goal "{n. u < abs (real_of_posnat n)} : FreeUltrafilterNat";
+by (cut_inst_tac [("u","u")] rabs_real_of_posnat_le_real_FreeUltrafilterNat 1);
+by (auto_tac (claset() addDs [FreeUltrafilterNat_Compl_mem],
+    simpset() addsimps [lemma]));
+qed "FreeUltrafilterNat_omega";
+
+(*-----------------------------------------------
+       Omega is a member of HInfinite
+ -----------------------------------------------*)
+Goalw [omega_def] "whr: HInfinite";
+by (auto_tac (claset() addSIs [FreeUltrafilterNat_HInfinite,
+    lemma_hyprel_refl,FreeUltrafilterNat_omega], simpset()));
+qed "HInfinite_omega";
+
+(*-----------------------------------------------
+       Epsilon is a member of Infinitesimal
+ -----------------------------------------------*)
+
+Goal "ehr : Infinitesimal";
+by (auto_tac (claset() addSIs [HInfinite_inverse_Infinitesimal,HInfinite_omega],
+    simpset() addsimps [hypreal_epsilon_inverse_omega]));
+qed "Infinitesimal_epsilon";
+Addsimps [Infinitesimal_epsilon];
+
+Goal "ehr : HFinite";
+by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],
+    simpset()));
+qed "HFinite_epsilon";
+Addsimps [HFinite_epsilon];
+
+Goal "ehr @= #0";
+by (simp_tac (simpset() addsimps [mem_infmal_iff RS sym]) 1);
+qed "epsilon_inf_close_zero";
+Addsimps [epsilon_inf_close_zero];
+
+(*------------------------------------------------------------------------
+  Needed for proof that we define a hyperreal [<X(n)] @= hypreal_of_real a given 
+  that ALL n. |X n - a| < 1/n. Used in proof of NSLIM => LIM.
+ -----------------------------------------------------------------------*)
+
+Goal "#0 < u ==> finite {n. u < inverse(real_of_posnat n)}";
+by (asm_simp_tac (simpset() addsimps [real_of_posnat_less_inverse_iff]) 1);
+by (rtac finite_real_of_posnat_less_real 1);
+qed "finite_inverse_real_of_posnat_gt_real";
+
+Goal "{n. u <= inverse(real_of_posnat n)} = \
+\    {n. u < inverse(real_of_posnat n)} Un {n. u = inverse(real_of_posnat n)}";
+by (auto_tac (claset() addDs [order_le_imp_less_or_eq],
+              simpset() addsimps [order_less_imp_le]));
+qed "lemma_real_le_Un_eq2";
+
+Goal "#0 < u ==> finite {n::nat. u = inverse(real_of_posnat  n)}";
+by (asm_simp_tac (simpset() addsimps [real_of_posnat_inverse_eq_iff]) 1);
+by (auto_tac (claset() addIs [lemma_finite_omega_set RSN 
+    (2,finite_subset)], simpset()));
+qed "lemma_finite_omega_set2";
+
+Goal "#0 < u ==> finite {n. u <= inverse(real_of_posnat n)}";
+by (auto_tac (claset(), 
+      simpset() addsimps [lemma_real_le_Un_eq2,lemma_finite_omega_set2,
+                          finite_inverse_real_of_posnat_gt_real]));
+qed "finite_inverse_real_of_posnat_ge_real";
+
+Goal "#0 < u ==> \
+\    {n. u <= inverse(real_of_posnat n)} ~: FreeUltrafilterNat";
+by (blast_tac (claset() addSIs [FreeUltrafilterNat_finite,
+                                finite_inverse_real_of_posnat_ge_real]) 1);
+qed "inverse_real_of_posnat_ge_real_FreeUltrafilterNat";
+
+(*--------------------------------------------------------------
+    The complement of  {n. u <= inverse(real_of_posnat n)} =
+    {n. inverse(real_of_posnat n) < u} is in FreeUltrafilterNat 
+    by property of (free) ultrafilters
+ --------------------------------------------------------------*)
+Goal "- {n. u <= inverse(real_of_posnat n)} = \
+\     {n. inverse(real_of_posnat n) < u}";
+by (auto_tac (claset() addSDs [order_le_less_trans],
+              simpset() addsimps [not_real_leE]));
+val lemma = result();
+
+Goal "#0 < u ==> \
+\     {n. inverse(real_of_posnat n) < u} : FreeUltrafilterNat";
+by (cut_inst_tac [("u","u")]  inverse_real_of_posnat_ge_real_FreeUltrafilterNat 1);
+by (auto_tac (claset() addDs [FreeUltrafilterNat_Compl_mem],
+    simpset() addsimps [lemma]));
+qed "FreeUltrafilterNat_inverse_real_of_posnat";
+
+(*--------------------------------------------------------------
+      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 (the theorem was previously in REALTOPOS.thy).
+ -------------------------------------------------------------*)
+(*-----------------------------------------------------
+    |X(n) - x| < 1/n ==> [<X n>] - hypreal_of_real x|: Infinitesimal 
+ -----------------------------------------------------*)
+Goal "ALL n. abs(X n + -x) < inverse(real_of_posnat n) \ 
+\    ==> Abs_hypreal(hyprel^^{X}) + -hypreal_of_real x : Infinitesimal";
+by (auto_tac (claset() addSIs [bexI] 
+           addDs [rename_numerals FreeUltrafilterNat_inverse_real_of_posnat,
+                  FreeUltrafilterNat_all,FreeUltrafilterNat_Int] 
+           addIs [order_less_trans, FreeUltrafilterNat_subset],
+      simpset() addsimps [hypreal_minus, 
+                          hypreal_of_real_def,hypreal_add,
+                          Infinitesimal_FreeUltrafilterNat_iff,hypreal_inverse,
+                          hypreal_of_real_of_posnat]));
+qed "real_seq_to_hypreal_Infinitesimal";
+
+Goal "ALL n. abs(X n + -x) < inverse(real_of_posnat n) \ 
+\     ==> Abs_hypreal(hyprel^^{X}) @= hypreal_of_real x";
+by (rtac (inf_close_minus_iff RS ssubst) 1);
+by (rtac (mem_infmal_iff RS subst) 1);
+by (etac real_seq_to_hypreal_Infinitesimal 1);
+qed "real_seq_to_hypreal_inf_close";
+
+Goal "ALL n. abs(x + -X n) < inverse(real_of_posnat n) \ 
+\              ==> Abs_hypreal(hyprel^^{X}) @= hypreal_of_real x";
+by (asm_full_simp_tac (simpset() addsimps [abs_minus_add_cancel,
+        real_seq_to_hypreal_inf_close]) 1);
+qed "real_seq_to_hypreal_inf_close2";
+
+Goal "ALL n. abs(X n + -Y n) < inverse(real_of_posnat n) \ 
+\     ==> Abs_hypreal(hyprel^^{X}) + \
+\         -Abs_hypreal(hyprel^^{Y}) : Infinitesimal";
+by (auto_tac (claset() addSIs [bexI] 
+                  addDs [rename_numerals 
+                         FreeUltrafilterNat_inverse_real_of_posnat,
+                         FreeUltrafilterNat_all,FreeUltrafilterNat_Int] 
+        addIs [order_less_trans, FreeUltrafilterNat_subset],
+     simpset() addsimps 
+            [Infinitesimal_FreeUltrafilterNat_iff,hypreal_minus,hypreal_add, 
+             hypreal_inverse,hypreal_of_real_of_posnat]));
+qed "real_seq_to_hypreal_Infinitesimal2";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/NSA.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,48 @@
+(*  Title       : NSA.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Infinite numbers, Infinitesimals,
+                  infinitely close relation etc.
+*) 
+
+NSA = HRealAbs + RComplete +
+
+constdefs
+
+  Infinitesimal  :: "hypreal set"
+   "Infinitesimal == {x. ALL r: SReal. 0 < r --> abs x < r}"
+
+  HFinite :: "hypreal set"
+   "HFinite == {x. EX r: SReal. abs x < r}"
+
+  HInfinite :: "hypreal set"
+   "HInfinite == {x. ALL r: SReal. r < abs x}"
+
+  (* standard part map *)
+  st        :: hypreal => hypreal
+   "st           == (%x. @r. x : HFinite & r:SReal & r @= x)"
+
+  monad     :: hypreal => hypreal set
+   "monad x      == {y. x @= y}"
+
+  galaxy    :: hypreal => hypreal set
+   "galaxy x     == {y. (x + -y) : HFinite}"
+ 
+  (* infinitely close *)
+  inf_close :: "[hypreal, hypreal] => bool"    (infixl "@=" 50)
+   "x @= y       == (x + -y) : Infinitesimal"     
+
+
+defs  
+
+   (*standard real numbers as a subset of the hyperreals*)
+   SReal_def      "SReal == {x. EX r. x = hypreal_of_real r}"
+
+syntax (symbols)
+    inf_close :: "[hypreal, hypreal] => bool"    (infixl "\\<approx>" 50)
+  
+end
+
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/NatStar.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,481 @@
+(*  Title       : NatStar.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : *-transforms in NSA which extends 
+                  sets of reals, and nat=>real, 
+                  nat=>nat functions
+*) 
+
+Goalw [starsetNat_def] 
+      "*sNat*(UNIV::nat set) = (UNIV::hypnat set)";
+by (auto_tac (claset(), simpset() addsimps [FreeUltrafilterNat_Nat_set]));
+qed "NatStar_real_set";
+
+Goalw [starsetNat_def] "*sNat* {} = {}";
+by (Step_tac 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (dres_inst_tac [("x","%n. xa n")] bspec 1);
+by (auto_tac (claset(), simpset() addsimps [FreeUltrafilterNat_empty]));
+qed "NatStar_empty_set";
+
+Addsimps [NatStar_empty_set];
+
+Goalw [starsetNat_def] 
+      "*sNat* (A Un B) = *sNat* A Un *sNat* B";
+by (Auto_tac);
+by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2));
+by (dtac FreeUltrafilterNat_Compl_mem 1);
+by (dtac bspec 1 THEN assume_tac 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (Auto_tac);
+by (Fuf_tac 1);
+qed "NatStar_Un";
+
+Goalw [starsetNat_n_def] 
+      "*sNatn* (%n. (A n) Un (B n)) = *sNatn* A Un *sNatn* B";
+by Auto_tac;
+by (dres_inst_tac [("x","Xa")] bspec 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 2);
+by (auto_tac (claset() addSDs [bspec], simpset()));
+by (TRYALL(Ultra_tac));
+qed "starsetNat_n_Un";
+
+Goalw [InternalNatSets_def]
+     "[| X : InternalNatSets; Y : InternalNatSets |] \
+\     ==> (X Un Y) : InternalNatSets";
+by (auto_tac (claset(),
+         simpset() addsimps [starsetNat_n_Un RS sym]));
+qed "InternalNatSets_Un";
+
+Goalw [starsetNat_def] "*sNat* (A Int B) = *sNat* A Int *sNat* B";
+by (Auto_tac);
+by (blast_tac (claset() addIs [FreeUltrafilterNat_Int,
+    FreeUltrafilterNat_subset]) 3);
+by (REPEAT(blast_tac (claset() addIs 
+    [FreeUltrafilterNat_subset]) 1));
+qed "NatStar_Int";
+
+Goalw [starsetNat_n_def] 
+      "*sNatn* (%n. (A n) Int (B n)) = *sNatn* A Int *sNatn* B";
+by (Auto_tac);
+by (auto_tac (claset() addSDs [bspec],
+         simpset()));
+by (TRYALL(Ultra_tac));
+qed "starsetNat_n_Int";
+
+Goalw [InternalNatSets_def]
+     "[| X : InternalNatSets; Y : InternalNatSets |] \
+\     ==> (X Int Y) : InternalNatSets";
+by (auto_tac (claset(),
+         simpset() addsimps [starsetNat_n_Int RS sym]));
+qed "InternalNatSets_Int";
+
+Goalw [starsetNat_def] "*sNat* (-A) = -(*sNat* A)";
+by (Auto_tac);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 2);
+by (REPEAT(Step_tac 1) THEN Auto_tac);
+by (TRYALL(Ultra_tac));
+qed "NatStar_Compl";
+
+Goalw [starsetNat_n_def] "*sNatn* ((%n. - A n)) = -(*sNatn* A)";
+by (Auto_tac);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 2);
+by (REPEAT(Step_tac 1) THEN Auto_tac);
+by (TRYALL(Ultra_tac));
+qed "starsetNat_n_Compl";
+
+Goalw [InternalNatSets_def]
+     "X :InternalNatSets ==> -X : InternalNatSets";
+by (auto_tac (claset(),
+         simpset() addsimps [starsetNat_n_Compl RS sym]));
+qed "InternalNatSets_Compl";
+
+Goalw [starsetNat_n_def] 
+      "*sNatn* (%n. (A n) - (B n)) = *sNatn* A - *sNatn* B";
+by (Auto_tac);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 2);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 3);
+by (auto_tac (claset() addSDs [bspec], simpset()));
+by (TRYALL(Ultra_tac));
+qed "starsetNat_n_diff";
+
+Goalw [InternalNatSets_def]
+     "[| X : InternalNatSets; Y : InternalNatSets |] \
+\     ==> (X - Y) : InternalNatSets";
+by (auto_tac (claset(), simpset() addsimps [starsetNat_n_diff RS sym]));
+qed "InternalNatSets_diff";
+
+Goalw [starsetNat_def] "A <= B ==> *sNat* A <= *sNat* B";
+by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1));
+qed "NatStar_subset";
+
+Goalw [starsetNat_def,hypnat_of_nat_def] 
+          "a : A ==> hypnat_of_nat a : *sNat* A";
+by (auto_tac (claset() addIs [FreeUltrafilterNat_subset],
+         simpset()));
+qed "NatStar_mem";
+
+Goalw [starsetNat_def] "hypnat_of_nat `` A <= *sNat* A";
+by (auto_tac (claset(), simpset() addsimps [hypnat_of_nat_def]));
+by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1);
+qed "NatStar_hypreal_of_real_image_subset";
+
+Goal "SHNat <= *sNat* (UNIV:: nat set)";
+by (simp_tac (simpset() addsimps [SHNat_hypnat_of_nat_iff,
+    NatStar_hypreal_of_real_image_subset]) 1);
+qed "NatStar_SHNat_subset";
+
+Goalw [starsetNat_def] 
+      "*sNat* X Int SHNat = hypnat_of_nat `` X";
+by (auto_tac (claset(),
+         simpset() addsimps 
+           [hypnat_of_nat_def,SHNat_def]));
+by (fold_tac [hypnat_of_nat_def]);
+by (rtac imageI 1 THEN rtac ccontr 1);
+by (dtac bspec 1);
+by (rtac lemma_hypnatrel_refl 1);
+by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2);
+by (Auto_tac);
+qed "NatStar_hypreal_of_real_Int";
+
+Goal "x ~: hypnat_of_nat `` A ==> ALL y: A. x ~= hypnat_of_nat y";
+by (Auto_tac);
+qed "lemma_not_hypnatA";
+
+Goalw [starsetNat_n_def,starsetNat_def] "*sNat* X = *sNatn* (%n. X)";
+by Auto_tac;
+qed "starsetNat_starsetNat_n_eq";
+
+Goalw [InternalNatSets_def] "(*sNat* X) : InternalNatSets";
+by (auto_tac (claset(),
+         simpset() addsimps [starsetNat_starsetNat_n_eq]));
+qed "InternalNatSets_starsetNat_n";
+Addsimps [InternalNatSets_starsetNat_n];
+
+Goal "X : InternalNatSets ==> UNIV - X : InternalNatSets";
+by (auto_tac (claset() addIs [InternalNatSets_Compl], simpset()));
+qed "InternalNatSets_UNIV_diff";
+
+(*------------------------------------------------------------------ 
+   Nonstandard extension of a set (defined using a constant 
+   sequence) as a special case of an internal set
+ -----------------------------------------------------------------*)
+
+Goalw [starsetNat_n_def,starsetNat_def] 
+     "ALL n. (As n = A) ==> *sNatn* As = *sNat* A";
+by (Auto_tac);
+qed "starsetNat_n_starsetNat";
+
+(*------------------------------------------------------
+   Theorems about nonstandard extensions of functions
+ ------------------------------------------------------*)
+
+(*------------------------------------------------------------------ 
+   Nonstandard extension of a function (defined using a constant 
+   sequence) as a special case of an internal function
+ -----------------------------------------------------------------*)
+
+Goalw [starfunNat_n_def,starfunNat_def] 
+     "ALL n. (F n = f) ==> *fNatn* F = *fNat* f";
+by (Auto_tac);
+qed "starfunNat_n_starfunNat";
+
+Goalw [starfunNat2_n_def,starfunNat2_def] 
+     "ALL n. (F n = f) ==> *fNat2n* F = *fNat2* f";
+by (Auto_tac);
+qed "starfunNat2_n_starfunNat2";
+
+Goalw [congruent_def] 
+      "congruent hypnatrel (%X. hypnatrel^^{%n. f (X n)})";
+by (safe_tac (claset()));
+by (ALLGOALS(Fuf_tac));
+qed "starfunNat_congruent";
+
+(* f::nat=>real *)
+Goalw [starfunNat_def]
+      "(*fNat* f) (Abs_hypnat(hypnatrel^^{%n. X n})) = \
+\      Abs_hypreal(hyprel ^^ {%n. f (X n)})";
+by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
+by (simp_tac (simpset() addsimps 
+   [hyprel_in_hypreal RS Abs_hypreal_inverse]) 1);
+by (Auto_tac THEN Fuf_tac 1);
+qed "starfunNat";
+
+(* f::nat=>nat *)
+Goalw [starfunNat2_def]
+      "(*fNat2* f) (Abs_hypnat(hypnatrel^^{%n. X n})) = \
+\      Abs_hypnat(hypnatrel ^^ {%n. f (X n)})";
+by (res_inst_tac [("f","Abs_hypnat")] arg_cong 1);
+by (simp_tac (simpset() addsimps 
+   [hypnatrel_in_hypnat RS Abs_hypnat_inverse,
+    [equiv_hypnatrel, starfunNat_congruent] MRS UN_equiv_class]) 1);
+qed "starfunNat2";
+
+(*---------------------------------------------
+  multiplication: ( *f ) x ( *g ) = *(f x g)  
+ ---------------------------------------------*)
+Goal "(*fNat* f) z * (*fNat* g) z = (*fNat* (%x. f x * g x)) z";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat,hypreal_mult]));
+qed "starfunNat_mult";
+
+Goal "(*fNat2* f) z * (*fNat2* g) z = (*fNat2* (%x. f x * g x)) z";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat2,hypnat_mult]));
+qed "starfunNat2_mult";
+
+(*---------------------------------------
+  addition: ( *f ) + ( *g ) = *(f + g)  
+ ---------------------------------------*)
+Goal "(*fNat* f) z + (*fNat* g) z = (*fNat* (%x. f x + g x)) z";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat,hypreal_add]));
+qed "starfunNat_add";
+
+Goal "(*fNat2* f) z + (*fNat2* g) z = (*fNat2* (%x. f x + g x)) z";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat2,hypnat_add]));
+qed "starfunNat2_add";
+
+Goal "(*fNat2* f) z - (*fNat2* g) z = (*fNat2* (%x. f x - g x)) z";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat2, hypnat_minus]));
+qed "starfunNat2_minus";
+
+(*--------------------------------------
+  composition: ( *f ) o ( *g ) = *(f o g)  
+ ---------------------------------------*)
+(***** ( *f::nat=>real ) o ( *g::nat=>nat ) = *(f o g) *****)
+ 
+Goal "(*fNat* f) o (*fNat2* g) = (*fNat* (f o g))";
+by (rtac ext 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat2, starfunNat]));
+qed "starfunNatNat2_o";
+
+Goal "(%x. (*fNat* f) ((*fNat2* g) x)) = (*fNat* (%x. f(g x)))";
+by (rtac ( simplify (simpset() addsimps [o_def]) starfunNatNat2_o) 1);
+qed "starfunNatNat2_o2";
+
+(***** ( *f::nat=>nat ) o ( *g::nat=>nat ) = *(f o g) *****)
+Goal "(*fNat2* f) o (*fNat2* g) = (*fNat2* (f o g))";
+by (rtac ext 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat2]));
+qed "starfunNat2_o";
+
+(***** ( *f::real=>real ) o ( *g::nat=>real ) = *(f o g) *****)
+
+Goal "(*f* f) o (*fNat* g) = (*fNat* (f o g))"; 
+by (rtac ext 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat,starfun]));
+qed "starfun_stafunNat_o";
+
+Goal "(%x. (*f* f) ((*fNat* g) x)) = (*fNat* (%x. f (g x)))"; 
+by (rtac ( simplify (simpset() addsimps [o_def]) starfun_stafunNat_o) 1);
+qed "starfun_stafunNat_o2";
+
+(*--------------------------------------
+  NS extension of constant function
+ --------------------------------------*)
+Goal "(*fNat* (%x. k)) z = hypreal_of_real k";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat, hypreal_of_real_def]));
+qed "starfunNat_const_fun";
+Addsimps [starfunNat_const_fun];
+
+Goal "(*fNat2* (%x. k)) z = hypnat_of_nat  k";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat2, hypnat_of_nat_def]));
+qed "starfunNat2_const_fun";
+
+Addsimps [starfunNat2_const_fun];
+
+Goal "- (*fNat* f) x = (*fNat* (%x. - f x)) x";
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat, hypreal_minus]));
+qed "starfunNat_minus";
+
+Goal "inverse ((*fNat* f) x) = (*fNat* (%x. inverse (f x))) x";
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat, hypreal_inverse]));
+qed "starfunNat_inverse";
+
+(*--------------------------------------------------------
+   extented function has same solution as its standard
+   version for natural arguments. i.e they are the same
+   for all natural arguments (c.f. Hoskins pg. 107- SEQ)
+ -------------------------------------------------------*)
+
+Goal "(*fNat* f) (hypnat_of_nat a) = hypreal_of_real (f a)";
+by (auto_tac (claset(),
+      simpset() addsimps [starfunNat,hypnat_of_nat_def,hypreal_of_real_def]));
+qed "starfunNat_eq";
+
+Addsimps [starfunNat_eq];
+
+Goal "(*fNat2* f) (hypnat_of_nat a) = hypnat_of_nat (f a)";
+by (auto_tac (claset(), simpset() addsimps [starfunNat2,hypnat_of_nat_def]));
+qed "starfunNat2_eq";
+
+Addsimps [starfunNat2_eq];
+
+Goal "(*fNat* f) (hypnat_of_nat a) @= hypreal_of_real (f a)";
+by (Auto_tac);
+qed "starfunNat_inf_close";
+
+
+(*-----------------------------------------------------------------
+    Example of transfer of a property from reals to hyperreals
+    --- used for limit comparison of sequences
+ ----------------------------------------------------------------*)
+Goal "ALL n. N <= n --> f n <= g n \
+\         ==> ALL n. hypnat_of_nat N <= n --> (*fNat* f) n <= (*fNat* g) n";
+by (Step_tac 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+         simpset() addsimps [starfunNat, hypnat_of_nat_def,hypreal_le,
+                             hypreal_less, hypnat_le,hypnat_less]));
+by (Ultra_tac 1);
+by Auto_tac;
+qed "starfun_le_mono";
+
+(*****----- and another -----*****) 
+Goal "ALL n. N <= n --> f n < g n \
+\         ==> ALL n. hypnat_of_nat N <= n --> (*fNat* f) n < (*fNat* g) n";
+by (Step_tac 1);
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+         simpset() addsimps [starfunNat, hypnat_of_nat_def,hypreal_le,
+                             hypreal_less, hypnat_le,hypnat_less]));
+by (Ultra_tac 1);
+by Auto_tac;
+qed "starfun_less_mono";
+
+(*----------------------------------------------------------------
+            NS extension when we displace argument by one
+ ---------------------------------------------------------------*)
+Goal "(*fNat* (%n. f (Suc n))) N = (*fNat* f) (N + 1hn)";
+by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+         simpset() addsimps [starfunNat, hypnat_one_def,hypnat_add]));
+qed "starfunNat_shift_one";
+
+(*----------------------------------------------------------------
+                 NS extension with rabs    
+ ---------------------------------------------------------------*)
+Goal "(*fNat* (%n. abs (f n))) N = abs((*fNat* f) N)";
+by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat, hypreal_hrabs]));
+qed "starfunNat_rabs";
+
+(*----------------------------------------------------------------
+       The hyperpow function as a NS extension of realpow
+ ----------------------------------------------------------------*)
+Goal "(*fNat* (%n. r ^ n)) N = (hypreal_of_real r) pow N";
+by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+         simpset() addsimps [hyperpow, hypreal_of_real_def,starfunNat]));
+qed "starfunNat_pow";
+
+Goal "(*fNat* (%n. (X n) ^ m)) N = (*fNat* X) N pow hypnat_of_nat m";
+by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),
+         simpset() addsimps [hyperpow, hypnat_of_nat_def,starfunNat]));
+qed "starfunNat_pow2";
+
+Goalw [hypnat_of_nat_def] "(*f* (%r. r ^ n)) R = (R) pow hypnat_of_nat n";
+by (res_inst_tac [("z","R")] eq_Abs_hypreal 1);
+by (auto_tac (claset(), simpset() addsimps [hyperpow,starfun]));
+qed "starfun_pow";
+
+(*----------------------------------------------------- 
+   hypreal_of_hypnat as NS extension of real_of_nat! 
+-------------------------------------------------------*)
+Goal "(*fNat* real_of_nat) = hypreal_of_hypnat";
+by (rtac ext 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [hypreal_of_hypnat,starfunNat]));
+qed "starfunNat_real_of_nat";
+
+Goal "N : HNatInfinite \
+\  ==> (*fNat* (%x. inverse(real_of_nat x))) N = inverse(hypreal_of_hypnat N)";
+by (res_inst_tac [("f1","inverse")]  (starfun_stafunNat_o2 RS subst) 1);
+by (subgoal_tac "hypreal_of_hypnat N ~= #0" 1);
+by (auto_tac (claset(), 
+       simpset() addsimps [starfunNat_real_of_nat, starfun_inverse_inverse]));
+qed "starfunNat_inverse_real_of_nat_eq";
+
+(*----------------------------------------------------------
+     Internal functions - some redundancy with *fNat* now
+ ---------------------------------------------------------*)
+Goalw [congruent_def] 
+      "congruent hypnatrel (%X. hypnatrel^^{%n. f n (X n)})";
+by (safe_tac (claset()));
+by (ALLGOALS(Fuf_tac));
+qed "starfunNat_n_congruent";
+
+Goalw [starfunNat_n_def]
+     "(*fNatn* f) (Abs_hypnat(hypnatrel^^{%n. X n})) = \
+\     Abs_hypreal(hyprel ^^ {%n. f n (X n)})";
+by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
+by Auto_tac;
+by (Ultra_tac 1);
+qed "starfunNat_n";
+
+(*-------------------------------------------------
+  multiplication: ( *fn ) x ( *gn ) = *(fn x gn)  
+ -------------------------------------------------*)
+Goal "(*fNatn* f) z * (*fNatn* g) z = (*fNatn* (% i x. f i x * g i x)) z";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat_n,hypreal_mult]));
+qed "starfunNat_n_mult";
+
+(*-----------------------------------------------
+  addition: ( *fn ) + ( *gn ) = *(fn + gn)  
+ -----------------------------------------------*)
+Goal "(*fNatn* f) z + (*fNatn* g) z = (*fNatn* (%i x. f i x + g i x)) z";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat_n,hypreal_add]));
+qed "starfunNat_n_add";
+
+(*-------------------------------------------------
+  subtraction: ( *fn ) + -( *gn ) = *(fn + -gn)  
+ -------------------------------------------------*)
+Goal "(*fNatn* f) z + -(*fNatn* g) z = (*fNatn* (%i x. f i x + -g i x)) z";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), 
+          simpset() addsimps [starfunNat_n, hypreal_minus,hypreal_add]));
+qed "starfunNat_n_add_minus";
+
+(*--------------------------------------------------
+  composition: ( *fn ) o ( *gn ) = *(fn o gn)  
+ -------------------------------------------------*)
+ 
+Goal "(*fNatn* (%i x. k)) z = hypreal_of_real  k";
+by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), 
+       simpset() addsimps [starfunNat_n, hypreal_of_real_def]));
+qed "starfunNat_n_const_fun";
+
+Addsimps [starfunNat_n_const_fun];
+
+Goal "- (*fNatn* f) x = (*fNatn* (%i x. - (f i) x)) x";
+by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat_n, hypreal_minus]));
+qed "starfunNat_n_minus";
+
+Goal "(*fNatn* f) (hypnat_of_nat n) = Abs_hypreal(hyprel ^^ {%i. f i n})";
+by (auto_tac (claset(), simpset() addsimps [starfunNat_n,hypnat_of_nat_def]));
+qed "starfunNat_n_eq";
+Addsimps [starfunNat_n_eq];
+
+Goal "((*fNat* f) = (*fNat* g)) = (f = g)";
+by Auto_tac;
+by (rtac ext 1 THEN rtac ccontr 1);
+by (dres_inst_tac [("x","hypnat_of_nat(x)")] fun_cong 1);
+by (auto_tac (claset(), simpset() addsimps [starfunNat,hypnat_of_nat_def]));
+qed "starfun_eq_iff";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/NatStar.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,46 @@
+(*  Title       : NatStar.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : defining *-transforms in NSA which extends 
+                  sets of reals, and nat=>real, nat=>nat functions
+*) 
+
+NatStar = HyperNat + RealPow + HyperPow + 
+
+constdefs
+
+  (* internal sets and nonstandard extensions -- see Star.thy as well *)
+
+    starsetNat :: nat set => hypnat set          ("*sNat* _" [80] 80)
+    "*sNat* A  == {x. ALL X: Rep_hypnat(x). {n::nat. X n : A}: FreeUltrafilterNat}"
+
+    starsetNat_n :: (nat => nat set) => hypnat set        ("*sNatn* _" [80] 80)
+    "*sNatn* As  == {x. ALL X: Rep_hypnat(x). {n::nat. X n : (As n)}: FreeUltrafilterNat}"   
+
+    InternalNatSets :: "hypnat set set"
+    "InternalNatSets == {X. EX As. X = *sNatn* As}"
+
+    (* star transform of functions f:Nat --> Real *)
+
+    starfunNat :: (nat => real) => hypnat => hypreal        ("*fNat* _" [80] 80)
+    "*fNat* f  == (%x. Abs_hypreal(UN X: Rep_hypnat(x). hyprel^^{%n. f (X n)}))" 
+
+    starfunNat_n :: (nat => (nat => real)) => hypnat => hypreal        ("*fNatn* _" [80] 80)
+    "*fNatn* F  == (%x. Abs_hypreal(UN X: Rep_hypnat(x). hyprel^^{%n. (F n)(X n)}))" 
+
+    InternalNatFuns :: (hypnat => hypreal) set
+    "InternalNatFuns == {X. EX F. X = *fNatn* F}"
+
+    (* star transform of functions f:Nat --> Nat *)
+
+    starfunNat2 :: (nat => nat) => hypnat => hypnat        ("*fNat2* _" [80] 80)
+    "*fNat2* f  == (%x. Abs_hypnat(UN X: Rep_hypnat(x). hypnatrel^^{%n. f (X n)}))" 
+
+    starfunNat2_n :: (nat => (nat => nat)) => hypnat => hypnat        ("*fNat2n* _" [80] 80)
+    "*fNat2n* F  == (%x. Abs_hypnat(UN X: Rep_hypnat(x). hypnatrel^^{%n. (F n)(X n)}))" 
+
+    InternalNatFuns2 :: (hypnat => hypnat) set
+    "InternalNatFuns2 == {X. EX F. X = *fNat2n* F}"
+end
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/ROOT.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,11 @@
+(*  Title:      HOL/Hyperreal/ROOT.ML
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1998  University of Cambridge
+
+Construction of the Hyperreals by the Ultrapower Construction
+and mechanization of Nonstandard Real Analysis
+by Jacques Fleuriot
+*)
+
+time_use_thy "Hyperreal";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/SEQ.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,1368 @@
+(*  Title       : SEQ.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Theory of sequence and series of real numbers
+*) 
+
+(*---------------------------------------------------------------------------
+   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 
+ -------------------------------------------------------------------------- *)
+
+Goalw [hypnat_omega_def] 
+      "(*fNat* (%n::nat. inverse(real_of_posnat n))) whn : Infinitesimal";
+by (auto_tac (claset(),
+      simpset() addsimps [Infinitesimal_FreeUltrafilterNat_iff,starfunNat]));
+by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
+by (auto_tac (claset(),
+       simpset() addsimps (map rename_numerals) 
+         [real_of_posnat_gt_zero,real_inverse_gt_zero,abs_eqI2,
+          FreeUltrafilterNat_inverse_real_of_posnat]));
+qed "SEQ_Infinitesimal";
+
+(*--------------------------------------------------------------------------
+                  Rules for LIMSEQ and NSLIMSEQ etc.
+ --------------------------------------------------------------------------*)
+
+(*** LIMSEQ ***)
+Goalw [LIMSEQ_def] 
+      "X ----> L ==> \
+\      ALL r. #0 < r --> (EX no. ALL n. no <= n --> abs(X n + -L) < r)";
+by (Asm_simp_tac 1);
+qed "LIMSEQD1";
+
+Goalw [LIMSEQ_def] 
+      "[| X ----> L; #0 < r|] ==> \
+\      EX no. ALL n. no <= n --> abs(X n + -L) < r";
+by (Asm_simp_tac 1);
+qed "LIMSEQD2";
+
+Goalw [LIMSEQ_def] 
+      "ALL r. #0 < r --> (EX no. ALL n. \
+\      no <= n --> abs(X n + -L) < r) ==> X ----> L";
+by (Asm_simp_tac 1);
+qed "LIMSEQI";
+
+Goalw [LIMSEQ_def] 
+      "(X ----> L) = \
+\      (ALL r. #0 <r --> (EX no. ALL n. no <= n --> abs(X n + -L) < r))";
+by (Simp_tac 1);
+qed "LIMSEQ_iff";
+
+(*** NSLIMSEQ ***)
+Goalw [NSLIMSEQ_def] 
+     "X ----NS> L ==> ALL N: HNatInfinite. (*fNat* X) N @= hypreal_of_real L";
+by (Asm_simp_tac 1);
+qed "NSLIMSEQD1";
+
+Goalw [NSLIMSEQ_def] 
+    "[| X ----NS> L; N: HNatInfinite |] ==> (*fNat* X) N @= hypreal_of_real L";
+by (Asm_simp_tac 1);
+qed "NSLIMSEQD2";
+
+Goalw [NSLIMSEQ_def] 
+     "ALL N: HNatInfinite. (*fNat* X) N @= hypreal_of_real L ==> X ----NS> L";
+by (Asm_simp_tac 1);
+qed "NSLIMSEQI";
+
+Goalw [NSLIMSEQ_def] 
+    "(X ----NS> L) = (ALL N: HNatInfinite. (*fNat* X) N @= hypreal_of_real L)";
+by (Simp_tac 1);
+qed "NSLIMSEQ_iff";
+
+(*----------------------------------------
+          LIMSEQ ==> NSLIMSEQ
+ ---------------------------------------*)
+Goalw [LIMSEQ_def,NSLIMSEQ_def] 
+      "X ----> L ==> X ----NS> L";
+by (auto_tac (claset(),simpset() addsimps 
+    [HNatInfinite_FreeUltrafilterNat_iff]));
+by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
+by (rtac (inf_close_minus_iff RS iffD2) 1);
+by (auto_tac (claset(),simpset() addsimps [starfunNat,
+    mem_infmal_iff RS sym,hypreal_of_real_def,
+    hypreal_minus,hypreal_add,
+    Infinitesimal_FreeUltrafilterNat_iff]));
+by (EVERY[rtac bexI 1, rtac lemma_hyprel_refl 2, Step_tac 1]);
+by (dres_inst_tac [("x","u")] spec 1 THEN Step_tac 1);
+by (dres_inst_tac [("x","no")] spec 1);
+by (Fuf_tac 1);
+by (blast_tac (claset() addDs [less_imp_le]) 1);
+qed "LIMSEQ_NSLIMSEQ";
+
+(*-------------------------------------------------------------
+          NSLIMSEQ ==> LIMSEQ
+    proving NS def ==> Standard def is trickier as usual 
+ -------------------------------------------------------------*)
+(* the following sequence f(n) defines a hypernatural *)
+(* lemmas etc. first *)
+Goal "!!(f::nat=>nat). ALL n. n <= f n \
+\          ==> {n. f n = 0} = {0} | {n. f n = 0} = {}";
+by (Auto_tac);
+by (dres_inst_tac [("x","xa")] spec 1);
+by (dres_inst_tac [("x","x")] spec 2);
+by (Auto_tac);
+val lemma_NSLIMSEQ1 = result();
+
+Goal "{n. f n <= Suc u} = {n. f n <= u} Un {n. f n = Suc u}";
+by (auto_tac (claset(),simpset() addsimps [le_Suc_eq]));
+val lemma_NSLIMSEQ2 = result();
+
+Goal "!!(f::nat=>nat). ALL n. n <= f n \
+\          ==> {n. f n = Suc u} <= {n. n <= Suc u}";
+by (Auto_tac);
+by (dres_inst_tac [("x","x")] spec 1);
+by (Auto_tac);
+val lemma_NSLIMSEQ3 = result();
+
+Goal "!!(f::nat=>nat). ALL n. n <= f n \ 
+\         ==> finite {n. f n <= u}";
+by (induct_tac "u" 1);
+by (auto_tac (claset(),simpset() addsimps [lemma_NSLIMSEQ2]));
+by (auto_tac (claset() addIs [(lemma_NSLIMSEQ3 RS finite_subset),
+    finite_nat_le_segment], simpset()));
+by (dtac lemma_NSLIMSEQ1 1 THEN Step_tac 1);
+by (ALLGOALS(Asm_simp_tac));
+qed "NSLIMSEQ_finite_set";
+
+Goal "- {n. u < (f::nat=>nat) n} = {n. f n <= u}";
+by (auto_tac (claset() addDs [less_le_trans],
+    simpset() addsimps [le_def]));
+qed "Compl_less_set";
+
+(* the index set is in the free ultrafilter *)
+Goal "!!(f::nat=>nat). ALL n. n <= f n \ 
+\         ==> {n. u < f n} : FreeUltrafilterNat";
+by (rtac (FreeUltrafilterNat_Compl_iff2 RS iffD2) 1);
+by (rtac FreeUltrafilterNat_finite 1);
+by (auto_tac (claset() addDs [NSLIMSEQ_finite_set],
+    simpset() addsimps [Compl_less_set]));
+qed "FreeUltrafilterNat_NSLIMSEQ";
+
+(* thus, the sequence defines an infinite hypernatural! *)
+Goal "ALL n. n <= f n \
+\         ==> Abs_hypnat (hypnatrel ^^ {f}) : HNatInfinite";
+by (auto_tac (claset(),simpset() addsimps [HNatInfinite_FreeUltrafilterNat_iff]));
+by (EVERY[rtac bexI 1, rtac lemma_hypnatrel_refl 2, Step_tac 1]);
+by (etac FreeUltrafilterNat_NSLIMSEQ 1);
+qed "HNatInfinite_NSLIMSEQ";
+
+val lemmaLIM = CLAIM  "{n. X (f n) + - L = Y n} Int {n. abs (Y n) < r} <= \
+\         {n. abs (X (f n) + - L) < r}";
+
+Goal "{n. abs (X (f n) + - L) < r} Int {n. r <= abs (X (f n) + - (L::real))} \
+\     = {}";
+by Auto_tac;  
+val lemmaLIM2 = result();
+
+Goal "[| #0 < r; ALL n. r <= abs (X (f n) + - L); \
+\          (*fNat* X) (Abs_hypnat (hypnatrel ^^ {f})) + \
+\          - hypreal_of_real  L @= 0 |] ==> False";
+by (auto_tac (claset(),simpset() addsimps [starfunNat,
+    mem_infmal_iff RS sym,hypreal_of_real_def,
+    hypreal_minus,hypreal_add,
+    Infinitesimal_FreeUltrafilterNat_iff]));
+by (dres_inst_tac [("x","r")] spec 1 THEN Step_tac 1);
+by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
+by (dtac (lemmaLIM RSN (2,FreeUltrafilterNat_subset)) 1);
+by (dtac FreeUltrafilterNat_all 1);
+by (thin_tac "{n. abs (Y n) < r} : FreeUltrafilterNat" 1);
+by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [lemmaLIM2,
+                                  FreeUltrafilterNat_empty]) 1);
+val lemmaLIM3 = result();
+
+Goalw [LIMSEQ_def,NSLIMSEQ_def] 
+      "X ----NS> L ==> X ----> L";
+by (rtac ccontr 1 THEN Asm_full_simp_tac 1);
+by (Step_tac 1);
+(* skolemization step *)
+by (dtac choice 1 THEN Step_tac 1);
+by (dres_inst_tac [("x","Abs_hypnat(hypnatrel^^{f})")] bspec 1);
+by (dtac (inf_close_minus_iff RS iffD1) 2);
+by (fold_tac [real_le_def]);
+by (blast_tac (claset() addIs [HNatInfinite_NSLIMSEQ]) 1);
+by (blast_tac (claset() addIs [rename_numerals lemmaLIM3]) 1);
+qed "NSLIMSEQ_LIMSEQ";
+
+(* Now the all important result is trivially proved! *)
+Goal "(f ----> L) = (f ----NS> L)";
+by (blast_tac (claset() addIs [LIMSEQ_NSLIMSEQ,NSLIMSEQ_LIMSEQ]) 1);
+qed "LIMSEQ_NSLIMSEQ_iff";
+
+(*-------------------------------------------------------------------
+                   Theorems about sequences
+ ------------------------------------------------------------------*)
+Goalw [NSLIMSEQ_def] "(%n. k) ----NS> k";
+by (Auto_tac);
+qed "NSLIMSEQ_const";
+
+Goalw [LIMSEQ_def] "(%n. k) ----> k";
+by (Auto_tac);
+qed "LIMSEQ_const";
+
+Goalw [NSLIMSEQ_def]
+      "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n + Y n) ----NS> a + b";
+by (auto_tac (claset() addIs [inf_close_add],
+    simpset() addsimps [starfunNat_add RS sym]));
+qed "NSLIMSEQ_add";
+
+Goal "[| X ----> a; Y ----> b |] ==> (%n. X n + Y n) ----> a + b";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
+                                           NSLIMSEQ_add]) 1);
+qed "LIMSEQ_add";
+
+Goalw [NSLIMSEQ_def]
+      "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n * Y n) ----NS> a * b";
+by (auto_tac (claset() addSIs [inf_close_mult_HFinite],
+    simpset() addsimps [hypreal_of_real_mult, starfunNat_mult RS sym]));
+qed "NSLIMSEQ_mult";
+
+Goal "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
+                                           NSLIMSEQ_mult]) 1);
+qed "LIMSEQ_mult";
+
+Goalw [NSLIMSEQ_def] 
+     "X ----NS> a ==> (%n. -(X n)) ----NS> -a";
+by (auto_tac (claset(), simpset() addsimps [starfunNat_minus RS sym]));
+qed "NSLIMSEQ_minus";
+
+Goal "X ----> a ==> (%n. -(X n)) ----> -a";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
+                                           NSLIMSEQ_minus]) 1);
+qed "LIMSEQ_minus";
+
+Goal "(%n. -(X n)) ----> -a ==> X ----> a";
+by (dtac LIMSEQ_minus 1);
+by (Asm_full_simp_tac 1);
+qed "LIMSEQ_minus_cancel";
+
+Goal "(%n. -(X n)) ----NS> -a ==> X ----NS> a";
+by (dtac NSLIMSEQ_minus 1);
+by (Asm_full_simp_tac 1);
+qed "NSLIMSEQ_minus_cancel";
+
+Goal "[| X ----NS> a; Y ----NS> b |] \
+\               ==> (%n. X n + -Y n) ----NS> a + -b";
+by (dres_inst_tac [("X","Y")] NSLIMSEQ_minus 1);
+by (auto_tac (claset(),simpset() addsimps [NSLIMSEQ_add]));
+qed "NSLIMSEQ_add_minus";
+
+Goal "[| X ----> a; Y ----> b |] \
+\               ==> (%n. X n + -Y n) ----> a + -b";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
+    NSLIMSEQ_add_minus]) 1);
+qed "LIMSEQ_add_minus";
+
+Goalw [real_diff_def] 
+     "[| X ----> a; Y ----> b |] ==> (%n. X n - Y n) ----> a - b";
+by (blast_tac (claset() addIs [LIMSEQ_add_minus]) 1);
+qed "LIMSEQ_diff";
+
+Goalw [real_diff_def] 
+     "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n - Y n) ----NS> a - b";
+by (blast_tac (claset() addIs [NSLIMSEQ_add_minus]) 1);
+qed "NSLIMSEQ_diff";
+
+(*---------------------------------------------------------------
+    Proof is like that of NSLIM_inverse.
+ --------------------------------------------------------------*)
+Goalw [NSLIMSEQ_def] 
+     "[| X ----NS> a;  a ~= #0 |] ==> (%n. inverse(X n)) ----NS> inverse(a)";
+by (Clarify_tac 1);
+by (dtac bspec 1);
+by (auto_tac (claset(), 
+              simpset() addsimps [starfunNat_inverse RS sym, 
+                                  hypreal_of_real_inf_close_inverse]));  
+qed "NSLIMSEQ_inverse";
+
+
+(*------ Standard version of theorem -------*)
+Goal "[| X ----> a; a ~= #0 |] ==> (%n. inverse(X n)) ----> inverse(a)";
+by (asm_full_simp_tac (simpset() addsimps [NSLIMSEQ_inverse,
+    LIMSEQ_NSLIMSEQ_iff]) 1);
+qed "LIMSEQ_inverse";
+
+Goal "[| X ----NS> a;  Y ----NS> b;  b ~= #0 |] \
+\     ==> (%n. X n / Y n) ----NS> a/b";
+by (asm_full_simp_tac (simpset() addsimps [NSLIMSEQ_mult, NSLIMSEQ_inverse, 
+                                           real_divide_def]) 1);
+qed "NSLIMSEQ_mult_inverse";
+
+Goal "[| X ----> a;  Y ----> b;  b ~= #0 |] ==> (%n. X n / Y n) ----> a/b";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_mult, LIMSEQ_inverse, 
+                                           real_divide_def]) 1);
+qed "LIMSEQ_divide";
+
+(*-----------------------------------------------
+            Uniqueness of limit
+ ----------------------------------------------*)
+Goalw [NSLIMSEQ_def] 
+     "[| X ----NS> a; X ----NS> b |] ==> a = b";
+by (REPEAT(dtac (HNatInfinite_whn RSN (2,bspec)) 1));
+by (auto_tac (claset() addDs [inf_close_trans3], simpset()));
+qed "NSLIMSEQ_unique";
+
+Goal "[| X ----> a; X ----> b |] ==> a = b";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
+    NSLIMSEQ_unique]) 1);
+qed "LIMSEQ_unique";
+
+(*-----------------------------------------------------------------
+    theorems about nslim and lim
+ ----------------------------------------------------------------*)
+Goalw [lim_def] "X ----> L ==> lim X = L";
+by (blast_tac (claset() addIs [LIMSEQ_unique]) 1);
+qed "limI";
+
+Goalw [nslim_def] "X ----NS> L ==> nslim X = L";
+by (blast_tac (claset() addIs [NSLIMSEQ_unique]) 1);
+qed "nslimI";
+
+Goalw [lim_def,nslim_def] "lim X = nslim X";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff]) 1);
+qed "lim_nslim_iff";
+
+(*------------------------------------------------------------------
+                      Convergence
+ -----------------------------------------------------------------*)
+Goalw [convergent_def]
+     "convergent X ==> EX L. (X ----> L)";
+by (assume_tac 1);
+qed "convergentD";
+
+Goalw [convergent_def]
+     "(X ----> L) ==> convergent X";
+by (Blast_tac 1);
+qed "convergentI";
+
+Goalw [NSconvergent_def]
+     "NSconvergent X ==> EX L. (X ----NS> L)";
+by (assume_tac 1);
+qed "NSconvergentD";
+
+Goalw [NSconvergent_def]
+     "(X ----NS> L) ==> NSconvergent X";
+by (Blast_tac 1);
+qed "NSconvergentI";
+
+Goalw [convergent_def,NSconvergent_def] 
+     "convergent X = NSconvergent X";
+by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff]) 1);
+qed "convergent_NSconvergent_iff";
+
+Goalw [NSconvergent_def,nslim_def] 
+     "NSconvergent X = (X ----NS> nslim X)";
+by (auto_tac (claset() addIs [someI], simpset()));
+qed "NSconvergent_NSLIMSEQ_iff";
+
+Goalw [convergent_def,lim_def] 
+     "convergent X = (X ----> lim X)";
+by (auto_tac (claset() addIs [someI], simpset()));
+qed "convergent_LIMSEQ_iff";
+
+(*-------------------------------------------------------------------
+         Subsequence (alternative definition) (e.g. Hoskins)
+ ------------------------------------------------------------------*)
+Goalw [subseq_def] "subseq f = (ALL n. (f n) < (f (Suc n)))";
+by (auto_tac (claset() addSDs [less_imp_Suc_add], simpset()));
+by (nat_ind_tac "k" 1);
+by (auto_tac (claset() addIs [less_trans], simpset()));
+qed "subseq_Suc_iff";
+
+(*-------------------------------------------------------------------
+                   Monotonicity
+ ------------------------------------------------------------------*)
+
+Goalw [monoseq_def]
+   "monoseq X = ((ALL n. X n <= X (Suc n)) \
+\                | (ALL n. X (Suc n) <= X n))";
+by (auto_tac (claset () addSDs [le_imp_less_or_eq], simpset()));
+by (auto_tac (claset() addSIs [lessI RS less_imp_le]
+                       addSDs [less_imp_Suc_add], 
+    simpset()));
+by (induct_tac "ka" 1);
+by (auto_tac (claset() addIs [order_trans], simpset()));
+by (EVERY1[rtac ccontr, rtac swap, Simp_tac]);
+by (induct_tac "k" 1);
+by (auto_tac (claset() addIs [order_trans], simpset()));
+qed "monoseq_Suc";
+
+Goalw [monoseq_def] 
+       "ALL m n. m <= n --> X m <= X n ==> monoseq X";
+by (Blast_tac 1);
+qed "monoI1";
+
+Goalw [monoseq_def] 
+       "ALL m n. m <= n --> X n <= X m ==> monoseq X";
+by (Blast_tac 1);
+qed "monoI2";
+
+Goal "ALL n. X n <= X (Suc n) ==> monoseq X";
+by (asm_simp_tac (simpset() addsimps [monoseq_Suc]) 1);
+qed "mono_SucI1";
+
+Goal "ALL n. X (Suc n) <= X n ==> monoseq X";
+by (asm_simp_tac (simpset() addsimps [monoseq_Suc]) 1);
+qed "mono_SucI2";
+
+(*-------------------------------------------------------------------
+                  Bounded Sequence
+ ------------------------------------------------------------------*)
+Goalw [Bseq_def] 
+      "Bseq X ==> EX K. #0 < K & (ALL n. abs(X n) <= K)";
+by (assume_tac 1);
+qed "BseqD";
+
+Goalw [Bseq_def]
+      "[| #0 < K; ALL n. abs(X n) <= K |] \
+\           ==> Bseq X";
+by (Blast_tac 1);
+qed "BseqI";
+
+Goal "(EX K. #0 < K & (ALL n. abs(X n) <= K)) = \
+\         (EX N. ALL n. abs(X n) <= real_of_posnat N)";
+by (auto_tac (claset(),simpset() addsimps 
+    (map rename_numerals) [real_gt_zero_preal_Ex,real_of_posnat_gt_zero]));
+by (cut_inst_tac [("x","real_of_preal y")] reals_Archimedean2 1);
+by (blast_tac (claset() addIs [order_le_less_trans, order_less_imp_le]) 1);
+by (auto_tac (claset(),simpset() addsimps [real_of_posnat_def]));
+qed "lemma_NBseq_def";
+
+(* alternative definition for Bseq *)
+Goalw [Bseq_def] 
+      "Bseq X = (EX N. ALL n. abs(X n) <= real_of_posnat N)";
+by (simp_tac (simpset() addsimps [lemma_NBseq_def]) 1);
+qed "Bseq_iff";
+
+Goal "(EX K. #0 < K & (ALL n. abs(X n) <= K)) = \
+\     (EX N. ALL n. abs(X n) < real_of_posnat N)";
+by (auto_tac (claset(),simpset() addsimps 
+    (map rename_numerals) [real_gt_zero_preal_Ex,real_of_posnat_gt_zero]));
+by (cut_inst_tac [("x","real_of_preal y")] reals_Archimedean2 1);
+by (blast_tac (claset() addIs [order_less_trans, order_le_less_trans]) 1);
+by (auto_tac (claset() addIs [order_less_imp_le],
+              simpset() addsimps [real_of_posnat_def]));
+qed "lemma_NBseq_def2";
+
+(* yet another definition for Bseq *)
+Goalw [Bseq_def] 
+      "Bseq X = (EX N. ALL n. abs(X n) < real_of_posnat N)";
+by (simp_tac (simpset() addsimps [lemma_NBseq_def2]) 1);
+qed "Bseq_iff1a";
+
+Goalw [NSBseq_def]
+      "[| NSBseq X; N: HNatInfinite |] \
+\           ==> (*fNat* X) N : HFinite";
+by (Blast_tac 1);
+qed "NSBseqD";
+
+Goalw [NSBseq_def]
+      "ALL N: HNatInfinite. (*fNat* X) N : HFinite \
+\           ==> NSBseq X";
+by (assume_tac 1);
+qed "NSBseqI";
+
+(*-----------------------------------------------------------
+       Standard definition ==> NS definition
+ ----------------------------------------------------------*)
+(* a few lemmas *)
+Goal "ALL n. abs(X n) <= K ==> \
+\     ALL n. abs(X((f::nat=>nat) n)) <= K";
+by (Auto_tac);
+val lemma_Bseq = result();
+
+Goalw [Bseq_def,NSBseq_def] "Bseq X ==> NSBseq X";
+by (Step_tac 1);
+by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),simpset() addsimps [starfunNat,    
+    HFinite_FreeUltrafilterNat_iff,
+    HNatInfinite_FreeUltrafilterNat_iff]));
+by (EVERY[rtac bexI 1, rtac lemma_hyprel_refl 2]);
+by (dres_inst_tac [("f","Xa")] lemma_Bseq 1); 
+by (res_inst_tac [("x","K+#1")] exI 1);
+by (rotate_tac 2 1 THEN dtac FreeUltrafilterNat_all 1);
+by (Ultra_tac 1);
+qed "Bseq_NSBseq";
+
+(*---------------------------------------------------------------
+       NS  definition ==> Standard definition
+ ---------------------------------------------------------------*)
+(* similar to NSLIM proof in REALTOPOS *)
+(*------------------------------------------------------------------- 
+   We need to get rid of the real variable and do so by proving the
+   following which relies on the Archimedean property of the reals
+   When we skolemize we then get the required function f::nat=>nat 
+   o/w we would be stuck with a skolem function f :: real=>nat which
+   is not what we want (read useless!)
+ -------------------------------------------------------------------*)
+ 
+Goal "ALL K. #0 < K --> (EX n. K < abs (X n)) \
+\          ==> ALL N. EX n. real_of_posnat  N < abs (X n)";
+by (Step_tac 1);
+by (cut_inst_tac [("n","N")] (rename_numerals real_of_posnat_gt_zero) 1);
+by (Blast_tac 1);
+val lemmaNSBseq = result();
+
+Goal "ALL K. #0 < K --> (EX n. K < abs (X n)) \
+\         ==> EX f. ALL N. real_of_posnat  N < abs (X (f N))";
+by (dtac lemmaNSBseq 1);
+by (dtac choice 1);
+by (Blast_tac 1);
+val lemmaNSBseq2 = result();
+
+Goal "ALL N. real_of_posnat  N < abs (X (f N)) \
+\         ==>  Abs_hypreal(hyprel^^{X o f}) : HInfinite";
+by (auto_tac (claset(),simpset() addsimps 
+    [HInfinite_FreeUltrafilterNat_iff,o_def]));
+by (EVERY[rtac bexI 1, rtac lemma_hyprel_refl 2, 
+    Step_tac 1]);
+by (cut_inst_tac [("u","u")] FreeUltrafilterNat_nat_gt_real 1);
+by (blast_tac (claset() addDs [FreeUltrafilterNat_all, FreeUltrafilterNat_Int] 
+                        addIs [order_less_trans, FreeUltrafilterNat_subset]) 1);
+qed "real_seq_to_hypreal_HInfinite";
+
+(*--------------------------------------------------------------------------------
+     Now prove that we can get out an infinite hypernatural as well 
+     defined using the skolem function f::nat=>nat above
+ --------------------------------------------------------------------------------*)
+
+Goal "{n. f n <= Suc u & real_of_posnat  n < abs (X (f n))} <= \
+\         {n. f n <= u & real_of_posnat  n < abs (X (f n))} \
+\         Un {n. real_of_posnat n < abs (X (Suc u))}";
+by (auto_tac (claset() addSDs [le_imp_less_or_eq] addIs [less_imp_le],
+    simpset() addsimps [less_Suc_eq]));
+val lemma_finite_NSBseq = result();
+
+Goal "finite {n. f n <= (u::nat) &  real_of_posnat n < abs(X(f n))}";
+by (induct_tac "u" 1);
+by (rtac (CLAIM "{n. f n <= (0::nat) & real_of_posnat n < abs (X (f n))} <= \
+\         {n. real_of_posnat n < abs (X 0)}"
+          RS finite_subset) 1);
+by (rtac finite_real_of_posnat_less_real 1);
+by (rtac (lemma_finite_NSBseq RS finite_subset) 1);
+by (auto_tac (claset() addIs [finite_real_of_posnat_less_real], simpset()));
+val lemma_finite_NSBseq2 = result();
+
+Goal "ALL N. real_of_posnat  N < abs (X (f N)) \
+\     ==> Abs_hypnat(hypnatrel^^{f}) : HNatInfinite";
+by (auto_tac (claset(),simpset() addsimps 
+    [HNatInfinite_FreeUltrafilterNat_iff]));
+by (EVERY[rtac bexI 1, rtac lemma_hypnatrel_refl 2, 
+    Step_tac 1]);
+by (rtac ccontr 1 THEN dtac FreeUltrafilterNat_Compl_mem 1);
+by (asm_full_simp_tac (simpset() addsimps 
+   [CLAIM_SIMP "- {n. u < (f::nat=>nat) n} \
+\   = {n. f n <= u}" [le_def]]) 1);
+by (dtac FreeUltrafilterNat_all 1);
+by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [CLAIM "({n. f n <= u} Int {n. real_of_posnat n < abs(X(f n))}) = \
+\          {n. f n <= (u::nat) &  real_of_posnat n < abs(X(f n))}",
+     lemma_finite_NSBseq2 RS FreeUltrafilterNat_finite]));
+qed "HNatInfinite_skolem_f";
+
+Goalw [Bseq_def,NSBseq_def] 
+      "NSBseq X ==> Bseq X";
+by (rtac ccontr 1);
+by (auto_tac (claset(),simpset() addsimps [real_le_def]));
+by (dtac lemmaNSBseq2 1 THEN Step_tac 1);
+by (forw_inst_tac [("X","X"),("f","f")] real_seq_to_hypreal_HInfinite 1);
+by (dtac (HNatInfinite_skolem_f RSN (2,bspec)) 1 THEN assume_tac 1);
+by (auto_tac (claset(),simpset() addsimps [starfunNat,
+    o_def,HFinite_HInfinite_iff]));
+qed "NSBseq_Bseq";
+
+(*----------------------------------------------------------------------
+  Equivalence of nonstandard and standard definitions 
+  for a bounded sequence
+ -----------------------------------------------------------------------*)
+Goal "(Bseq X) = (NSBseq X)";
+by (blast_tac (claset() addSIs [NSBseq_Bseq,Bseq_NSBseq]) 1);
+qed "Bseq_NSBseq_iff";
+
+(*----------------------------------------------------------------------
+   A convergent sequence is bounded
+   (Boundedness as a necessary condition for convergence)
+ -----------------------------------------------------------------------*)
+(* easier --- nonstandard version - no existential as usual *)
+Goalw [NSconvergent_def,NSBseq_def,NSLIMSEQ_def] 
+          "NSconvergent X ==> NSBseq X";
+by (blast_tac (claset() addDs [HFinite_hypreal_of_real RS 
+               (inf_close_sym RSN (2,inf_close_HFinite))]) 1);
+qed "NSconvergent_NSBseq";
+
+(* standard version - easily now proved using *)
+(* equivalence of NS and standard definitions *)
+Goal "convergent X ==> Bseq X";
+by (asm_full_simp_tac (simpset() addsimps [NSconvergent_NSBseq,
+    convergent_NSconvergent_iff,Bseq_NSBseq_iff]) 1);
+qed "convergent_Bseq";
+
+(*----------------------------------------------------------------------
+             Results about Ubs and Lubs of bounded sequences
+ -----------------------------------------------------------------------*)
+Goalw [Bseq_def]
+  "!!(X::nat=>real). Bseq X ==> \
+\  EX U. isUb (UNIV::real set) {x. EX n. X n = x} U";
+by (auto_tac (claset() addIs [isUbI,setleI],
+    simpset() addsimps [abs_le_interval_iff]));
+qed "Bseq_isUb";
+
+(*----------------------------------------------------------------------
+   Use completeness of reals (supremum property) 
+   to show that any bounded sequence has a lub 
+-----------------------------------------------------------------------*)
+Goal
+  "!!(X::nat=>real). Bseq X ==> \
+\  EX U. isLub (UNIV::real set) {x. EX n. X n = x} U";
+by (blast_tac (claset() addIs [reals_complete,
+    Bseq_isUb]) 1);
+qed "Bseq_isLub";
+
+(* nonstandard version of premise will be *)
+(* handy when we work in NS universe      *)
+Goal   "NSBseq X ==> \
+\  EX U. isUb (UNIV::real set) {x. EX n. X n = x} U";
+by (asm_full_simp_tac (simpset() addsimps 
+    [Bseq_NSBseq_iff RS sym,Bseq_isUb]) 1);
+qed "NSBseq_isUb";
+
+Goal
+  "NSBseq X ==> \
+\  EX U. isLub (UNIV::real set) {x. EX n. X n = x} U";
+by (asm_full_simp_tac (simpset() addsimps 
+    [Bseq_NSBseq_iff RS sym,Bseq_isLub]) 1);
+qed "NSBseq_isLub";
+
+(*--------------------------------------------------------------------
+             Bounded and monotonic sequence converges              
+ --------------------------------------------------------------------*)
+(* lemmas *)
+Goal 
+     "!!(X::nat=>real). [| ALL m n. m <= n -->  X m <= X n; \
+\                 isLub (UNIV::real set) {x. EX n. X n = x} (X ma) \
+\              |] ==> ALL n. ma <= n --> X n = X ma";
+by (Step_tac 1);
+by (dres_inst_tac [("y","X n")] isLubD2 1);
+by (ALLGOALS(blast_tac (claset() addDs [real_le_anti_sym])));
+val lemma_converg1 = result();
+
+(*------------------------------------------------------------------- 
+   The best of both world: Easier to prove this result as a standard
+   theorem and then use equivalence to "transfer" it into the
+   equivalent nonstandard form if needed!
+ -------------------------------------------------------------------*)
+Goalw [LIMSEQ_def] 
+         "ALL n. m <= n --> X n = X m \
+\         ==> EX L. (X ----> L)";  
+by (res_inst_tac [("x","X m")] exI 1);
+by (Step_tac 1);
+by (res_inst_tac [("x","m")] exI 1);
+by (Step_tac 1);
+by (dtac spec 1 THEN etac impE 1);
+by (Auto_tac);
+qed "Bmonoseq_LIMSEQ";
+
+(* Now same theorem in terms of NS limit *)
+Goal "ALL n. m <= n --> X n = X m \
+\         ==> EX L. (X ----NS> L)";  
+by (auto_tac (claset() addSDs [Bmonoseq_LIMSEQ],
+    simpset() addsimps [LIMSEQ_NSLIMSEQ_iff]));
+qed "Bmonoseq_NSLIMSEQ";
+
+(* a few more lemmas *)
+Goal "!!(X::nat=>real). \
+\ [| ALL m. X m ~= U;  isLub UNIV {x. EX n. X n = x} U |] ==> ALL m. X m < U";
+by (Step_tac 1);
+by (dres_inst_tac [("y","X m")] isLubD2 1);
+by (auto_tac (claset() addSDs [order_le_imp_less_or_eq],
+              simpset()));
+val lemma_converg2 = result();
+
+Goal "!!(X ::nat=>real). ALL m. X m <= U ==> \
+\         isUb UNIV {x. EX n. X n = x} U";
+by (rtac (setleI RS isUbI) 1);
+by (Auto_tac);
+val lemma_converg3 = result();
+
+(* FIXME: U - T < U redundant *)
+Goal "!!(X::nat=> real). \
+\              [| ALL m. X m ~= U; \
+\                 isLub UNIV {x. EX n. X n = x} U; \
+\                 #0 < T; \
+\                 U + - T < U \
+\              |] ==> EX m. U + -T < X m & X m < U";
+by (dtac lemma_converg2 1 THEN assume_tac 1);
+by (rtac ccontr 1 THEN Asm_full_simp_tac 1);
+by (fold_tac [real_le_def]);
+by (dtac lemma_converg3 1);
+by (dtac isLub_le_isUb 1 THEN assume_tac 1);
+by (auto_tac (claset() addDs [order_less_le_trans],
+              simpset() addsimps [real_minus_zero_le_iff]));
+val lemma_converg4 = result();
+
+(*-------------------------------------------------------------------
+  A standard proof of the theorem for monotone increasing sequence
+ ------------------------------------------------------------------*)
+
+Goalw [convergent_def] 
+     "[| Bseq X; ALL m n. m <= n --> X m <= X n |] \
+\                ==> convergent X";
+by (forward_tac [Bseq_isLub] 1);
+by (Step_tac 1);
+by (case_tac "EX m. X m = U" 1 THEN Auto_tac);
+by (blast_tac (claset() addDs [lemma_converg1,
+    Bmonoseq_LIMSEQ]) 1);
+(* second case *)
+by (res_inst_tac [("x","U")] exI 1);
+by (rtac LIMSEQI 1 THEN Step_tac 1);
+by (forward_tac [lemma_converg2] 1 THEN assume_tac 1);
+by (dtac lemma_converg4 1 THEN Auto_tac);
+by (res_inst_tac [("x","m")] exI 1 THEN Step_tac 1);
+by (subgoal_tac "X m <= X n" 1 THEN Fast_tac 2);
+by (rotate_tac 3 1 THEN dres_inst_tac [("x","n")] spec 1);
+by (arith_tac 1);
+qed "Bseq_mono_convergent";
+
+(* NS version of theorem *)
+Goalw [convergent_def] 
+     "[| NSBseq X; ALL m n. m <= n --> X m <= X n |] \
+\                ==> NSconvergent X";
+by (auto_tac (claset() addIs [Bseq_mono_convergent], 
+    simpset() addsimps [convergent_NSconvergent_iff RS sym,
+    Bseq_NSBseq_iff RS sym]));
+qed "NSBseq_mono_NSconvergent";
+
+Goalw [convergent_def] 
+      "(convergent X) = (convergent (%n. -(X n)))";
+by (auto_tac (claset() addDs [LIMSEQ_minus], simpset()));
+by (dtac LIMSEQ_minus 1 THEN Auto_tac);
+qed "convergent_minus_iff";
+
+Goalw [Bseq_def] "Bseq (%n. -(X n)) = Bseq X";
+by (Asm_full_simp_tac 1);
+qed "Bseq_minus_iff";
+
+(*--------------------------------
+   **** main mono theorem ****
+ -------------------------------*)
+Goalw [monoseq_def] "[| Bseq X; monoseq X |] ==> convergent X";
+by (Step_tac 1);
+by (rtac (convergent_minus_iff RS ssubst) 2);
+by (dtac (Bseq_minus_iff RS ssubst) 2);
+by (auto_tac (claset() addSIs [Bseq_mono_convergent], simpset()));
+qed "Bseq_monoseq_convergent";
+
+(*----------------------------------------------------------------
+          A few more equivalence theorems for boundedness 
+ ---------------------------------------------------------------*)
+ 
+(***--- alternative formulation for boundedness---***)
+Goalw [Bseq_def] 
+   "Bseq X = (EX k x. #0 < k & (ALL n. abs(X(n) + -x) <= k))";
+by (Step_tac 1);
+by (res_inst_tac [("x","k + abs(x)")] exI 2);
+by (res_inst_tac [("x","K")] exI 1);
+by (res_inst_tac [("x","0")] exI 1);
+by (Auto_tac);
+by (ALLGOALS (dres_inst_tac [("x","n")] spec));
+by (ALLGOALS arith_tac);
+qed "Bseq_iff2";
+
+(***--- alternative formulation for boundedness ---***)
+Goal "Bseq X = (EX k N. #0 < k & (ALL n. abs(X(n) + -X(N)) <= k))";
+by (Step_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [Bseq_def]) 1);
+by (Step_tac 1);
+by (res_inst_tac [("x","K + abs(X N)")] exI 1);
+by (Auto_tac);
+by (arith_tac 1);
+by (res_inst_tac [("x","N")] exI 1);
+by (Step_tac 1);
+by (dres_inst_tac [("x","n")] spec 1);
+by (arith_tac 1);
+by (auto_tac (claset(), simpset() addsimps [Bseq_iff2]));
+qed "Bseq_iff3";
+
+Goalw [Bseq_def] "(ALL n. k <= f n & f n <= K) ==> Bseq f";
+by (res_inst_tac [("x","(abs(k) + abs(K)) + #1")] exI 1);
+by (Auto_tac);
+by (dres_inst_tac [("x","n")] spec 2);
+by (ALLGOALS arith_tac);
+qed "BseqI2";
+
+(*-------------------------------------------------------------------
+   Equivalence between NS and standard definitions of Cauchy seqs
+ ------------------------------------------------------------------*)
+(*-------------------------------
+      Standard def => NS def
+ -------------------------------*)
+Goal "Abs_hypnat (hypnatrel ^^ {x}) : HNatInfinite \
+\         ==> {n. M <= x n} : FreeUltrafilterNat";
+by (auto_tac (claset(),
+              simpset() addsimps [HNatInfinite_FreeUltrafilterNat_iff]));
+by (dres_inst_tac [("x","M")] spec 1);
+by (ultra_tac (claset(),simpset() addsimps [less_imp_le]) 1);
+val lemmaCauchy1 = result();
+
+Goal "{n. ALL m n. M <= m & M <= (n::nat) --> abs (X m + - X n) < u} Int \
+\     {n. M <= xa n} Int {n. M <= x n} <= \
+\     {n. abs (X (xa n) + - X (x n)) < u}";
+by (Blast_tac 1);
+val lemmaCauchy2 = result();
+
+Goalw [Cauchy_def,NSCauchy_def] 
+      "Cauchy X ==> NSCauchy X";
+by (Step_tac 1);
+by (res_inst_tac [("z","M")] eq_Abs_hypnat 1);
+by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
+by (rtac (inf_close_minus_iff RS iffD2) 1);
+by (rtac (mem_infmal_iff RS iffD1) 1);
+by (auto_tac (claset(),simpset() addsimps [starfunNat,
+    hypreal_minus,hypreal_add,Infinitesimal_FreeUltrafilterNat_iff]));
+by (EVERY[rtac bexI 1, Auto_tac]);
+by (dtac spec 1 THEN Auto_tac);
+by (dres_inst_tac [("M","M")] lemmaCauchy1 1);
+by (dres_inst_tac [("M","M")] lemmaCauchy1 1);
+by (res_inst_tac [("x1","xa")] 
+    (lemmaCauchy2 RSN (2,FreeUltrafilterNat_subset)) 1);
+by (rtac FreeUltrafilterNat_Int 1 THEN assume_tac 2);
+by (auto_tac (claset() addIs [FreeUltrafilterNat_Int,
+        FreeUltrafilterNat_Nat_set], simpset()));
+qed "Cauchy_NSCauchy";
+
+(*-----------------------------------------------
+     NS def => Standard def -- rather long but 
+     straightforward proof in this case
+ ---------------------------------------------*)
+Goalw [Cauchy_def,NSCauchy_def] 
+      "NSCauchy X ==> Cauchy X";
+by (EVERY1[Step_tac, rtac ccontr,Asm_full_simp_tac]);
+by (dtac choice 1 THEN auto_tac (claset(),simpset() 
+         addsimps [all_conj_distrib]));
+by (dtac choice 1 THEN step_tac (claset() addSDs 
+         [all_conj_distrib RS iffD1]) 1);
+by (REPEAT(dtac HNatInfinite_NSLIMSEQ 1));
+by (dtac bspec 1 THEN assume_tac 1);
+by (dres_inst_tac [("x","Abs_hypnat (hypnatrel ^^ {fa})")] bspec 1 
+    THEN auto_tac (claset(),simpset() addsimps [starfunNat]));
+by (dtac (inf_close_minus_iff RS iffD1) 1);
+by (dtac (mem_infmal_iff RS iffD2) 1);
+by (auto_tac (claset(),simpset() addsimps [hypreal_minus,
+    hypreal_add,Infinitesimal_FreeUltrafilterNat_iff]));
+by (dres_inst_tac [("x","e")] spec 1 THEN Auto_tac);
+by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
+by (dtac (CLAIM "{n. X (f n) + - X (fa n) = Y n} Int \
+\         {n. abs (Y n) < e} <= \
+\         {n. abs (X (f n) + - X (fa n)) < e}" RSN 
+          (2,FreeUltrafilterNat_subset)) 1);
+by (thin_tac "{n. abs (Y n) < e} : FreeUltrafilterNat" 1);
+by (dtac FreeUltrafilterNat_all 1);
+by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
+by (asm_full_simp_tac (simpset() addsimps 
+    [CLAIM "{n. abs (X (f n) + - X (fa n)) < e} Int \
+\         {M. ~ abs (X (f M) + - X (fa M)) < e} = {}",
+     FreeUltrafilterNat_empty]) 1);
+qed "NSCauchy_Cauchy";
+
+(*----- Equivalence -----*)
+Goal "NSCauchy X = Cauchy X";
+by (blast_tac (claset() addSIs[NSCauchy_Cauchy,
+    Cauchy_NSCauchy]) 1);
+qed "NSCauchy_Cauchy_iff";
+
+(*-------------------------------------------------------
+  Cauchy sequence is bounded -- this is the standard 
+  proof mechanization rather than the nonstandard proof 
+ -------------------------------------------------------*)
+
+(***-------------  VARIOUS LEMMAS --------------***)
+Goal "ALL n. M <= n --> abs (X M + - X n) < (#1::real) \
+\         ==>  ALL n. M <= n --> abs(X n) < #1 + abs(X M)";
+by (Step_tac 1);
+by (dtac spec 1 THEN Auto_tac);
+by (arith_tac 1);
+val lemmaCauchy = result();
+
+Goal "(n < Suc M) = (n <= M)";
+by Auto_tac;
+qed "less_Suc_cancel_iff";
+
+(* FIXME: Long. Maximal element in subsequence *)
+Goal "EX m. m <= M & (ALL n. n <= M --> \
+\         abs ((X::nat=> real) n) <= abs (X m))";
+by (induct_tac "M" 1);
+by (res_inst_tac [("x","0")] exI 1);
+by (Asm_full_simp_tac 1);
+by (Step_tac 1);
+by (cut_inst_tac [("R1.0","abs (X (Suc n))"),("R2.0","abs(X m)")]
+        real_linear 1);
+by (Step_tac 1);
+by (res_inst_tac [("x","m")] exI 1);
+by (res_inst_tac [("x","m")] exI 2);
+by (res_inst_tac [("x","Suc n")] exI 3);
+by (ALLGOALS(Asm_full_simp_tac));
+by (Step_tac 1);
+by (ALLGOALS(eres_inst_tac [("m1","na")] 
+    (le_imp_less_or_eq RS disjE)));
+by (ALLGOALS(asm_full_simp_tac (simpset() addsimps 
+    [less_Suc_cancel_iff, order_less_imp_le])));
+by (blast_tac (claset() addIs [order_le_less_trans RS order_less_imp_le]) 1);
+qed "SUP_rabs_subseq";
+
+(* lemmas to help proof - mostly trivial *)
+Goal "[| ALL m::nat. m <= M --> P M m; \
+\        ALL m. M <= m --> P M m |] \
+\     ==> ALL m. P M m";
+by (Step_tac 1);
+by (REPEAT(dres_inst_tac [("x","m")] spec 1));
+by (auto_tac (claset() addEs [less_asym],
+    simpset() addsimps [le_def]));
+val lemma_Nat_covered = result();
+
+Goal "[| ALL n. n <= M --> abs ((X::nat=>real) n) <= a;  a < b |] \
+\     ==> ALL n. n <= M --> abs(X n) <= b";
+by (blast_tac (claset() addIs [order_le_less_trans RS order_less_imp_le]) 1);
+val lemma_trans1 = result();
+
+Goal "[| ALL n. M <= n --> abs ((X::nat=>real) n) < a; \
+\        a < b |] \
+\     ==> ALL n. M <= n --> abs(X n)<= b";
+by (blast_tac (claset() addIs [order_less_trans RS order_less_imp_le]) 1);
+val lemma_trans2 = result();
+
+Goal "[| ALL n. n <= M --> abs (X n) <= a; \
+\        a = b |] \
+\     ==> ALL n. n <= M --> abs(X n) <= b";
+by (Auto_tac);
+val lemma_trans3 = result();
+
+Goal "ALL n. M <= n --> abs ((X::nat=>real) n) < a \
+\             ==>  ALL n. M <= n --> abs (X n) <= a";
+by (blast_tac (claset() addIs [order_less_imp_le]) 1);
+val lemma_trans4 = result();
+
+(*---------------------------------------------------------- 
+   Trickier than expected --- proof is more involved than
+   outlines sketched by various authors would suggest
+ ---------------------------------------------------------*)
+Goalw [Cauchy_def,Bseq_def] "Cauchy X ==> Bseq X";
+by (dres_inst_tac [("x","#1")] spec 1);
+by (etac (rename_numerals real_zero_less_one RSN (2,impE)) 1);
+by (Step_tac 1);
+by (dres_inst_tac [("x","M")] spec 1);
+by (Asm_full_simp_tac 1);
+by (dtac lemmaCauchy 1);
+by (cut_inst_tac [("M","M"),("X","X")] SUP_rabs_subseq 1);
+by (Step_tac 1);
+by (cut_inst_tac [("R1.0","abs(X m)"),
+     ("R2.0","#1 + abs(X M)")] real_linear 1);
+by (Step_tac 1);
+by (dtac lemma_trans1 1 THEN assume_tac 1);
+by (dtac lemma_trans2 3 THEN assume_tac 3);
+by (dtac lemma_trans3 2 THEN assume_tac 2);
+by (dtac (abs_add_one_gt_zero RS order_less_trans) 3);
+by (dtac lemma_trans4 1);
+by (dtac lemma_trans4 2);
+by (res_inst_tac [("x","#1 + abs(X M)")] exI 1);
+by (res_inst_tac [("x","#1 + abs(X M)")] exI 2);
+by (res_inst_tac [("x","abs(X m)")] exI 3);
+by (auto_tac (claset() addSEs [lemma_Nat_covered],
+              simpset()));
+by (ALLGOALS arith_tac);
+qed "Cauchy_Bseq";
+
+(*------------------------------------------------
+  Cauchy sequence is bounded -- NSformulation
+ ------------------------------------------------*)
+Goal "NSCauchy X ==> NSBseq X";
+by (asm_full_simp_tac (simpset() addsimps [Cauchy_Bseq,
+    Bseq_NSBseq_iff RS sym,NSCauchy_Cauchy_iff]) 1);
+qed "NSCauchy_NSBseq";
+
+
+(*-----------------------------------------------------------------
+          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.
+ -----------------------------------------------------------------*)
+Goalw [NSconvergent_def,NSLIMSEQ_def]
+      "NSCauchy X = NSconvergent X";
+by (Step_tac 1);
+by (forward_tac [NSCauchy_NSBseq] 1);
+by (auto_tac (claset() addIs [inf_close_trans2], 
+    simpset() addsimps 
+    [NSBseq_def,NSCauchy_def]));
+by (dtac (HNatInfinite_whn RSN (2,bspec)) 1);
+by (dtac (HNatInfinite_whn RSN (2,bspec)) 1);
+by (auto_tac (claset() addSDs [st_part_Ex], simpset() 
+              addsimps [SReal_iff]));
+by (blast_tac (claset() addIs [inf_close_trans3]) 1);
+qed "NSCauchy_NSconvergent_iff";
+
+(* Standard proof for free *)
+Goal "Cauchy X = convergent X";
+by (simp_tac (simpset() addsimps [NSCauchy_Cauchy_iff RS sym,
+    convergent_NSconvergent_iff, NSCauchy_NSconvergent_iff]) 1);
+qed "Cauchy_convergent_iff";
+
+(*-----------------------------------------------------------------
+     We can now try and derive a few properties of sequences
+     starting with the limit comparison property for sequences
+ -----------------------------------------------------------------*)
+Goalw [NSLIMSEQ_def]
+       "[| f ----NS> l; g ----NS> m; \
+\                  EX N. ALL n. N <= n --> f(n) <= g(n) \
+\               |] ==> l <= m";
+by (Step_tac 1);
+by (dtac starfun_le_mono 1);
+by (REPEAT(dtac (HNatInfinite_whn RSN (2,bspec)) 1));
+by (dres_inst_tac [("x","whn")] spec 1);
+by (REPEAT(dtac (bex_Infinitesimal_iff2 RS iffD2) 1));
+by Auto_tac;
+by (auto_tac (claset() addIs 
+    [hypreal_of_real_le_add_Infininitesimal_cancel2], simpset()));
+qed "NSLIMSEQ_le";
+
+(* standard version *)
+Goal "[| f ----> l; g ----> m; \
+\        EX N. ALL n. N <= n --> f(n) <= g(n) |] \
+\     ==> l <= m";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
+    NSLIMSEQ_le]) 1);
+qed "LIMSEQ_le";
+
+(*---------------
+    Also...
+ --------------*)
+Goal "[| X ----> r; ALL n. a <= X n |] ==> a <= r";
+by (rtac LIMSEQ_le 1);
+by (rtac LIMSEQ_const 1);
+by (Auto_tac);
+qed "LIMSEQ_le_const";
+
+Goal "[| X ----NS> r; ALL n. a <= X n |] ==> a <= r";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
+    LIMSEQ_le_const]) 1);
+qed "NSLIMSEQ_le_const";
+
+Goal "[| X ----> r; ALL n. X n <= a |] ==> r <= a";
+by (rtac LIMSEQ_le 1);
+by (rtac LIMSEQ_const 2);
+by (Auto_tac);
+qed "LIMSEQ_le_const2";
+
+Goal "[| X ----NS> r; ALL n. X n <= a |] ==> r <= a";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
+    LIMSEQ_le_const2]) 1);
+qed "NSLIMSEQ_le_const2";
+
+(*-----------------------------------------------------
+            Shift a convergent series by 1
+  We use the fact that Cauchyness and convergence
+  are equivalent and also that the successor of an
+  infinite hypernatural is also infinite.
+ -----------------------------------------------------*)
+Goal "f ----NS> l ==> (%n. f(Suc n)) ----NS> l";
+by (forward_tac [NSconvergentI RS 
+    (NSCauchy_NSconvergent_iff RS iffD2)] 1);
+by (auto_tac (claset(),simpset() addsimps [NSCauchy_def,
+    NSLIMSEQ_def,starfunNat_shift_one]));
+by (dtac bspec 1 THEN assume_tac 1);
+by (dtac bspec 1 THEN assume_tac 1);
+by (dtac (SHNat_one RSN (2,HNatInfinite_SHNat_add)) 1);
+by (blast_tac (claset() addIs [inf_close_trans3]) 1);
+qed "NSLIMSEQ_Suc";
+
+(* standard version *)
+Goal "f ----> l ==> (%n. f(Suc n)) ----> l";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
+    NSLIMSEQ_Suc]) 1);
+qed "LIMSEQ_Suc";
+
+Goal "(%n. f(Suc n)) ----NS> l ==> f ----NS> l";
+by (forward_tac [NSconvergentI RS 
+    (NSCauchy_NSconvergent_iff RS iffD2)] 1);
+by (auto_tac (claset(),simpset() addsimps [NSCauchy_def,
+    NSLIMSEQ_def,starfunNat_shift_one]));
+by (dtac bspec 1 THEN assume_tac 1);
+by (dtac bspec 1 THEN assume_tac 1);
+by (ftac (SHNat_one RSN (2,HNatInfinite_SHNat_diff)) 1);
+by (rotate_tac 2 1);
+by (auto_tac (claset() addSDs [bspec] addIs [inf_close_trans3],
+    simpset()));
+qed "NSLIMSEQ_imp_Suc";
+
+Goal "(%n. f(Suc n)) ----> l ==> f ----> l";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff]) 1);
+by (etac NSLIMSEQ_imp_Suc 1);
+qed "LIMSEQ_imp_Suc";
+
+Goal "((%n. f(Suc n)) ----> l) = (f ----> l)";
+by (blast_tac (claset() addIs [LIMSEQ_imp_Suc,LIMSEQ_Suc]) 1);
+qed "LIMSEQ_Suc_iff";
+
+Goal "((%n. f(Suc n)) ----NS> l) = (f ----NS> l)";
+by (blast_tac (claset() addIs [NSLIMSEQ_imp_Suc,NSLIMSEQ_Suc]) 1);
+qed "NSLIMSEQ_Suc_iff";
+
+(*-----------------------------------------------------
+       A sequence tends to zero iff its abs does
+ ----------------------------------------------------*)
+(* we can prove this directly since proof is trivial *)
+Goalw [LIMSEQ_def] 
+      "((%n. abs(f n)) ----> #0) = (f ----> #0)";
+by (simp_tac (simpset() addsimps [abs_idempotent]) 1);
+qed "LIMSEQ_rabs_zero";
+
+(*-----------------------------------------------------*)
+(* We prove the NS version from the standard one       *)
+(* Actually pure NS proof seems more complicated       *)
+(* than the direct standard one above!                 *)
+(*-----------------------------------------------------*)
+
+Goal "((%n. abs(f n)) ----NS> #0) = (f ----NS> #0)";
+by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
+             LIMSEQ_rabs_zero]) 1);
+qed "NSLIMSEQ_rabs_zero";
+
+(*----------------------------------------
+    Also we have for a general limit 
+        (NS proof much easier)
+ ---------------------------------------*)
+Goalw [NSLIMSEQ_def] 
+       "f ----NS> l ==> (%n. abs(f n)) ----NS> abs(l)";
+by (auto_tac (claset() addIs [inf_close_hrabs], simpset() 
+    addsimps [starfunNat_rabs,hypreal_of_real_hrabs RS sym]));
+qed "NSLIMSEQ_imp_rabs";
+
+(* standard version *)
+Goal "f ----> l ==> (%n. abs(f n)) ----> abs(l)";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
+    NSLIMSEQ_imp_rabs]) 1);
+qed "LIMSEQ_imp_rabs";
+
+(*-----------------------------------------------------
+       An unbounded sequence's inverse tends to 0
+  ----------------------------------------------------*)
+(* standard proof seems easier *)
+Goalw [LIMSEQ_def] 
+      "ALL y. EX N. ALL n. N <= n --> y < f(n) \
+\      ==> (%n. inverse(f n)) ----> #0";
+by (Step_tac 1 THEN Asm_full_simp_tac 1);
+by (dres_inst_tac [("x","inverse r")] spec 1 THEN Step_tac 1);
+by (res_inst_tac [("x","N")] exI 1 THEN Step_tac 1);
+by (dtac spec 1 THEN Auto_tac);
+by (forward_tac [real_inverse_gt_0] 1);
+by (forward_tac [order_less_trans] 1 THEN assume_tac 1);
+by (forw_inst_tac [("x","f n")] real_inverse_gt_0 1);
+by (asm_simp_tac (simpset() addsimps [abs_eqI2]) 1);
+by (res_inst_tac [("t","r")] (real_inverse_inverse RS subst) 1);
+by (auto_tac (claset() addIs [real_inverse_less_iff RS iffD2], 
+              simpset() delsimps [real_inverse_inverse]));
+qed "LIMSEQ_inverse_zero";
+
+Goal "ALL y. EX N. ALL n. N <= n --> y < f(n) \
+\     ==> (%n. inverse(f n)) ----NS> #0";
+by (asm_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
+                  LIMSEQ_inverse_zero]) 1);
+qed "NSLIMSEQ_inverse_zero";
+
+(*--------------------------------------------------------------
+             Sequence  1/n --> 0 as n --> infinity 
+ -------------------------------------------------------------*)
+Goal "(%n. inverse(real_of_posnat n)) ----> #0";
+by (rtac LIMSEQ_inverse_zero 1 THEN Step_tac 1);
+by (cut_inst_tac [("x","y")] reals_Archimedean2 1);
+by (Step_tac 1 THEN res_inst_tac [("x","n")] exI 1);
+by (Step_tac 1 THEN etac (le_imp_less_or_eq RS disjE) 1);
+by (dtac (real_of_posnat_less_iff RS iffD2) 1);
+by (auto_tac (claset() addEs [order_less_trans], simpset()));
+qed "LIMSEQ_inverse_real_of_posnat";
+
+Goal "(%n. inverse(real_of_posnat n)) ----NS> #0";
+by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
+    LIMSEQ_inverse_real_of_posnat]) 1);
+qed "NSLIMSEQ_inverse_real_of_posnat";
+
+(*--------------------------------------------
+    Sequence  r + 1/n --> r as n --> infinity 
+    now easily proved
+ --------------------------------------------*)
+Goal "(%n. r + inverse(real_of_posnat n)) ----> r";
+by (cut_facts_tac [[LIMSEQ_const,LIMSEQ_inverse_real_of_posnat]
+    MRS LIMSEQ_add] 1);
+by (Auto_tac);
+qed "LIMSEQ_inverse_real_of_posnat_add";
+
+Goal "(%n. r + inverse(real_of_posnat n)) ----NS> r";
+by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
+    LIMSEQ_inverse_real_of_posnat_add]) 1);
+qed "NSLIMSEQ_inverse_real_of_posnat_add";
+
+(*--------------
+    Also...
+ --------------*)
+
+Goal "(%n. r + -inverse(real_of_posnat n)) ----> r";
+by (cut_facts_tac [[LIMSEQ_const,LIMSEQ_inverse_real_of_posnat]
+    MRS LIMSEQ_add_minus] 1);
+by (Auto_tac);
+qed "LIMSEQ_inverse_real_of_posnat_add_minus";
+
+Goal "(%n. r + -inverse(real_of_posnat n)) ----NS> r";
+by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
+    LIMSEQ_inverse_real_of_posnat_add_minus]) 1);
+qed "NSLIMSEQ_inverse_real_of_posnat_add_minus";
+
+Goal "(%n. r*( #1 + -inverse(real_of_posnat n))) ----> r";
+by (cut_inst_tac [("b","#1")] ([LIMSEQ_const,
+    LIMSEQ_inverse_real_of_posnat_add_minus] MRS LIMSEQ_mult) 1);
+by (Auto_tac);
+qed "LIMSEQ_inverse_real_of_posnat_add_minus_mult";
+
+Goal "(%n. r*( #1 + -inverse(real_of_posnat n))) ----NS> r";
+by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
+    LIMSEQ_inverse_real_of_posnat_add_minus_mult]) 1);
+qed "NSLIMSEQ_inverse_real_of_posnat_add_minus_mult";
+
+(*---------------------------------------------------------------
+                          Real Powers
+ --------------------------------------------------------------*)
+Goal "(X ----NS> a) --> ((%n. (X n) ^ m) ----NS> a ^ m)";
+by (induct_tac "m" 1);
+by (auto_tac (claset() addIs [NSLIMSEQ_mult,NSLIMSEQ_const],
+    simpset()));
+qed_spec_mp "NSLIMSEQ_pow";
+
+Goal "X ----> a ==> (%n. (X n) ^ m) ----> a ^ m";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
+    NSLIMSEQ_pow]) 1);
+qed "LIMSEQ_pow";
+
+(*----------------------------------------------------------------
+               0 <= x < #1 ==> (x ^ n ----> 0)
+  Proof will use (NS) Cauchy equivalence for convergence and
+  also fact that bounded and monotonic sequence converges.  
+ ---------------------------------------------------------------*)
+Goalw [Bseq_def] 
+      "[| #0 <= x; x < #1 |] ==> Bseq (%n. x ^ n)";
+by (res_inst_tac [("x","#1")] exI 1);
+by (auto_tac (claset() addDs [conjI RS realpow_le2] 
+                       addIs [order_less_imp_le], 
+       simpset() addsimps [real_zero_less_one,abs_eqI1,realpow_abs RS sym] ));
+qed "Bseq_realpow";
+
+Goal "[| #0 <= x; x < #1 |] ==> monoseq (%n. x ^ n)";
+by (blast_tac (claset() addSIs [mono_SucI2,realpow_Suc_le3]) 1);
+qed "monoseq_realpow";
+
+Goal "[| #0 <= x; x < #1 |] ==> convergent (%n. x ^ n)";
+by (blast_tac (claset() addSIs [Bseq_monoseq_convergent,
+                                Bseq_realpow,monoseq_realpow]) 1);
+qed "convergent_realpow";
+
+(* We now use NS criterion to bring proof of theorem through *)
+
+
+Goalw [NSLIMSEQ_def]
+     "[| #0 <= x; x < #1 |] ==> (%n. x ^ n) ----NS> #0";
+by (auto_tac (claset() addSDs [convergent_realpow],
+              simpset() addsimps [convergent_NSconvergent_iff]));
+by (forward_tac [NSconvergentD] 1);
+by (auto_tac (claset(),
+        simpset() addsimps [NSLIMSEQ_def, NSCauchy_NSconvergent_iff RS sym,
+                            NSCauchy_def, starfunNat_pow]));
+by (forward_tac [HNatInfinite_add_one] 1);
+by (dtac bspec 1 THEN assume_tac 1);
+by (dtac bspec 1 THEN assume_tac 1);
+by (dres_inst_tac [("x","N + 1hn")] bspec 1 THEN assume_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [hyperpow_add]) 1);
+by (dtac inf_close_mult_subst_SReal 1 THEN assume_tac 1);
+by (dtac inf_close_trans3 1 THEN assume_tac 1);
+by (auto_tac (claset(),
+              simpset() delsimps [hypreal_of_real_mult]
+			addsimps [hypreal_of_real_mult RS sym]));
+qed "NSLIMSEQ_realpow_zero";
+
+(*---------------  standard version ---------------*)
+Goal "[| #0 <= x; x < #1 |] ==> (%n. x ^ n) ----> #0";
+by (asm_simp_tac (simpset() addsimps [NSLIMSEQ_realpow_zero,
+                                      LIMSEQ_NSLIMSEQ_iff]) 1);
+qed "LIMSEQ_realpow_zero";
+
+Goal "#1 < x ==> (%n. a / (x ^ n)) ----> #0";
+by (cut_inst_tac [("a","a"),("x1","inverse x")] 
+    ([LIMSEQ_const, LIMSEQ_realpow_zero] MRS LIMSEQ_mult) 1);
+by (auto_tac (claset(), 
+              simpset() addsimps [real_divide_def, realpow_inverse])); 
+by (asm_simp_tac (simpset() addsimps [real_inverse_eq_divide,
+                                      pos_real_divide_less_eq]) 1); 
+qed "LIMSEQ_divide_realpow_zero";
+
+(*----------------------------------------------------------------
+               Limit of c^n for |c| < 1  
+ ---------------------------------------------------------------*)
+Goal "abs(c) < #1 ==> (%n. abs(c) ^ n) ----> #0";
+by (blast_tac (claset() addSIs [LIMSEQ_realpow_zero,abs_ge_zero]) 1);
+qed "LIMSEQ_rabs_realpow_zero";
+
+Goal "abs(c) < #1 ==> (%n. abs(c) ^ n) ----NS> #0";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_rabs_realpow_zero,
+    LIMSEQ_NSLIMSEQ_iff RS sym]) 1);
+qed "NSLIMSEQ_rabs_realpow_zero";
+
+Goal "abs(c) < #1 ==> (%n. c ^ n) ----> #0";
+by (rtac (LIMSEQ_rabs_zero RS iffD1) 1);
+by (auto_tac (claset() addIs [LIMSEQ_rabs_realpow_zero],
+              simpset() addsimps [realpow_abs RS sym]));
+qed "LIMSEQ_rabs_realpow_zero2";
+
+Goal "abs(c) < #1 ==> (%n. c ^ n) ----NS> #0";
+by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_rabs_realpow_zero2,
+    LIMSEQ_NSLIMSEQ_iff RS sym]) 1);
+qed "NSLIMSEQ_rabs_realpow_zero2";
+
+(***---------------------------------------------------------------
+                 Hyperreals and Sequences
+ ---------------------------------------------------------------***)
+(*** A bounded sequence is a finite hyperreal ***)
+Goal "NSBseq X ==> Abs_hypreal(hyprel^^{X}) : HFinite";
+by (auto_tac (claset() addSIs [bexI,lemma_hyprel_refl] addIs 
+       [FreeUltrafilterNat_all RS FreeUltrafilterNat_subset],
+       simpset() addsimps [HFinite_FreeUltrafilterNat_iff,
+        Bseq_NSBseq_iff RS sym, Bseq_iff1a]));
+qed "NSBseq_HFinite_hypreal";
+
+(*** A sequence converging to zero defines an infinitesimal ***)
+Goalw [NSLIMSEQ_def] 
+      "X ----NS> #0 ==> Abs_hypreal(hyprel^^{X}) : Infinitesimal";
+by (dres_inst_tac [("x","whn")] bspec 1);
+by (simp_tac (simpset() addsimps [HNatInfinite_whn]) 1);
+by (auto_tac (claset(),simpset() addsimps [hypnat_omega_def,
+    mem_infmal_iff RS sym,starfunNat,hypreal_of_real_zero]));
+qed "NSLIMSEQ_zero_Infinitesimal_hypreal";
+
+(***---------------------------------------------------------------
+    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 "EX 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 <= f(n)";
+Goal "subseq f ==> EX n. N1 <= n & N2 <= f(n)";
+ ---------------------------------------------------------------***)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/SEQ.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,63 @@
+(*  Title       : SEQ.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Convergence of sequences and series
+*) 
+
+SEQ = NatStar + HyperPow + 
+
+constdefs
+
+  (* Standard definition of convergence of sequence *)           
+  LIMSEQ :: [nat=>real,real] => bool    ("((_)/ ----> (_))" [60, 60] 60)
+  "X ----> L == (ALL r. #0 < r --> (EX no. ALL n. no <= n --> abs (X n + -L) < r))"
+  
+  (* Nonstandard definition of convergence of sequence *)
+  NSLIMSEQ :: [nat=>real,real] => bool    ("((_)/ ----NS> (_))" [60, 60] 60)
+  "X ----NS> L == (ALL N: HNatInfinite. (*fNat* X) N @= hypreal_of_real L)"
+
+  (* define value of limit using choice operator*)
+  lim :: (nat => real) => real
+  "lim X == (@L. (X ----> L))"
+
+  nslim :: (nat => real) => real
+  "nslim X == (@L. (X ----NS> L))"
+
+  (* Standard definition of convergence *)
+  convergent :: (nat => real) => bool
+  "convergent X == (EX L. (X ----> L))"
+
+  (* Nonstandard definition of convergence *)
+  NSconvergent :: (nat => real) => bool
+  "NSconvergent X == (EX L. (X ----NS> L))"
+
+  (* Standard definition for bounded sequence *)
+  Bseq :: (nat => real) => bool
+  "Bseq X == (EX K. (#0 < K & (ALL n. abs(X n) <= K)))"
+ 
+  (* Nonstandard definition for bounded sequence *)
+  NSBseq :: (nat=>real) => bool
+  "NSBseq X == (ALL N: HNatInfinite. (*fNat* X) N : HFinite)" 
+
+  (* Definition for monotonicity *)
+  monoseq :: (nat=>real)=>bool
+  "monoseq X == ((ALL (m::nat) n. m <= n --> X m <= X n) |
+                 (ALL m n. m <= n --> X n <= X m))"   
+
+  (* Definition of subsequence *)
+  subseq :: (nat => nat) => bool
+  "subseq f == (ALL m n. m < n --> (f m) < (f n))"
+
+  (** Cauchy condition **)
+
+  (* Standard definition *)
+  Cauchy :: (nat => real) => bool
+  "Cauchy X == (ALL e. (#0 < e -->
+                       (EX M. (ALL m n. M <= m & M <= n 
+                             --> abs((X m) + -(X n)) < e))))"
+
+  NSCauchy :: (nat => real) => bool
+  "NSCauchy X == (ALL M: HNatInfinite. ALL N: HNatInfinite.
+                      (*fNat* X) M @= (*fNat* X) N)"
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/Series.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,610 @@
+(*  Title       : Series.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+                : 1999  University of Edinburgh
+    Description : Finite summation and infinite series
+*) 
+
+Goal "sumr (Suc n) n f = #0";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "sumr_Suc_zero";
+Addsimps [sumr_Suc_zero];
+
+Goal "sumr m m f = #0";
+by (induct_tac "m" 1);
+by (Auto_tac);
+qed "sumr_eq_bounds";
+Addsimps [sumr_eq_bounds];
+
+Goal "sumr m (Suc m) f = f(m)";
+by (Auto_tac);
+qed "sumr_Suc_eq";
+Addsimps [sumr_Suc_eq];
+
+Goal "sumr (m+k) k f = #0";
+by (induct_tac "k" 1);
+by (Auto_tac);
+qed "sumr_add_lbound_zero";
+Addsimps [sumr_add_lbound_zero];
+
+Goal "sumr m n f + sumr m n g = sumr m n (%n. f n + g n)";
+by (induct_tac "n" 1);
+by (auto_tac (claset(),simpset() addsimps real_add_ac));
+qed "sumr_add";
+
+Goal "r * sumr m n f = sumr m n (%n. r * f n)";
+by (induct_tac "n" 1);
+by (Auto_tac);
+by (auto_tac (claset(),simpset() addsimps 
+    [real_add_mult_distrib2]));
+qed "sumr_mult";
+
+Goal "n < p --> sumr 0 n f + sumr n p f = sumr 0 p f";
+by (induct_tac "p" 1);
+by (auto_tac (claset() addSDs [CLAIM "n < Suc na ==> n <= na",
+    leI] addDs [le_anti_sym], simpset()));
+qed_spec_mp "sumr_split_add";
+
+Goal "n < p ==> sumr 0 p f + \
+\                - sumr 0 n f = sumr n p f";
+by (dres_inst_tac [("f1","f")] (sumr_split_add RS sym) 1);
+by (asm_simp_tac (simpset() addsimps real_add_ac) 1);
+qed "sumr_split_add_minus";
+
+Goal "abs(sumr m n f) <= sumr m n (%i. abs(f i))";
+by (induct_tac "n" 1);
+by (auto_tac (claset() addIs [abs_triangle_ineq RS order_trans,
+                              real_add_le_mono1], 
+              simpset()));
+qed "sumr_rabs";
+
+Goal "!!f g. (ALL r. m <= r & r < n --> f r = g r) \
+\                --> sumr m n f = sumr m n g";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed_spec_mp "sumr_fun_eq";
+
+Goal "sumr 0 n (%i. r) = real_of_nat n*r";
+by (induct_tac "n" 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [real_add_mult_distrib,real_of_nat_zero]));
+qed "sumr_const";
+Addsimps [sumr_const];
+
+Goal "sumr 0 n f + -(real_of_nat n*r) = sumr 0 n (%i. f i + -r)";
+by (full_simp_tac (simpset() addsimps [sumr_add RS sym,
+    real_minus_mult_eq2] delsimps [real_minus_mult_eq2 RS sym]) 1);
+qed "sumr_add_mult_const";
+
+Goalw [real_diff_def] 
+     "sumr 0 n f - (real_of_nat n*r) = sumr 0 n (%i. f i - r)";
+by (full_simp_tac (simpset() addsimps [sumr_add_mult_const]) 1);
+qed "sumr_diff_mult_const";
+
+Goal "n < m --> sumr m n f = #0";
+by (induct_tac "n" 1);
+by (auto_tac (claset() addDs [less_imp_le], simpset()));
+qed_spec_mp "sumr_less_bounds_zero";
+Addsimps [sumr_less_bounds_zero];
+
+Goal "sumr m n (%i. - f i) = - sumr m n f";
+by (induct_tac "n" 1);
+by (case_tac "Suc n <= m" 2);
+by (auto_tac (claset(),simpset() addsimps 
+    [real_minus_add_distrib]));
+qed "sumr_minus";
+
+Goal "sumr (m+k) (n+k) f = sumr m n (%i. f(i + k))";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "sumr_shift_bounds";
+
+Goal "sumr 0 (#2*n) (%i. (#-1) ^ Suc i) = #0";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "sumr_minus_one_realpow_zero";
+Addsimps [sumr_minus_one_realpow_zero];
+
+Goal "m < Suc n ==> Suc n - m = Suc (n - m)";
+by (dtac less_imp_Suc_add 1);
+by (Auto_tac);
+val Suc_diff_n = result();
+
+Goal "(ALL n. m <= Suc n --> f n = r) & m <= na \
+\                --> sumr m na f = (real_of_nat (na - m) * r)";
+by (induct_tac "na" 1);
+by (Step_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [le_Suc_eq])  2);
+by (Step_tac 1);
+by (dres_inst_tac [("x","n")] spec 3);
+by (auto_tac (claset() addSDs [le_imp_less_or_eq],
+    simpset()));
+by (asm_simp_tac (simpset() addsimps [real_add_mult_distrib,
+    Suc_diff_n]) 1);
+qed_spec_mp "sumr_interval_const";
+
+Goal "(ALL n. m <= n --> f n = r) & m <= na \
+\     --> sumr m na f = (real_of_nat (na - m) * r)";
+by (induct_tac "na" 1);
+by (Step_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [le_Suc_eq])  2);
+by (Step_tac 1);
+by (dres_inst_tac [("x","n")] spec 3);
+by (auto_tac (claset() addSDs [le_imp_less_or_eq],
+    simpset()));
+by (asm_simp_tac (simpset() addsimps [Suc_diff_n,
+    real_add_mult_distrib]) 1);
+qed_spec_mp "sumr_interval_const2";
+
+Goal "(m <= n)= (m < Suc n)";
+by (Auto_tac);
+qed "le_eq_less_Suc";
+
+Goal "(ALL n. m <= n --> #0 <= f n) & m < na \
+\                --> sumr 0 m f <= sumr 0 na f";
+by (induct_tac "na" 1);
+by (Step_tac 1);
+by (ALLGOALS(asm_full_simp_tac (simpset() addsimps 
+    [le_eq_less_Suc RS sym])));
+by (ALLGOALS(dres_inst_tac [("x","n")] spec));
+by (Step_tac 1);
+by (dtac le_imp_less_or_eq 1 THEN Step_tac 1);
+by (dtac real_add_le_mono 2);
+by (dres_inst_tac [("i","sumr 0 m f")] (order_refl RS real_add_le_mono) 1);
+by (Auto_tac);
+qed_spec_mp "sumr_le";
+
+Goal "!!f g. (ALL r. m <= r & r < n --> f r <= g r) \
+\                --> sumr m n f <= sumr m n g";
+by (induct_tac "n" 1);
+by (auto_tac (claset() addIs [real_add_le_mono],
+    simpset() addsimps [le_def]));
+qed_spec_mp "sumr_le2";
+
+Goal "(ALL n. #0 <= f n) --> #0 <= sumr m n f";
+by (induct_tac "n" 1);
+by Auto_tac;
+by (dres_inst_tac [("x","n")] spec 1);
+by (arith_tac 1);
+qed_spec_mp "sumr_ge_zero";
+
+Goal "(ALL n. m <= n --> #0 <= f n) --> #0 <= sumr m n f";
+by (induct_tac "n" 1);
+by Auto_tac;
+by (dres_inst_tac [("x","n")] spec 1);
+by (arith_tac 1);
+qed_spec_mp "sumr_ge_zero2";
+
+Goal "#0 <= sumr m n (%n. abs (f n))";
+by (induct_tac "n" 1);
+by Auto_tac;
+by (arith_tac 1);
+qed "sumr_rabs_ge_zero";
+Addsimps [sumr_rabs_ge_zero];
+AddSIs [sumr_rabs_ge_zero]; 
+
+Goal "abs (sumr m n (%n. abs (f n))) = (sumr m n (%n. abs (f n)))";
+by (induct_tac "n" 1);
+by (Auto_tac THEN arith_tac 1);
+qed "rabs_sumr_rabs_cancel";
+Addsimps [rabs_sumr_rabs_cancel];
+
+Goal "ALL n. N <= n --> f n = #0 \
+\     ==> ALL m n. N <= m --> sumr m n f = #0";   
+by (Step_tac 1);
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "sumr_zero";
+
+Goal "Suc N <= n --> N <= n - 1";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed_spec_mp "Suc_le_imp_diff_ge";
+
+Goal "ALL n. N <= n --> f (Suc n) = #0 \
+\     ==> ALL m n. Suc N <= m --> sumr m n f = #0";   
+by (rtac sumr_zero 1 THEN Step_tac 1);
+by (subgoal_tac "0 < n" 1);
+by (dtac Suc_le_imp_diff_ge 1);
+by (Auto_tac);
+qed "Suc_le_imp_diff_ge2";
+
+(* proved elsewhere? *)
+Goal "(0 < n) = (EX m. n = Suc m)";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "gt_zero_eq_Ex";
+
+Goal "sumr 1 n (%n. f(n) * #0 ^ n) = #0";
+by (induct_tac "n" 1);
+by (auto_tac (claset(),simpset() addsimps [gt_zero_eq_Ex]));
+qed "sumr_one_lb_realpow_zero";
+Addsimps [sumr_one_lb_realpow_zero];
+
+Goalw [real_diff_def] "sumr m n f - sumr m n g = sumr m n (%n. f n - g n)";
+by (simp_tac (simpset() addsimps [sumr_add RS sym,sumr_minus]) 1);
+qed "sumr_diff";
+
+Goal "(ALL p. (m <= p & p < m + n --> (f p = g p))) --> sumr m n f = sumr m n g";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed_spec_mp "sumr_subst";
+
+Goal "(ALL p. m <= p & p < m + n --> (f(p) <= K)) \
+\     --> (sumr m (m + n) f <= (real_of_nat n * K))";
+by (induct_tac "n" 1);
+by (auto_tac (claset() addIs [real_add_le_mono],
+    simpset() addsimps [real_add_mult_distrib]));
+qed_spec_mp "sumr_bound";
+
+Goal "(ALL p. 0 <= p & p < n --> (f(p) <= K)) \
+\     --> (sumr 0 n f <= (real_of_nat n * K))";
+by (induct_tac "n" 1);
+by (auto_tac (claset() addIs [real_add_le_mono],
+    simpset() addsimps [real_add_mult_distrib]));
+qed_spec_mp "sumr_bound2";
+
+Goal "sumr 0 n (%m. sumr (m * k) (m*k + k) f) = sumr 0 (n * k) f";
+by (subgoal_tac "k = 0 | 0 < k" 1);
+by (Auto_tac);
+by (induct_tac "n" 1);
+by (auto_tac (claset(),simpset() addsimps [sumr_split_add,add_commute]));
+qed "sumr_group";
+Addsimps [sumr_group];
+
+(*-------------------------------------------------------------------
+                        Infinite sums
+    Theorems follow from properties of limits and sums
+ -------------------------------------------------------------------*)
+(*----------------------
+   suminf is the sum   
+ ---------------------*)
+Goalw [sums_def,summable_def] 
+      "f sums l ==> summable f";
+by (Blast_tac 1);
+qed "sums_summable";
+
+Goalw [summable_def,suminf_def]
+     "summable f ==> f sums (suminf f)";
+by (blast_tac (claset() addIs [someI2]) 1);
+qed "summable_sums";
+
+Goalw [summable_def,suminf_def,sums_def]
+     "summable f ==> (%n. sumr 0 n f) ----> (suminf f)";
+by (blast_tac (claset() addIs [someI2]) 1);
+qed "summable_sumr_LIMSEQ_suminf";
+
+(*-------------------
+    sum is unique                    
+ ------------------*)
+Goal "f sums s ==> (s = suminf f)";
+by (forward_tac [sums_summable RS summable_sums] 1);
+by (auto_tac (claset() addSIs [LIMSEQ_unique],
+    simpset() addsimps [sums_def]));
+qed "sums_unique";
+
+(*
+Goalw [sums_def,LIMSEQ_def] 
+     "(ALL m. n <= Suc m --> f(m) = #0) ==> f sums (sumr 0 n f)";
+by (Step_tac 1);
+by (res_inst_tac [("x","n")] exI 1);
+by (Step_tac 1 THEN forward_tac [le_imp_less_or_eq] 1);
+by (Step_tac 1);
+by (dres_inst_tac [("f","f")] sumr_split_add_minus 1);
+by (ALLGOALS (Asm_simp_tac));
+by (dtac (conjI RS sumr_interval_const) 1);
+by (Auto_tac);
+qed "series_zero";
+next one was called series_zero2
+**********************)
+
+Goalw [sums_def,LIMSEQ_def] 
+     "(ALL m. n <= m --> f(m) = #0) ==> f sums (sumr 0 n f)";
+by (Step_tac 1);
+by (res_inst_tac [("x","n")] exI 1);
+by (Step_tac 1 THEN forward_tac [le_imp_less_or_eq] 1);
+by (Step_tac 1);
+by (dres_inst_tac [("f","f")] sumr_split_add_minus 1);
+by (ALLGOALS (Asm_simp_tac));
+by (dtac (conjI RS sumr_interval_const2) 1);
+by (Auto_tac);
+qed "series_zero";
+
+Goalw [sums_def] "x sums x0 ==> (%n. c * x(n)) sums (c * x0)";
+by (auto_tac (claset() addSIs [LIMSEQ_mult] addIs [LIMSEQ_const], 
+              simpset() addsimps [sumr_mult RS sym]));
+qed "sums_mult";
+
+Goal "x sums x' ==> (%n. x(n)/c) sums (x'/c)";
+by (asm_simp_tac (simpset() addsimps [real_divide_def, sums_mult,
+                              inst "w" "inverse c" real_mult_commute]) 1); 
+qed "sums_divide";
+
+Goalw [sums_def] "[| x sums x0; y sums y0 |] ==> (%n. x n - y n) sums (x0-y0)";
+by (auto_tac (claset() addIs [LIMSEQ_diff],
+              simpset() addsimps [sumr_diff RS sym]));
+qed "sums_diff";
+
+Goal "summable f ==> suminf f * c = suminf(%n. f n * c)";
+by (auto_tac (claset() addSIs [sums_unique,sums_mult,summable_sums],
+    simpset() addsimps [real_mult_commute]));
+qed "suminf_mult";
+
+Goal "summable f ==> c * suminf f  = suminf(%n. c * f n)";
+by (auto_tac (claset() addSIs [sums_unique,sums_mult,summable_sums],
+    simpset()));
+qed "suminf_mult2";
+
+Goal "[| summable f; summable g |]  \
+\     ==> suminf f - suminf g  = suminf(%n. f n - g n)";
+by (auto_tac (claset() addSIs [sums_diff,sums_unique,summable_sums], simpset()));
+qed "suminf_diff";
+
+Goalw [sums_def] 
+       "x sums x0 ==> (%n. - x n) sums - x0";
+by (auto_tac (claset() addSIs [LIMSEQ_minus], simpset() addsimps [sumr_minus]));
+qed "sums_minus";
+
+Goal "[|summable f; 0 < k |] \
+\     ==> (%n. sumr (n*k) (n*k + k) f) sums (suminf f)";
+by (dtac summable_sums 1);
+by (auto_tac (claset(),simpset() addsimps [sums_def,LIMSEQ_def]));
+by (dres_inst_tac [("x","r")] spec 1 THEN Step_tac 1); 
+by (res_inst_tac [("x","no")] exI 1 THEN Step_tac 1); 
+by (dres_inst_tac [("x","n*k")] spec 1); 
+by (auto_tac (claset() addSDs [not_leE], simpset()));
+by (dres_inst_tac [("j","no")] less_le_trans 1);
+by (Auto_tac);
+qed "sums_group";
+
+Goal "[|summable f; ALL d. #0 < (f(n + (2 * d))) + f(n + ((2 * d) + 1))|] \
+\     ==> sumr 0 n f < suminf f";
+by (dtac summable_sums 1);
+by (auto_tac (claset(),simpset() addsimps [sums_def,LIMSEQ_def]));
+by (dres_inst_tac [("x","f(n) + f(n + 1)")] spec 1);
+by (Auto_tac);
+by (rtac ccontr 2 THEN dtac real_leI 2);
+by (subgoal_tac "sumr 0 (n + 2) f <= sumr 0 (2 * (Suc no) + n) f" 2);
+by (induct_tac "no" 3 THEN Simp_tac 3);
+by (res_inst_tac [("y","sumr 0 (2*(Suc na)+n) f")] order_trans 3);
+by (assume_tac 3);
+by (dres_inst_tac [("x","Suc na")] spec 3);
+by (dres_inst_tac [("x","0")] spec 1);
+by (Asm_full_simp_tac 1);
+by (asm_full_simp_tac (simpset() addsimps add_ac) 2);
+by (rotate_tac 1 1 THEN dres_inst_tac [("x","2 * (Suc no) + n")] spec 1);
+by (Step_tac 1 THEN Asm_full_simp_tac 1);
+by (subgoal_tac "suminf f + (f(n) + f(n + 1)) <= \
+\                 sumr 0 (2 * (Suc no) + n) f" 1);
+by (res_inst_tac [("y","sumr 0 (n+2) f")] order_trans 2);
+by (assume_tac 3);
+by (res_inst_tac [("y","sumr 0 n f + (f(n) + f(n + 1))")] order_trans 2);
+by (REPEAT(Asm_simp_tac 2));
+by (subgoal_tac "suminf f <= sumr 0 (2 * (Suc no) + n) f" 1);
+by (res_inst_tac [("y","suminf f + (f(n) + f(n + 1))")] order_trans 2);
+by (assume_tac 3);
+by (dres_inst_tac [("x","0")] spec 2);
+by (Asm_full_simp_tac 2);
+by (subgoal_tac "#0 <= sumr 0 (2 * Suc no + n) f + - suminf f" 1);
+by (dtac (rename_numerals abs_eqI1) 1 );
+by (Asm_full_simp_tac 1);
+by (auto_tac (claset(),simpset() addsimps [real_le_def]));
+qed "sumr_pos_lt_pair";
+
+(*-----------------------------------------------------------------
+   Summable series of positive terms has limit >= any partial sum 
+ ----------------------------------------------------------------*)
+Goal 
+     "[| summable f; ALL m. n <= m --> #0 <= f(m) |] \
+\          ==> sumr 0 n f <= suminf f";
+by (dtac summable_sums 1);
+by (rewtac sums_def);
+by (cut_inst_tac [("k","sumr 0 n f")] LIMSEQ_const 1);
+by (etac LIMSEQ_le 1 THEN Step_tac 1);
+by (res_inst_tac [("x","n")] exI 1 THEN Step_tac 1);
+by (dtac le_imp_less_or_eq 1 THEN Step_tac 1);
+by (auto_tac (claset() addIs [sumr_le], simpset()));
+qed "series_pos_le";
+
+Goal "[| summable f; ALL m. n <= m --> #0 < f(m) |] \
+\          ==> sumr 0 n f < suminf f";
+by (res_inst_tac [("y","sumr 0 (Suc n) f")] order_less_le_trans 1);
+by (rtac series_pos_le 2);
+by (Auto_tac);
+by (dres_inst_tac [("x","m")] spec 1);
+by (Auto_tac);
+qed "series_pos_less";
+
+(*-------------------------------------------------------------------
+                    sum of geometric progression                 
+ -------------------------------------------------------------------*)
+
+Goal "x ~= #1 ==> sumr 0 n (%n. x ^ n) = (x ^ n - #1) / (x - #1)";
+by (induct_tac "n" 1);
+by (Auto_tac);
+by (res_inst_tac [("c1","x - #1")] (real_mult_right_cancel RS iffD1) 1);
+by (auto_tac (claset(),
+       simpset() addsimps [real_mult_assoc, real_add_mult_distrib]));
+by (auto_tac (claset(),
+       simpset() addsimps [real_add_mult_distrib2,
+                           real_diff_def, real_mult_commute]));
+qed "sumr_geometric";
+
+Goal "abs(x) < #1 ==> (%n. x ^ n) sums (#1/(#1 - x))";
+by (case_tac "x = #1" 1);
+by (auto_tac (claset() addSDs [LIMSEQ_rabs_realpow_zero2],
+             simpset() addsimps [sumr_geometric ,sums_def,
+                                 real_diff_def, real_add_divide_distrib]));
+by (subgoal_tac "#1 / (#1 + - x) = #0/(x-#1) + - #1/(x-#1)" 1);
+by (asm_full_simp_tac (simpset() addsimps [real_divide_eq_cancel1,
+                 real_divide_minus_eq RS sym, real_diff_def]) 2); 
+by (etac ssubst 1); 
+by (rtac LIMSEQ_add 1 THEN rtac LIMSEQ_divide 1);
+by (auto_tac (claset() addIs [LIMSEQ_const], 
+              simpset() addsimps [real_diff_def, LIMSEQ_rabs_realpow_zero2]));
+qed "geometric_sums";
+
+(*-------------------------------------------------------------------
+    Cauchy-type criterion for convergence of series (c.f. Harrison)
+ -------------------------------------------------------------------*)
+Goalw [summable_def,sums_def,convergent_def]
+      "summable f = convergent (%n. sumr 0 n f)";
+by (Auto_tac);
+qed "summable_convergent_sumr_iff";
+
+Goal "summable f = \
+\     (ALL e. #0 < e --> (EX N. ALL m n. N <= m --> abs(sumr m n f) < e))";
+by (auto_tac (claset(),simpset() addsimps [summable_convergent_sumr_iff,
+    Cauchy_convergent_iff RS sym,Cauchy_def]));
+by (ALLGOALS(dtac spec) THEN Auto_tac);
+by (res_inst_tac [("x","M")] exI 1);
+by (res_inst_tac [("x","N")] exI 2);
+by (Auto_tac);
+by (ALLGOALS(cut_inst_tac [("m","m"),("n","n")] less_linear));
+by (Auto_tac);
+by (ftac (le_less_trans RS less_imp_le) 1 THEN assume_tac 1);
+by (dres_inst_tac [("x","n")] spec 1);
+by (dres_inst_tac [("x","m")] spec 1);
+by (auto_tac (claset() addIs [(abs_minus_add_cancel RS subst)],
+    simpset() addsimps [sumr_split_add_minus,abs_minus_add_cancel]));
+qed "summable_Cauchy";
+
+(*-------------------------------------------------------------------
+     Terms of a convergent series tend to zero
+     > Goalw [LIMSEQ_def] "summable f ==> f ----> #0";
+     Proved easily in HSeries after proving nonstandard case.
+ -------------------------------------------------------------------*)
+(*-------------------------------------------------------------------
+                    Comparison test
+ -------------------------------------------------------------------*)
+Goal "[| EX N. ALL n. N <= n --> abs(f n) <= g n; \
+\        summable g \
+\     |] ==> summable f";
+by (auto_tac (claset(),simpset() addsimps [summable_Cauchy]));
+by (dtac spec 1 THEN Auto_tac);
+by (res_inst_tac [("x","N + Na")] exI 1 THEN Auto_tac);
+by (rotate_tac 2 1);
+by (dres_inst_tac [("x","m")] spec 1);
+by (Auto_tac THEN rotate_tac 2 1 THEN dres_inst_tac [("x","n")] spec 1);
+by (res_inst_tac [("y","sumr m n (%k. abs(f k))")] order_le_less_trans 1);
+by (rtac sumr_rabs 1);
+by (res_inst_tac [("y","sumr m n g")] order_le_less_trans 1);
+by (auto_tac (claset() addIs [sumr_le2],
+    simpset() addsimps [abs_interval_iff]));
+qed "summable_comparison_test";
+
+Goal "[| EX N. ALL n. N <= n --> abs(f n) <= g n; \
+\        summable g \
+\     |] ==> summable (%k. abs (f k))";
+by (rtac summable_comparison_test 1);
+by (auto_tac (claset(),simpset() addsimps [abs_idempotent]));
+qed "summable_rabs_comparison_test";
+
+(*------------------------------------------------------------------*)
+(*       Limit comparison property for series (c.f. jrh)            *)
+(*------------------------------------------------------------------*)
+Goal "[|ALL n. f n <= g n; summable f; summable g |] \
+\     ==> suminf f <= suminf g";
+by (REPEAT(dtac summable_sums 1));
+by (auto_tac (claset() addSIs [LIMSEQ_le], simpset() addsimps [sums_def]));
+by (rtac exI 1);
+by (auto_tac (claset() addSIs [sumr_le2], simpset()));
+qed "summable_le";
+
+Goal "[|ALL n. abs(f n) <= g n; summable g |] \
+\     ==> summable f & suminf f <= suminf g";
+by (auto_tac (claset() addIs [summable_comparison_test]
+    addSIs [summable_le], simpset()));
+by (auto_tac (claset(),simpset() addsimps [abs_le_interval_iff]));
+qed "summable_le2";
+
+(*-------------------------------------------------------------------
+         Absolute convergence imples normal convergence
+ -------------------------------------------------------------------*)
+Goal "summable (%n. abs (f n)) ==> summable f";
+by (auto_tac (claset(),simpset() addsimps [sumr_rabs,summable_Cauchy]));
+by (dtac spec 1 THEN Auto_tac);
+by (res_inst_tac [("x","N")] exI 1 THEN Auto_tac);
+by (dtac spec 1 THEN Auto_tac);
+by (res_inst_tac [("y","sumr m n (%n. abs(f n))")] order_le_less_trans 1);
+by (auto_tac (claset() addIs [sumr_rabs], simpset()));
+qed "summable_rabs_cancel";
+
+(*-------------------------------------------------------------------
+         Absolute convergence of series
+ -------------------------------------------------------------------*)
+Goal "summable (%n. abs (f n)) ==> abs(suminf f) <= suminf (%n. abs(f n))";
+by (auto_tac (claset() addIs [LIMSEQ_le,LIMSEQ_imp_rabs,summable_rabs_cancel,
+                              summable_sumr_LIMSEQ_suminf,sumr_rabs], 
+              simpset()));
+qed "summable_rabs";
+
+(*-------------------------------------------------------------------
+                        The ratio test
+ -------------------------------------------------------------------*)
+Goal "[| c <= #0; abs x <= c * abs y |] ==> x = (#0::real)";
+by (dtac order_le_imp_less_or_eq 1);
+by Auto_tac;
+by (dres_inst_tac [("x1","y")] 
+    (abs_ge_zero RS rename_numerals real_mult_le_zero) 1);
+by (asm_full_simp_tac (simpset() addsimps [real_mult_commute]) 1);
+by (dtac order_trans 1 THEN assume_tac 1);
+by (TRYALL(arith_tac));
+qed "rabs_ratiotest_lemma";
+
+(* lemmas *)
+
+Goal "(k::nat) <= l ==> (EX n. l = k + n)";
+by (dtac le_imp_less_or_eq 1);
+by (auto_tac (claset() addDs [less_imp_Suc_add], simpset()));
+qed "le_Suc_ex";
+
+Goal "((k::nat) <= l) = (EX n. l = k + n)";
+by (auto_tac (claset(),simpset() addsimps [le_Suc_ex]));
+qed "le_Suc_ex_iff";
+
+(*All this trouble just to get #0<c *)
+Goal "[| ALL n. N <= n --> abs(f(Suc n)) <= c*abs(f n) |] \
+\     ==> #0 < c | summable f";
+by (simp_tac (simpset() addsimps [linorder_not_le RS sym]) 1); 
+by (asm_full_simp_tac (simpset() addsimps [summable_Cauchy]) 1);
+by (Step_tac 1 THEN subgoal_tac "ALL n. N <= n --> f (Suc n) = #0" 1);
+by (blast_tac (claset() addIs [rabs_ratiotest_lemma]) 2);
+by (res_inst_tac [("x","Suc N")] exI 1 THEN Step_tac 1);
+by (dtac Suc_le_imp_diff_ge2 1 THEN Auto_tac);
+qed "ratio_test_lemma2";
+
+
+Goal "[| c < #1; ALL n. N <= n --> abs(f(Suc n)) <= c*abs(f n) |] \
+\     ==> summable f";
+by (forward_tac [ratio_test_lemma2] 1);
+by Auto_tac;  
+by (res_inst_tac [("g","%n. (abs (f N)/(c ^ N))*c ^ n")] 
+    summable_comparison_test 1);
+by (res_inst_tac [("x","N")] exI 1 THEN Step_tac 1);
+by (dtac (le_Suc_ex_iff RS iffD1) 1);
+by (auto_tac (claset(),simpset() addsimps [realpow_add,
+    rename_numerals realpow_not_zero]));
+by (induct_tac "na" 1 THEN Auto_tac);
+by (res_inst_tac [("y","c*abs(f(N + n))")] order_trans 1);
+by (auto_tac (claset() addIs [real_mult_le_le_mono1],
+    simpset() addsimps [summable_def]));
+by (asm_full_simp_tac (simpset() addsimps real_mult_ac) 1);
+by (res_inst_tac [("x","abs(f N) * (#1/(#1 - c)) / (c ^ N)")] exI 1);
+by (rtac sums_divide 1); 
+by (rtac sums_mult 1); 
+by (auto_tac (claset() addSIs [sums_mult,geometric_sums], 
+              simpset() addsimps [real_abs_def]));
+qed "ratio_test";
+
+
+(*----------------------------------------------------------------------------*)
+(* Differentiation of finite sum                                              *)
+(*----------------------------------------------------------------------------*)
+
+Goal "(ALL r. m <= r & r < (m + n) --> DERIV (%x. f r x) x :> (f' r x)) \
+\     --> DERIV (%x. sumr m n (%n. f n x)) x :> sumr m n (%r. f' r x)";
+by (induct_tac "n" 1);
+by (auto_tac (claset() addIs [DERIV_add], simpset()));
+qed "DERIV_sumr";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/Series.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,27 @@
+(*  Title       : Series.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Finite summation and infinite series
+*) 
+
+
+Series = SEQ + Lim +
+
+consts sumr :: "[nat,nat,(nat=>real)] => real"
+primrec
+   sumr_0   "sumr m 0 f = #0"
+   sumr_Suc "sumr m (Suc n) f = (if n < m then #0 
+                               else sumr m n f + f(n))"
+
+constdefs
+   sums  :: [nat=>real,real] => bool     (infixr 80)
+   "f sums s  == (%n. sumr 0 n f) ----> s"
+
+   summable :: (nat=>real) => bool
+   "summable f == (EX s. f sums s)"
+
+   suminf   :: (nat=>real) => real
+   "suminf f == (@s. f sums s)"
+end
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/Star.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,492 @@
+(*  Title       : STAR.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : *-transforms
+*) 
+
+(*--------------------------------------------------------
+   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) 
+*)
+Goal "ALL x. EX y. Q x y ==> EX (f :: nat => nat). ALL x. Q x (f x)";
+by (res_inst_tac [("x","%x. LEAST y. Q x y")] exI 1);
+by (blast_tac (claset() addIs [LeastI]) 1);
+qed "no_choice";
+
+(*------------------------------------------------------------
+    Properties of the *-transform applied to sets of reals
+ ------------------------------------------------------------*)
+
+Goalw [starset_def] "*s*(UNIV::real set) = (UNIV::hypreal set)";
+by (Auto_tac);
+qed "STAR_real_set";
+Addsimps [STAR_real_set];
+
+Goalw [starset_def] "*s* {} = {}";
+by (Step_tac 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (dres_inst_tac [("x","%n. xa n")] bspec 1);
+by (Auto_tac);
+qed "STAR_empty_set";
+Addsimps [STAR_empty_set];
+
+Goalw [starset_def] "*s* (A Un B) = *s* A Un *s* B";
+by (Auto_tac);
+by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2));
+by (dtac FreeUltrafilterNat_Compl_mem 1);
+by (dtac bspec 1 THEN assume_tac 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (Auto_tac);
+by (Fuf_tac 1);
+qed "STAR_Un";
+
+Goalw [starset_def] "*s* (A Int B) = *s* A Int *s* B";
+by (Auto_tac);
+by (blast_tac (claset() addIs [FreeUltrafilterNat_Int,
+               FreeUltrafilterNat_subset]) 3);
+by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1));
+qed "STAR_Int";
+
+Goalw [starset_def] "*s* -A = -(*s* A)";
+by (Auto_tac);
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 2);
+by (REPEAT(Step_tac 1) THEN Auto_tac);
+by (Fuf_empty_tac 1);
+by (dtac FreeUltrafilterNat_Compl_mem 1);
+by (Fuf_tac 1);
+qed "STAR_Compl";
+
+goal Set.thy "(A - B) = (A Int (- B))";
+by (Step_tac 1);
+qed "set_diff_iff2";
+
+Goal "x ~: *s* F ==> x : *s* (- F)";
+by (auto_tac (claset(),simpset() addsimps [STAR_Compl]));
+qed "STAR_mem_Compl";
+
+Goal "*s* (A - B) = *s* A - *s* B";
+by (auto_tac (claset(),simpset() addsimps 
+         [set_diff_iff2,STAR_Int,STAR_Compl]));
+qed "STAR_diff";
+
+Goalw [starset_def] "A <= B ==> *s* A <= *s* B";
+by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1));
+qed "STAR_subset";
+
+Goalw [starset_def,hypreal_of_real_def] 
+          "a : A ==> hypreal_of_real a : *s* A";
+by (auto_tac (claset() addIs [FreeUltrafilterNat_subset],simpset()));
+qed "STAR_mem";
+
+Goalw [starset_def] "hypreal_of_real `` A <= *s* A";
+by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_def]));
+by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1);
+qed "STAR_hypreal_of_real_image_subset";
+
+Goalw [starset_def] "*s* X Int SReal = hypreal_of_real `` X";
+by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_def,SReal_def]));
+by (fold_tac [hypreal_of_real_def]);
+by (rtac imageI 1 THEN rtac ccontr 1);
+by (dtac bspec 1);
+by (rtac lemma_hyprel_refl 1);
+by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2);
+by (Auto_tac);
+qed "STAR_hypreal_of_real_Int";
+
+Goal "x ~: hypreal_of_real `` A ==> ALL y: A. x ~= hypreal_of_real y";
+by (Auto_tac);
+qed "lemma_not_hyprealA";
+
+Goal "- {n. X n = xa} = {n. X n ~= xa}";
+by (Auto_tac);
+qed "lemma_Compl_eq";
+
+Goalw [starset_def]
+    "ALL n. (X n) ~: M \
+\         ==> Abs_hypreal(hyprel^^{X}) ~: *s* M";
+by (Auto_tac THEN rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
+by (Auto_tac);
+qed "STAR_real_seq_to_hypreal";
+
+Goalw [starset_def] "*s* {x} = {hypreal_of_real x}";
+by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_def]));
+by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addIs [FreeUltrafilterNat_subset],simpset()));
+qed "STAR_singleton";
+Addsimps [STAR_singleton];
+
+Goal "x ~: F ==> hypreal_of_real x ~: *s* F";
+by (auto_tac (claset(),simpset() addsimps
+    [starset_def,hypreal_of_real_def]));
+by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
+by (Auto_tac);
+qed "STAR_not_mem";
+
+Goal "[| x : *s* A; A <= B |] ==> x : *s* B";
+by (blast_tac (claset() addDs [STAR_subset]) 1);
+qed "STAR_subset_closed";
+
+(*------------------------------------------------------------------ 
+   Nonstandard extension of a set (defined using a constant 
+   sequence) as a special case of an internal set
+ -----------------------------------------------------------------*)
+
+Goalw [starset_n_def,starset_def] 
+     "ALL n. (As n = A) ==> *sn* As = *s* A";
+by (Auto_tac);
+qed "starset_n_starset";
+
+
+(*----------------------------------------------------------------*)
+(* Theorems about nonstandard extensions of functions             *)
+(*----------------------------------------------------------------*)
+
+(*----------------------------------------------------------------*) 
+(* Nonstandard extension of a function (defined using a           *)
+(* constant sequence) as a special case of an internal function   *)
+(*----------------------------------------------------------------*)
+
+Goalw [starfun_n_def,starfun_def] 
+     "ALL n. (F n = f) ==> *fn* F = *f* f";
+by (Auto_tac);
+qed "starfun_n_starfun";
+
+
+(* 
+   Prove that hrabs is a nonstandard extension of rabs without
+   use of congruence property (proved after this for general
+   nonstandard extensions of real valued functions). This makes 
+   proof much longer- see comments at end of HREALABS.thy where
+   we proved a congruence theorem for hrabs. 
+
+   NEW!!! No need to prove all the lemmas anymore. Use the ultrafilter
+   tactic! 
+*)
+  
+Goalw [is_starext_def] "is_starext abs abs";
+by (Step_tac 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by Auto_tac;
+by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
+by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
+by (auto_tac (claset() addSDs [spec],
+              simpset() addsimps [hypreal_minus,hrabs_def,
+                rename_numerals hypreal_zero_def,
+                hypreal_le_def, hypreal_less_def]));
+by (TRYALL(Ultra_tac));
+by (TRYALL(arith_tac));
+qed "hrabs_is_starext_rabs";
+
+Goal "[| X: Rep_hypreal z; Y: Rep_hypreal z |] \
+\     ==> {n. X n = Y n} : FreeUltrafilterNat";
+by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
+by (Auto_tac THEN Fuf_tac 1);
+qed "Rep_hypreal_FreeUltrafilterNat";
+
+(*-----------------------------------------------------------------------
+    Nonstandard extension of functions- congruence 
+ -----------------------------------------------------------------------*) 
+
+Goalw [congruent_def] "congruent hyprel (%X. hyprel^^{%n. f (X n)})";
+by (safe_tac (claset()));
+by (ALLGOALS(Fuf_tac));
+qed "starfun_congruent";
+
+Goalw [starfun_def]
+      "(*f* f) (Abs_hypreal(hyprel^^{%n. X n})) = \
+\      Abs_hypreal(hyprel ^^ {%n. f (X n)})";
+by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
+by (simp_tac (simpset() addsimps 
+   [hyprel_in_hypreal RS Abs_hypreal_inverse,[equiv_hyprel,
+   starfun_congruent] MRS UN_equiv_class]) 1);
+qed "starfun";
+
+(*-------------------------------------------
+  multiplication: ( *f ) x ( *g ) = *(f x g)  
+ ------------------------------------------*)
+Goal "(*f* f) xa * (*f* g) xa = (*f* (%x. f x * g x)) xa";
+by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [starfun,hypreal_mult]));
+qed "starfun_mult";
+Addsimps [starfun_mult RS sym];
+
+(*---------------------------------------
+  addition: ( *f ) + ( *g ) = *(f + g)  
+ ---------------------------------------*)
+Goal "(*f* f) xa + (*f* g) xa = (*f* (%x. f x + g x)) xa";
+by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [starfun,hypreal_add]));
+qed "starfun_add";
+Addsimps [starfun_add RS sym];
+
+(*--------------------------------------------
+  subtraction: ( *f ) + -( *g ) = *(f + -g)  
+ -------------------------------------------*)
+
+Goal "- (*f* f) x = (*f* (%x. - f x)) x";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [starfun, hypreal_minus]));
+qed "starfun_minus";
+Addsimps [starfun_minus RS sym];
+
+(*FIXME: delete*)
+Goal "(*f* f) xa + -(*f* g) xa = (*f* (%x. f x + -g x)) xa";
+by (Simp_tac 1);
+qed "starfun_add_minus";
+Addsimps [starfun_add_minus RS sym];
+
+Goalw [hypreal_diff_def,real_diff_def]
+  "(*f* f) xa  - (*f* g) xa = (*f* (%x. f x - g x)) xa";
+by (rtac starfun_add_minus 1);
+qed "starfun_diff";
+Addsimps [starfun_diff RS sym];
+
+(*--------------------------------------
+  composition: ( *f ) o ( *g ) = *(f o g)  
+ ---------------------------------------*)
+
+Goal "(%x. (*f* f) ((*f* g) x)) = *f* (%x. f (g x))"; 
+by (rtac ext 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [starfun]));
+qed "starfun_o2";
+
+Goalw [o_def] "(*f* f) o (*f* g) = (*f* (f o g))";
+by (simp_tac (simpset() addsimps [starfun_o2]) 1);
+qed "starfun_o";
+
+(*--------------------------------------
+  NS extension of constant function
+ --------------------------------------*)
+Goal "(*f* (%x. k)) xa = hypreal_of_real  k";
+by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [starfun,
+    hypreal_of_real_def]));
+qed "starfun_const_fun";
+
+Addsimps [starfun_const_fun];
+
+(*----------------------------------------------------
+   the NS extension of the identity function
+ ----------------------------------------------------*)
+
+Goal "x @= hypreal_of_real a ==> (*f* (%x. x)) x @= hypreal_of_real  a";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [starfun]));
+qed "starfun_Idfun_inf_close";
+
+Goal "(*f* (%x. x)) x = x";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [starfun]));
+qed "starfun_Id";
+Addsimps [starfun_Id];  
+
+(*----------------------------------------------------------------------
+      the *-function is a (nonstandard) extension of the function
+ ----------------------------------------------------------------------*)
+
+Goalw [is_starext_def] "is_starext (*f* f) f";
+by (Auto_tac);
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addSIs [bexI] ,simpset() addsimps [starfun]));
+qed "is_starext_starfun";
+
+(*----------------------------------------------------------------------
+     Any nonstandard extension is in fact the *-function
+ ----------------------------------------------------------------------*)
+
+Goalw [is_starext_def] "is_starext F f ==> F = *f* f";
+by (rtac ext 1);
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (dres_inst_tac [("x","x")] spec 1);
+by (dres_inst_tac [("x","(*f* f) x")] spec 1);
+by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem],
+    simpset() addsimps [starfun]));
+by (Fuf_empty_tac 1);
+qed "is_starfun_starext";
+
+Goal "(is_starext F f) = (F = *f* f)";
+by (blast_tac (claset() addIs [is_starfun_starext,is_starext_starfun]) 1);
+qed "is_starext_starfun_iff";
+
+(*--------------------------------------------------------
+   extented function has same solution as its standard
+   version for real arguments. i.e they are the same
+   for all real arguments
+ -------------------------------------------------------*)
+Goal "(*f* f) (hypreal_of_real a) = hypreal_of_real (f a)";
+by (auto_tac (claset(),simpset() addsimps 
+     [starfun,hypreal_of_real_def]));
+qed "starfun_eq";
+
+Addsimps [starfun_eq];
+
+Goal "(*f* f) (hypreal_of_real a) @= hypreal_of_real (f a)";
+by (Auto_tac);
+qed "starfun_inf_close";
+
+(* useful for NS definition of derivatives *)
+Goal "(*f* (%h. f (x + h))) xa  = (*f* f) (hypreal_of_real  x + xa)";
+by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [starfun,
+    hypreal_of_real_def,hypreal_add]));
+qed "starfun_lambda_cancel";
+
+Goal "(*f* (%h. f(g(x + h)))) xa = (*f* (f o g)) (hypreal_of_real x + xa)";
+by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [starfun,
+    hypreal_of_real_def,hypreal_add]));
+qed "starfun_lambda_cancel2";
+
+Goal "[| (*f* f) xa @= l; (*f* g) xa @= m; \
+\                 l: HFinite; m: HFinite  \
+\              |] ==>  (*f* (%x. f x * g x)) xa @= l * m";
+by (dtac inf_close_mult_HFinite 1);
+by (REPEAT(assume_tac 1));
+by (auto_tac (claset() addIs [inf_close_sym RSN (2,inf_close_HFinite)],
+              simpset()));
+qed "starfun_mult_HFinite_inf_close";
+
+Goal "[| (*f* f) xa @= l; (*f* g) xa @= m \
+\              |] ==>  (*f* (%x. f x + g x)) xa @= l + m";
+by (auto_tac (claset() addIs [inf_close_add], simpset()));
+qed "starfun_add_inf_close";
+
+(*----------------------------------------------------
+    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          *)
+
+Goal "*f* abs = abs";
+by (rtac (hrabs_is_starext_rabs RS 
+          (is_starext_starfun_iff RS iffD1) RS sym) 1);
+qed "starfun_rabs_hrabs";
+
+Goal "(*f* inverse) x = inverse(x)";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+            simpset() addsimps [starfun, hypreal_inverse, hypreal_zero_def]));
+qed "starfun_inverse_inverse";
+Addsimps [starfun_inverse_inverse];
+
+Goal "inverse ((*f* f) x) = (*f* (%x. inverse (f x))) x";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),
+              simpset() addsimps [starfun, hypreal_inverse]));
+qed "starfun_inverse";
+Addsimps [starfun_inverse RS sym];
+
+Goalw [hypreal_divide_def,real_divide_def]
+  "(*f* f) xa  / (*f* g) xa = (*f* (%x. f x / g x)) xa";
+by Auto_tac;
+qed "starfun_divide";
+Addsimps [starfun_divide RS sym];
+
+Goal "inverse ((*f* f) x) = (*f* (%x. inverse (f x))) x";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addIs [FreeUltrafilterNat_subset]
+                       addSDs [FreeUltrafilterNat_Compl_mem],
+    simpset() addsimps [starfun, hypreal_inverse, hypreal_zero_def]));
+qed "starfun_inverse2";
+
+(*-------------------------------------------------------------
+    General lemma/theorem needed for proofs in elementary
+    topology of the reals
+ ------------------------------------------------------------*)
+Goalw [starset_def] 
+      "(*f* f) x : *s* A ==> x : *s* {x. f x : A}";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps [starfun]));
+by (dres_inst_tac [("x","%n. f (Xa n)")] bspec 1);
+by (Auto_tac THEN Fuf_tac 1);
+qed "starfun_mem_starset";
+
+(*------------------------------------------------------------
+   Alternative definition for hrabs with rabs function
+   applied entrywise to equivalence class representative.
+   This is easily proved using starfun and ns extension thm
+ ------------------------------------------------------------*)
+Goal "abs (Abs_hypreal (hyprel ^^ {X})) = \
+\                 Abs_hypreal(hyprel ^^ {%n. abs (X n)})";
+by (simp_tac (simpset() addsimps [starfun_rabs_hrabs RS sym,starfun]) 1);
+qed "hypreal_hrabs";
+
+(*----------------------------------------------------------------
+   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 ( *f* f). See second NS set extension below.
+ ----------------------------------------------------------------*)
+Goalw [starset_def]
+   "*s* {x. abs (x + - y) < r} = \
+\    {x. abs(x + -hypreal_of_real y) < hypreal_of_real r}";
+by (Step_tac 1);
+by (ALLGOALS(res_inst_tac [("z","x")] eq_Abs_hypreal));
+by (auto_tac (claset() addSIs [exI] addSDs [bspec],
+          simpset() addsimps [hypreal_minus, hypreal_of_real_def,hypreal_add,
+                              hypreal_hrabs,hypreal_less_def]));
+by (Fuf_tac 1);
+qed "STAR_rabs_add_minus";
+
+Goalw [starset_def]
+  "*s* {x. abs (f x + - y) < r} = \
+\      {x. abs((*f* f) x + -hypreal_of_real y) < hypreal_of_real r}";
+by (Step_tac 1);
+by (ALLGOALS(res_inst_tac [("z","x")] eq_Abs_hypreal));
+by (auto_tac (claset() addSIs [exI] addSDs [bspec],
+    simpset() addsimps [hypreal_minus, hypreal_of_real_def,hypreal_add,
+    hypreal_hrabs,hypreal_less_def,starfun]));
+by (Fuf_tac 1);
+qed "STAR_starfun_rabs_add_minus";
+
+(*-------------------------------------------------------------------
+   Another charaterization of Infinitesimal and one of @= relation. 
+   In this theory since hypreal_hrabs proved here. (To Check:) Maybe 
+   move both if possible? 
+ -------------------------------------------------------------------*)
+Goal "(x:Infinitesimal) = \
+\     (EX X:Rep_hypreal(x). \
+\       ALL m. {n. abs(X n) < inverse(real_of_posnat m)}:FreeUltrafilterNat)";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (auto_tac (claset() addSIs [bexI,lemma_hyprel_refl],
+	simpset() addsimps [Infinitesimal_hypreal_of_posnat_iff,
+	    hypreal_of_real_of_posnat,hypreal_of_real_def,hypreal_inverse,
+	    hypreal_hrabs,hypreal_less])); 
+by (dres_inst_tac [("x","n")] spec 1);
+by (Fuf_tac 1);
+qed  "Infinitesimal_FreeUltrafilterNat_iff2";
+
+Goal "(Abs_hypreal(hyprel^^{X}) @= Abs_hypreal(hyprel^^{Y})) = \
+\     (ALL m. {n. abs (X n + - Y n) < \
+\                 inverse(real_of_posnat m)} : FreeUltrafilterNat)";
+by (rtac (inf_close_minus_iff RS ssubst) 1);
+by (rtac (mem_infmal_iff RS subst) 1);
+by (auto_tac (claset(), 
+              simpset() addsimps [hypreal_minus, hypreal_add,
+                                  Infinitesimal_FreeUltrafilterNat_iff2]));
+by (dres_inst_tac [("x","m")] spec 1);
+by (Fuf_tac 1);
+qed "inf_close_FreeUltrafilterNat_iff";
+
+Goal "inj starfun";
+by (rtac injI 1);
+by (rtac ext 1 THEN rtac ccontr 1);
+by (dres_inst_tac [("x","Abs_hypreal(hyprel ^^{%n. xa})")] fun_cong 1);
+by (auto_tac (claset(),simpset() addsimps [starfun]));
+qed "inj_starfun";
+
+
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/Star.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,39 @@
+(*  Title       : Star.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : defining *-transforms in NSA which extends sets of reals, 
+                  and real=>real functions
+*) 
+
+Star = NSA +
+
+constdefs
+    (* nonstandard extension of sets *)
+    starset :: real set => hypreal set          ("*s* _" [80] 80)
+    "*s* A  == {x. ALL X: Rep_hypreal(x). {n::nat. X n : A}: FreeUltrafilterNat}"
+
+    (* internal sets *)
+    starset_n :: (nat => real set) => hypreal set        ("*sn* _" [80] 80)
+    "*sn* As  == {x. ALL X: Rep_hypreal(x). {n::nat. X n : (As n)}: FreeUltrafilterNat}"   
+    
+    InternalSets :: "hypreal set set"
+    "InternalSets == {X. EX As. X = *sn* As}"
+
+    (* nonstandard extension of function *)
+    is_starext  ::  [hypreal => hypreal, real => real] => bool
+    "is_starext F f == (ALL x y. EX X: Rep_hypreal(x). EX Y: Rep_hypreal(y).
+                        ((y = (F x)) = ({n. Y n = f(X n)} : FreeUltrafilterNat)))"
+    
+    starfun :: (real => real) => hypreal => hypreal        ("*f* _" [80] 80)
+    "*f* f  == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel^^{%n. f (X n)}))" 
+
+    (* internal functions *)
+    starfun_n :: (nat => (real => real)) => hypreal => hypreal        ("*fn* _" [80] 80)
+    "*fn* F  == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel^^{%n. (F n)(X n)}))" 
+
+    InternalFuns :: (hypreal => hypreal) set
+    "InternalFuns == {X. EX F. X = *fn* F}"
+end
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/Zorn.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,285 @@
+(*  Title       : Zorn.ML
+    ID          : $Id$
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Zorn's Lemma -- adapted proofs from lcp's ZF/Zorn.ML
+*) 
+
+(*---------------------------------------------------------------
+      Section 1.  Mathematical Preamble 
+ ---------------------------------------------------------------*)
+
+Goal "(ALL x:C. x<=A | B<=x) ==> Union(C)<=A | B<=Union(C)";
+by (Blast_tac 1);
+qed "Union_lemma0";
+
+(*-- similar to subset_cs in ZF/subset.thy --*)
+bind_thms ("thissubset_SIs",
+    [subset_refl,Union_least, UN_least, Un_least, 
+     Inter_greatest, Int_greatest,
+     Un_upper1, Un_upper2, Int_lower1, Int_lower2]);
+
+
+(*A claset for subset reasoning*)
+val thissubset_cs = claset() 
+    delrules [subsetI, subsetCE]
+    addSIs thissubset_SIs
+    addIs  [Union_upper, Inter_lower];
+
+(* increasingD2 of ZF/Zorn.ML *) 
+Goalw [succ_def] "x <= succ S x";
+by (rtac (split_if RS iffD2) 1);
+by (auto_tac (claset(),simpset() addsimps [super_def,
+               maxchain_def,psubset_def]));
+by (rtac swap 1 THEN assume_tac 1);
+by (rtac someI2 1);
+by (ALLGOALS(Blast_tac));
+qed "Abrial_axiom1";
+
+val [TFin_succI, Pow_TFin_UnionI] = TFin.intrs;
+val TFin_UnionI = PowI RS Pow_TFin_UnionI;
+bind_thm ("TFin_succI", TFin_succI);
+bind_thm ("Pow_TFin_UnionI", Pow_TFin_UnionI);
+bind_thm ("TFin_UnionI", TFin_UnionI);
+
+val major::prems = Goal  
+          "[| n : TFin S; \
+\            !!x. [| x: TFin S; P(x) |] ==> P(succ S x); \
+\            !!Y. [| Y <= TFin S; Ball Y P |] ==> P(Union Y) |] \
+\         ==> P(n)";
+by (rtac (major RS TFin.induct) 1);
+by (ALLGOALS (fast_tac (claset() addIs prems)));
+qed "TFin_induct";
+
+(*Perform induction on n, then prove the major premise using prems. *)
+fun TFin_ind_tac a prems i = 
+    EVERY [induct_thm_tac TFin_induct a i,
+           rename_last_tac a ["1"] (i+1),
+           rename_last_tac a ["2"] (i+2),
+           ares_tac prems i];
+
+Goal "x <= y ==> x <= succ S y";
+by (etac (Abrial_axiom1 RSN (2,subset_trans)) 1);
+qed "succ_trans";
+
+(*Lemma 1 of section 3.1*)
+Goal "[| n: TFin S;  m: TFin S;  \
+\        ALL x: TFin S. x <= m --> x = m | succ S x <= m \
+\     |] ==> n <= m | succ S m <= n";
+by (etac TFin_induct 1);
+by (etac Union_lemma0 2);               (*or just Blast_tac*)
+by (blast_tac (thissubset_cs addIs [succ_trans]) 1);
+qed "TFin_linear_lemma1";
+
+(* Lemma 2 of section 3.2 *)
+Goal "m: TFin S ==> ALL n: TFin S. n<=m --> n=m | succ S n<=m";
+by (etac TFin_induct 1);
+by (rtac (impI RS ballI) 1);
+(*case split using TFin_linear_lemma1*)
+by (res_inst_tac [("n1","n"), ("m1","x")] 
+    (TFin_linear_lemma1 RS disjE) 1  THEN  REPEAT (assume_tac 1));
+by (dres_inst_tac [("x","n")] bspec 1 THEN assume_tac 1);
+by (blast_tac (thissubset_cs addIs [succ_trans]) 1);
+by (REPEAT (ares_tac [disjI1,equalityI] 1));
+(*second induction step*)
+by (rtac (impI RS ballI) 1);
+by (rtac (Union_lemma0 RS disjE) 1);
+by (rtac disjI2 3);
+by (REPEAT (ares_tac [disjI1,equalityI] 2));
+by (rtac ballI 1);
+by (ball_tac 1);
+by (set_mp_tac 1);
+by (res_inst_tac [("n1","n"), ("m1","x")] 
+    (TFin_linear_lemma1 RS disjE) 1  THEN  REPEAT (assume_tac 1));
+by (blast_tac thissubset_cs 1);
+by (rtac (Abrial_axiom1 RS subset_trans RS disjI1) 1);
+by (assume_tac 1);
+qed "TFin_linear_lemma2";
+
+(*a more convenient form for Lemma 2*)
+Goal "[| n<=m;  m: TFin S;  n: TFin S |] ==> n=m | succ S n<=m";
+by (rtac (TFin_linear_lemma2 RS bspec RS mp) 1);
+by (REPEAT (assume_tac 1));
+qed "TFin_subsetD";
+
+(*Consequences from section 3.3 -- Property 3.2, the ordering is total*)
+Goal "[| m: TFin S;  n: TFin S|] ==> n<=m | m<=n";
+by (rtac (TFin_linear_lemma2 RSN (3,TFin_linear_lemma1) RS disjE) 1);
+by (REPEAT (assume_tac 1) THEN etac disjI2 1);
+by (blast_tac (thissubset_cs addIs [Abrial_axiom1 RS subset_trans]) 1);
+qed "TFin_subset_linear";
+
+(*Lemma 3 of section 3.3*)
+Goal "[| n: TFin S;  m: TFin S;  m = succ S m |] ==> n<=m";
+by (etac TFin_induct 1);
+by (dtac TFin_subsetD 1);
+by (REPEAT (assume_tac 1));
+by (fast_tac (claset() addEs [ssubst]) 1);
+by (blast_tac (thissubset_cs) 1);
+qed "eq_succ_upper";
+
+(*Property 3.3 of section 3.3*)
+Goal "m: TFin S ==> (m = succ S m) = (m = Union(TFin S))";
+by (rtac iffI 1);
+by (rtac (Union_upper RS equalityI) 1);
+by (rtac (eq_succ_upper RS Union_least) 2);
+by (REPEAT (assume_tac 1));
+by (etac ssubst 1);
+by (rtac (Abrial_axiom1 RS equalityI) 1);
+by (blast_tac (thissubset_cs addIs [TFin_UnionI, TFin_succI]) 1);
+qed "equal_succ_Union";
+
+(*-------------------------------------------------------------------------
+    Section 4.  Hausdorff's Theorem: every set contains a maximal chain 
+    NB: We assume the partial ordering is <=, the subset relation! 
+ -------------------------------------------------------------------------*)
+
+Goalw [chain_def] "({} :: 'a set set) : chain S";
+by (Auto_tac);
+qed "empty_set_mem_chain";
+
+Goalw [super_def] "super S c <= chain S";
+by (Fast_tac 1);
+qed "super_subset_chain";
+
+Goalw [maxchain_def] "maxchain S <= chain S";
+by (Fast_tac 1);
+qed "maxchain_subset_chain";
+
+Goalw [succ_def] "c ~: chain S ==> succ S c = c";
+by (fast_tac (claset() addSIs [if_P]) 1);
+qed "succI1";
+
+Goalw [succ_def] "c: maxchain S ==> succ S c = c";
+by (fast_tac (claset() addSIs [if_P]) 1);
+qed "succI2";
+
+Goalw [succ_def] "c: chain S - maxchain S ==> \
+\                         succ S c = (@c'. c': super S c)";
+by (fast_tac (claset() addSIs [if_not_P]) 1);
+qed "succI3";
+
+Goal "c: chain S - maxchain S ==> ? d. d: super S c";
+by (rewrite_goals_tac [super_def,maxchain_def]);
+by (Auto_tac);
+qed "mem_super_Ex";
+
+Goal "c: chain S - maxchain S ==> \
+\                         (@c'. c': super S c): super S c";
+by (etac (mem_super_Ex RS exE) 1);
+by (rtac someI2 1);
+by (Auto_tac);
+qed "select_super";
+
+Goal "c: chain S - maxchain S ==> \
+\                         (@c'. c': super S c) ~= c";
+by (rtac notI 1);
+by (dtac select_super 1);
+by (asm_full_simp_tac (simpset() addsimps [super_def,psubset_def]) 1);
+qed "select_not_equals";
+
+Goal "c: chain S - maxchain S ==> \
+\                         succ S c ~= c";
+by (ftac succI3 1);
+by (Asm_simp_tac 1);
+by (rtac select_not_equals 1);
+by (assume_tac 1);
+qed "succ_not_equals";
+
+Goal "c: TFin S ==> (c :: 'a set set): chain S";
+by (etac TFin_induct 1);
+by (asm_simp_tac (simpset() addsimps [succ_def,
+    select_super RS (super_subset_chain RS subsetD)]
+                   addsplits [split_if]) 1);
+by (rewtac chain_def);
+by (rtac CollectI 1);
+by Safe_tac;
+by (dtac bspec 1 THEN assume_tac 1);
+by (res_inst_tac  [("m1","Xa"), ("n1","X")] (TFin_subset_linear RS disjE) 2);
+by (ALLGOALS(Blast_tac));
+qed "TFin_chain_lemm4";
+ 
+Goal "EX c. (c :: 'a set set): maxchain S";
+by (res_inst_tac [("x", "Union(TFin S)")] exI 1);
+by (rtac classical 1);
+by (subgoal_tac "succ S (Union(TFin S)) = Union(TFin S)" 1);
+by (resolve_tac [equal_succ_Union RS iffD2 RS sym] 2);
+by (resolve_tac [subset_refl RS TFin_UnionI] 2);
+by (rtac refl 2);
+by (cut_facts_tac [subset_refl RS TFin_UnionI RS TFin_chain_lemm4] 1);
+by (dtac (DiffI RS succ_not_equals) 1);
+by (ALLGOALS(Blast_tac));
+qed "Hausdorff";
+
+
+(*---------------------------------------------------------------
+  Section 5.  Zorn's Lemma: if all chains have upper bounds 
+                               there is  a maximal element 
+ ----------------------------------------------------------------*)
+Goalw [chain_def]
+    "[| c: chain S; z: S; \
+\             ALL x:c. x<=(z:: 'a set) |] ==> {z} Un c : chain S";
+by (Blast_tac 1);
+qed "chain_extend";
+
+Goalw [chain_def] "[| c: chain S; x: c |] ==> x <= Union(c)";
+by (Auto_tac);
+qed "chain_Union_upper";
+
+Goalw [chain_def] "c: chain S ==> ! x: c. x <= Union(c)";
+by (Auto_tac);
+qed "chain_ball_Union_upper";
+
+Goal "[| c: maxchain S; u: S; Union(c) <= u |] ==> Union(c) = u";
+by (rtac ccontr 1);
+by (asm_full_simp_tac (simpset() addsimps [maxchain_def]) 1);
+by (etac conjE 1);
+by (subgoal_tac "({u} Un c): super S c" 1);
+by (Asm_full_simp_tac 1);
+by (rewrite_tac [super_def,psubset_def]);
+by (blast_tac (claset() addIs [chain_extend] addDs [chain_Union_upper]) 1);
+qed "maxchain_Zorn";
+
+Goal "ALL c: chain S. Union(c): S ==> \
+\     EX y: S. ALL z: S. y <= z --> y = z";
+by (cut_facts_tac [Hausdorff,maxchain_subset_chain] 1);
+by (etac exE 1);
+by (dtac subsetD 1 THEN assume_tac 1);
+by (dtac bspec 1 THEN assume_tac 1);
+by (res_inst_tac [("x","Union(c)")] bexI 1);
+by (rtac ballI 1 THEN rtac impI 1);
+by (blast_tac (claset() addSDs [maxchain_Zorn]) 1);
+by (assume_tac 1);
+qed "Zorn_Lemma";
+
+(*-------------------------------------------------------------
+             Alternative version of Zorn's Lemma
+ --------------------------------------------------------------*)
+Goal "ALL (c:: 'a set set): chain S. EX y : S. ALL x : c. x <= y ==> \
+\     EX y : S. ALL x : S. (y :: 'a set) <= x --> y = x";
+by (cut_facts_tac [Hausdorff,maxchain_subset_chain] 1);
+by (EVERY1[etac exE, dtac subsetD, assume_tac]);
+by (EVERY1[dtac bspec, assume_tac, etac bexE]);
+by (res_inst_tac [("x","y")] bexI 1);
+by (assume_tac 2);
+by (EVERY1[rtac ballI, rtac impI, rtac ccontr]);
+by (forw_inst_tac [("z","x")]  chain_extend 1);
+by (assume_tac 1 THEN Blast_tac 1);
+by (rewrite_tac [maxchain_def,super_def,psubset_def]);
+by (blast_tac (claset() addSEs [equalityCE]) 1);
+qed "Zorn_Lemma2";
+
+(** misc. lemmas **)
+
+Goalw [chain_def] "[| c : chain S; x: c; y: c |] ==> x <= y | y <= x";
+by (Blast_tac 1);
+qed "chainD";
+
+Goalw [chain_def] "!!(c :: 'a set set). c: chain S ==> c <= S";
+by (Blast_tac 1);
+qed "chainD2";
+
+(* proved elsewhere? *) 
+Goal "x : Union(c) ==> EX m:c. x:m";
+by (Blast_tac 1);
+qed "mem_UnionD";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/Zorn.thy	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,34 @@
+(*  Title       : Zorn.thy
+    ID          : $Id$
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Zorn's Lemma -- See lcp's Zorn.thy in ZF
+*) 
+
+Zorn = Main +
+
+constdefs
+  chain     ::  'a::ord set => 'a set set
+    "chain S  == {F. F <= S & (ALL x: F. ALL y: F. x <= y | y <= x)}" 
+
+  super     ::  ['a::ord set,'a set] => 'a set set
+    "super S c == {d. d: chain(S) & c < d}"
+
+  maxchain  ::  'a::ord set => 'a set set
+    "maxchain S == {c. c: chain S & super S c = {}}"
+
+  succ      ::  ['a::ord set,'a set] => 'a set
+    "succ S c == if (c ~: chain S| c: maxchain S) 
+                 then c else (@c'. c': (super S c))" 
+
+consts 
+  "TFin" :: 'a::ord set => 'a set set
+
+inductive "TFin(S)"
+  intrs
+    succI        "x : TFin S ==> succ S x : TFin S"
+    Pow_UnionI   "Y : Pow(TFin S) ==> Union(Y) : TFin S"
+           
+  monos          Pow_mono
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/fuf.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,78 @@
+(*  Title       : HOL/Real/Hyperreal/fuf.ML
+    ID          : $Id$
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+                  1999  University of Edinburgh
+
+Simple tactics to help proofs involving our free ultrafilter
+(FreeUltrafilterNat). We rely on the fact that filters satisfy the
+finite intersection property.
+*)
+
+local
+
+exception FUFempty;
+
+fun get_fuf_hyps [] zs = zs
+|   get_fuf_hyps (x::xs) zs =
+        case (concl_of x) of
+        (_ $ (Const ("Not",_) $ (Const ("op :",_) $ _ $
+             Const ("HyperDef.FreeUltrafilterNat",_)))) =>  get_fuf_hyps xs
+                                              ((x RS FreeUltrafilterNat_Compl_mem)::zs)
+       |(_ $ (Const ("op :",_) $ _ $
+             Const ("HyperDef.FreeUltrafilterNat",_)))  =>  get_fuf_hyps xs (x::zs)
+       | _ => get_fuf_hyps xs zs;
+
+fun inter_prems [] = raise FUFempty
+|   inter_prems [x] = x
+|   inter_prems (x::y::ys) =
+      inter_prems (([x,y] MRS FreeUltrafilterNat_Int) :: ys);
+
+in
+
+(*---------------------------------------------------------------
+   solves goals of the form
+    [| A1: FUF; A2: FUF; ...; An: FUF |] ==> B : FUF
+   where A1 Int A2 Int ... Int An <= B
+ ---------------------------------------------------------------*)
+
+fun fuf_tac css i = METAHYPS(fn prems =>
+                    (rtac ((inter_prems (get_fuf_hyps prems [])) RS
+                           FreeUltrafilterNat_subset) 1) THEN
+                    auto_tac css) i;
+
+fun Fuf_tac i = fuf_tac (clasimpset ()) i;
+
+
+(*---------------------------------------------------------------
+   solves goals of the form
+    [| A1: FUF; A2: FUF; ...; An: FUF |] ==> P
+   where A1 Int A2 Int ... Int An <= {} since {} ~: FUF
+   (i.e. uses fact that FUF is a proper filter)
+ ---------------------------------------------------------------*)
+
+fun fuf_empty_tac css i = METAHYPS (fn prems =>
+  rtac ((inter_prems (get_fuf_hyps prems [])) RS
+    (FreeUltrafilterNat_subset RS (FreeUltrafilterNat_empty RS notE))) 1
+                     THEN auto_tac css) i;
+
+fun Fuf_empty_tac i = fuf_empty_tac (clasimpset ()) i;
+
+
+(*---------------------------------------------------------------
+  All in one -- not really needed.
+ ---------------------------------------------------------------*)
+
+fun fuf_auto_tac css i = SOLVE (fuf_empty_tac css i) ORELSE TRY (fuf_tac css i);
+fun Fuf_auto_tac i = fuf_auto_tac (clasimpset ()) i;
+
+
+(*---------------------------------------------------------------
+   In fact could make this the only tactic: just need to
+   use contraposition and then look for empty set.
+ ---------------------------------------------------------------*)
+
+fun ultra_tac css i = rtac ccontr i THEN fuf_empty_tac css i;
+fun Ultra_tac i = ultra_tac (clasimpset ()) i;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/hypreal_arith.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,39 @@
+(*  Title:      HOL/Real/hypreal_arith.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow, TU Muenchen
+    Copyright   1999 TU Muenchen
+
+Augmentation of hypreal linear arithmetic with rational coefficient handling
+*)
+
+local
+
+(* reduce contradictory <= to False *)
+val simps = [True_implies_equals,inst "w" "number_of ?v" hypreal_add_mult_distrib2,
+             hypreal_divide_1,hypreal_times_divide1_eq,hypreal_times_divide2_eq];
+
+val simprocs = [hypreal_cancel_numeral_factors_divide];
+
+fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
+
+val hypreal_mult_mono_thms =
+ [(rotate_prems 1 hypreal_mult_less_mono2,
+   cvar(hypreal_mult_less_mono2, hd(prems_of hypreal_mult_less_mono2))),
+  (hypreal_mult_le_mono2,
+   cvar(hypreal_mult_le_mono2, hd(tl(prems_of hypreal_mult_le_mono2))))]
+
+in
+
+val hypreal_arith_setup =
+ [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
+   {add_mono_thms = add_mono_thms,
+    mult_mono_thms = mult_mono_thms @ hypreal_mult_mono_thms,
+    inj_thms = inj_thms,
+    lessD = lessD,
+    simpset = simpset addsimps simps addsimprocs simprocs})];
+
+end;
+
+(*
+Procedure "assoc_fold" needed?
+*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/hypreal_arith0.ML	Sat Dec 30 22:03:47 2000 +0100
@@ -0,0 +1,123 @@
+(*  Title:      HOL/Hyperreal/hypreal_arith.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow, TU Muenchen
+    Copyright   1999 TU Muenchen
+
+Instantiation of the generic linear arithmetic package for type hypreal.
+*)
+
+val hypreal_mult_le_mono2 = rotate_prems 1 hypreal_mult_le_le_mono1;
+
+local
+
+(* reduce contradictory <= to False *)
+val simps = 
+    [order_less_irrefl, zero_eq_numeral_0, one_eq_numeral_1,
+     add_hypreal_number_of, minus_hypreal_number_of, diff_hypreal_number_of,
+     mult_hypreal_number_of, eq_hypreal_number_of, less_hypreal_number_of,
+     le_hypreal_number_of_eq_not_less, hypreal_diff_def,
+     hypreal_minus_add_distrib, hypreal_minus_minus, hypreal_mult_assoc];
+
+val add_rules =
+    map rename_numerals
+        [hypreal_add_zero_left, hypreal_add_zero_right,
+         hypreal_add_minus, hypreal_add_minus_left,
+         hypreal_mult_0, hypreal_mult_0_right,
+         hypreal_mult_1, hypreal_mult_1_right,
+         hypreal_mult_minus_1, hypreal_mult_minus_1_right];
+
+val simprocs = [Hyperreal_Times_Assoc.conv, 
+                Hyperreal_Numeral_Simprocs.combine_numerals]@
+               Hyperreal_Numeral_Simprocs.cancel_numerals;
+
+val mono_ss = simpset() addsimps
+                [hypreal_add_le_mono,hypreal_add_less_mono,
+                 hypreal_add_less_le_mono,hypreal_add_le_less_mono];
+
+val add_mono_thms_hypreal =
+  map (fn s => prove_goal (the_context ()) s
+                 (fn prems => [cut_facts_tac prems 1, asm_simp_tac mono_ss 1]))
+    ["(i <= j) & (k <= l) ==> i + k <= j + (l::hypreal)",
+     "(i  = j) & (k <= l) ==> i + k <= j + (l::hypreal)",
+     "(i <= j) & (k  = l) ==> i + k <= j + (l::hypreal)",
+     "(i  = j) & (k  = l) ==> i + k  = j + (l::hypreal)",
+     "(i < j) & (k = l)   ==> i + k < j + (l::hypreal)",
+     "(i = j) & (k < l)   ==> i + k < j + (l::hypreal)",
+     "(i < j) & (k <= l)  ==> i + k < j + (l::hypreal)",
+     "(i <= j) & (k < l)  ==> i + k < j + (l::hypreal)",
+     "(i < j) & (k < l)   ==> i + k < j + (l::hypreal)"];
+
+val hypreal_arith_simproc_pats =
+  map (fn s => Thm.read_cterm (Theory.sign_of (the_context ()))
+       (s, HOLogic.boolT))
+      ["(m::hypreal) < n", "(m::hypreal) <= n", "(m::hypreal) = n"];
+
+fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
+
+val hypreal_mult_mono_thms =
+ [(rotate_prems 1 hypreal_mult_less_mono2,
+   cvar(hypreal_mult_less_mono2, hd(prems_of hypreal_mult_less_mono2))),
+  (hypreal_mult_le_mono2,
+   cvar(hypreal_mult_le_mono2, hd(tl(prems_of hypreal_mult_le_mono2))))]
+
+in
+
+val fast_hypreal_arith_simproc = mk_simproc
+  "fast_hypreal_arith" hypreal_arith_simproc_pats Fast_Arith.lin_arith_prover;
+
+val hypreal_arith_setup =
+ [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
+   {add_mono_thms = add_mono_thms @ add_mono_thms_hypreal,
+    mult_mono_thms = mult_mono_thms @ hypreal_mult_mono_thms,
+    inj_thms = inj_thms, (*FIXME: add hypreal*)
+    lessD = lessD,  (*We don't change LA_Data_Ref.lessD because the hypreal ordering is dense!*)
+    simpset = simpset addsimps (add_rules @ simps)
+                      addsimprocs simprocs}),
+  arith_discrete ("HyperDef.hypreal",false),
+  Simplifier.change_simpset_of (op addsimprocs) [fast_hypreal_arith_simproc]];
+
+end;
+
+
+(* Some test data [omitting examples that assume the ordering to be discrete!]
+Goal "!!a::hypreal. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d";
+by (fast_arith_tac 1);
+qed "";
+
+Goal "!!a::hypreal. [| a <= b; b+b <= c |] ==> a+a <= c";
+by (fast_arith_tac 1);
+qed "";
+
+Goal "!!a::hypreal. [| a+b <= i+j; a<=b; i<=j |] ==> a+a <= j+j";
+by (fast_arith_tac 1);
+qed "";
+
+Goal "!!a::hypreal. a+b+c <= i+j+k & a<=b & b<=c & i<=j & j<=k --> a+a+a <= k+k+k";
+by (arith_tac 1);
+qed "";
+
+Goal "!!a::hypreal. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
+\     ==> a <= l";
+by (fast_arith_tac 1);
+qed "";
+
+Goal "!!a::hypreal. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
+\     ==> a+a+a+a <= l+l+l+l";
+by (fast_arith_tac 1);
+qed "";
+
+Goal "!!a::hypreal. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
+\     ==> a+a+a+a+a <= l+l+l+l+i";
+by (fast_arith_tac 1);
+qed "";
+
+Goal "!!a::hypreal. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
+\     ==> a+a+a+a+a+a <= l+l+l+l+i+l";
+by (fast_arith_tac 1);
+qed "";
+
+Goal "!!a::hypreal. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
+\     ==> #6*a <= #5*l+i";
+by (fast_arith_tac 1);
+qed "";
+*)
--- a/src/HOL/IsaMakefile	Sat Dec 30 22:03:46 2000 +0100
+++ b/src/HOL/IsaMakefile	Sat Dec 30 22:03:47 2000 +0100
@@ -7,7 +7,7 @@
 ## targets
 
 default: HOL
-images: HOL HOL-Real TLA
+images: HOL HOL-Real HOL-Hyperreal TLA
 
 #Note: keep targets sorted (except for HOL-Library)
 test: \
@@ -108,21 +108,9 @@
 
 HOL-Real: HOL $(OUT)/HOL-Real
 
-$(OUT)/HOL-Real: $(OUT)/HOL Real/Hyperreal/Filter.ML \
-  Real/Hyperreal/Filter.thy Real/Hyperreal/HRealAbs.ML \
-  Real/Hyperreal/HRealAbs.thy Real/Hyperreal/HyperDef.ML \
-  Real/Hyperreal/HyperDef.thy Real/Hyperreal/HyperNat.ML \
-  Real/Hyperreal/HyperNat.thy Real/Hyperreal/HyperOrd.ML \
-  Real/Hyperreal/HyperOrd.thy Real/Hyperreal/HyperPow.ML \
-  Real/Hyperreal/HyperPow.thy Real/Hyperreal/Hyperreal.thy \
-  Real/Hyperreal/Lim.ML Real/Hyperreal/Lim.thy Real/Hyperreal/NSA.ML \
-  Real/Hyperreal/NSA.thy Real/Hyperreal/NatStar.ML \
-  Real/Hyperreal/NatStar.thy Real/Hyperreal/SEQ.ML \
-  Real/Hyperreal/SEQ.thy Real/Hyperreal/Series.ML \
-  Real/Hyperreal/Series.thy Real/Hyperreal/Star.ML \
-  Real/Hyperreal/Star.thy Real/Hyperreal/Zorn.ML \
-  Real/Hyperreal/Zorn.thy Real/Hyperreal/fuf.ML Real/Lubs.ML \
-  Real/Lubs.thy Real/PNat.ML Real/PNat.thy Real/PRat.ML Real/PRat.thy \
+$(OUT)/HOL-Real: $(OUT)/HOL \
+  Real/Lubs.ML Real/Lubs.thy Real/PNat.ML Real/PNat.thy \
+  Real/PRat.ML Real/PRat.thy \
   Real/PReal.ML Real/PReal.thy Real/RComplete.ML Real/RComplete.thy \
   Real/ROOT.ML Real/Real.thy Real/RealAbs.ML Real/RealAbs.thy \
   Real/RealArith0.ML Real/RealArith0.thy Real/real_arith0.ML \
@@ -132,6 +120,29 @@
   Real/RealPow.thy Real/real_arith.ML
 	@cd Real; $(ISATOOL) usedir -b $(OUT)/HOL HOL-Real
 
+## HOL-Hyperreal
+
+HOL-Hyperreal: HOL-Real $(OUT)/HOL-Hyperreal
+
+$(OUT)/HOL-Hyperreal: $(OUT)/HOL-Real \
+  Hyperreal/Filter.ML Hyperreal/Filter.thy \
+  Hyperreal/HRealAbs.ML Hyperreal/HRealAbs.thy \
+  Hyperreal/HyperArith0.ML Hyperreal/HyperArith0.thy Hyperreal/HyperArith.thy \
+  Hyperreal/hypreal_arith0.ML Hyperreal/hypreal_arith.ML \
+  Hyperreal/HyperBin.ML Hyperreal/HyperBin.thy \
+  Hyperreal/HyperDef.ML Hyperreal/HyperDef.thy Hyperreal/HyperNat.ML \
+  Hyperreal/HyperNat.thy Hyperreal/HyperOrd.ML \
+  Hyperreal/HyperOrd.thy Hyperreal/HyperPow.ML \
+  Hyperreal/HyperPow.thy Hyperreal/Hyperreal.thy \
+  Hyperreal/Lim.ML Hyperreal/Lim.thy Hyperreal/NSA.ML \
+  Hyperreal/NSA.thy Hyperreal/NatStar.ML \
+  Hyperreal/NatStar.thy Hyperreal/SEQ.ML \
+  Hyperreal/SEQ.thy Hyperreal/Series.ML \
+  Hyperreal/Series.thy Hyperreal/Star.ML \
+  Hyperreal/Star.thy Hyperreal/Zorn.ML \
+  Hyperreal/Zorn.thy Hyperreal/fuf.ML
+	@cd Hyperreal; $(ISATOOL) usedir -b $(OUT)/HOL-Real HOL-Hyperreal
+
 
 ## HOL-Real-ex
 
--- a/src/HOL/Real/Hyperreal/Filter.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,554 +0,0 @@
-(*  Title       : Filter.ML
-    ID          : $Id$
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Filters and Ultrafilter
-*) 
-
-(*------------------------------------------------------------------
-      Properties of Filters and Freefilters - 
-      rules for intro, destruction etc.
- ------------------------------------------------------------------*)
-
-Goalw [is_Filter_def] "is_Filter X S ==> X <= Pow(S)";
-by (Blast_tac 1);
-qed "is_FilterD1";
-
-Goalw [is_Filter_def] "is_Filter X S ==> X ~= {}";
-by (Blast_tac 1);
-qed "is_FilterD2";
-
-Goalw [is_Filter_def] "is_Filter X S ==> {} ~: X";
-by (Blast_tac 1);
-qed "is_FilterD3";
-
-Goalw [Filter_def] "is_Filter X S ==> X : Filter S";
-by (Blast_tac 1);
-qed "mem_FiltersetI";
-
-Goalw [Filter_def] "X : Filter S ==> is_Filter X S";
-by (Blast_tac 1);
-qed "mem_FiltersetD";
-
-Goal "X : Filter S ==> {} ~: X";
-by (etac (mem_FiltersetD RS is_FilterD3) 1);
-qed "Filter_empty_not_mem";
-
-bind_thm ("Filter_empty_not_memE",(Filter_empty_not_mem RS notE));
-
-Goalw [Filter_def,is_Filter_def] 
-      "[| X: Filter S; A: X; B: X |] ==> A Int B : X";
-by (Blast_tac 1);
-qed "mem_FiltersetD1";
-
-Goalw [Filter_def,is_Filter_def] 
-      "[| X: Filter S; A: X; A <= B; B <= S|] ==> B : X";
-by (Blast_tac 1);
-qed "mem_FiltersetD2";
-
-Goalw [Filter_def,is_Filter_def] 
-      "[| X: Filter S; A: X |] ==> A : Pow S";
-by (Blast_tac 1);
-qed "mem_FiltersetD3";
-
-Goalw [Filter_def,is_Filter_def] 
-      "X: Filter S  ==> S : X";
-by (Blast_tac 1);
-qed "mem_FiltersetD4";
-
-Goalw [is_Filter_def] 
-      "[| X <= Pow(S);\
-\              S : X; \
-\              X ~= {}; \
-\              {} ~: X; \
-\              ALL u: X. ALL v: X. u Int v : X; \
-\              ALL u v. u: X & u<=v & v<=S --> v: X \
-\           |] ==> is_Filter X S";
-by (Blast_tac 1); 
-qed "is_FilterI";
-
-Goal "[| X <= Pow(S);\
-\              S : X; \
-\              X ~= {}; \
-\              {} ~: X; \
-\              ALL u: X. ALL v: X. u Int v : X; \
-\              ALL u v. u: X & u<=v & v<=S --> v: X \
-\           |] ==> X: Filter S";
-by (blast_tac (claset() addIs [mem_FiltersetI,is_FilterI]) 1);
-qed "mem_FiltersetI2";
-
-Goalw [is_Filter_def]
-      "is_Filter X S ==> X <= Pow(S) & \
-\                          S : X & \
-\                          X ~= {} & \
-\                          {} ~: X  & \
-\                          (ALL u: X. ALL v: X. u Int v : X) & \
-\                          (ALL u v. u: X & u <= v & v<=S --> v: X)";
-by (Fast_tac 1);
-qed "is_FilterE_lemma";
-
-Goalw [is_Filter_def]
-      "X : Filter S ==> X <= Pow(S) &\
-\                          S : X & \
-\                          X ~= {} & \
-\                          {} ~: X  & \
-\                          (ALL u: X. ALL v: X. u Int v : X) & \
-\                          (ALL u v. u: X & u <= v & v<=S --> v: X)";
-by (etac (mem_FiltersetD RS is_FilterE_lemma) 1);
-qed "memFiltersetE_lemma";
-
-Goalw [Filter_def,Freefilter_def] 
-      "X: Freefilter S ==> X: Filter S";
-by (Fast_tac 1);
-qed "Freefilter_Filter";
-
-Goalw [Freefilter_def] 
-      "X: Freefilter S ==> ALL y: X. ~finite(y)";
-by (Blast_tac 1);
-qed "mem_Freefilter_not_finite";
-
-Goal "[| X: Freefilter S; x: X |] ==> ~ finite x";
-by (blast_tac (claset() addSDs [mem_Freefilter_not_finite]) 1);
-qed "mem_FreefiltersetD1";
-
-bind_thm ("mem_FreefiltersetE1", (mem_FreefiltersetD1 RS notE));
-
-Goal "[| X: Freefilter S; finite x|] ==> x ~: X";
-by (blast_tac (claset() addSDs [mem_Freefilter_not_finite]) 1);
-qed "mem_FreefiltersetD2";
-
-Goalw [Freefilter_def] 
-      "[| X: Filter S; ALL x. ~(x: X & finite x) |] ==> X: Freefilter S";
-by (Blast_tac 1);
-qed "mem_FreefiltersetI1";
-
-Goalw [Freefilter_def]
-      "[| X: Filter S; ALL x. (x ~: X | ~ finite x) |] ==> X: Freefilter S";
-by (Blast_tac 1);
-qed "mem_FreefiltersetI2";
-
-Goal "[| X: Filter S; A: X; B: X |] ==> A Int B ~= {}";
-by (forw_inst_tac [("A","A"),("B","B")] mem_FiltersetD1 1);
-by (auto_tac (claset() addSDs [Filter_empty_not_mem],simpset()));
-qed "Filter_Int_not_empty";
-
-bind_thm ("Filter_Int_not_emptyE",(Filter_Int_not_empty RS notE));
-
-(*----------------------------------------------------------------------------------
-              Ultrafilters and Free ultrafilters
- ----------------------------------------------------------------------------------*)
-
-Goalw [Ultrafilter_def] "X : Ultrafilter S ==> X: Filter S";
-by (Blast_tac 1);
-qed "Ultrafilter_Filter";
-
-Goalw [Ultrafilter_def] 
-      "X : Ultrafilter S ==> !A: Pow(S). A : X | S - A: X";
-by (Blast_tac 1);
-qed "mem_UltrafiltersetD2";
-
-Goalw [Ultrafilter_def] 
-      "[|X : Ultrafilter S; A <= S; A ~: X |] ==> S - A: X";
-by (Blast_tac 1);
-qed "mem_UltrafiltersetD3";
-
-Goalw [Ultrafilter_def] 
-      "[|X : Ultrafilter S; A <= S; S - A ~: X |] ==> A: X";
-by (Blast_tac 1);
-qed "mem_UltrafiltersetD4";
-
-Goalw [Ultrafilter_def]
-     "[| X: Filter S; \
-\             ALL A: Pow(S). A: X | S - A : X |] ==> X: Ultrafilter S";
-by (Blast_tac 1);
-qed "mem_UltrafiltersetI";
-
-Goalw [Ultrafilter_def,FreeUltrafilter_def]
-     "X: FreeUltrafilter S ==> X: Ultrafilter S";
-by (Blast_tac 1);
-qed "FreeUltrafilter_Ultrafilter";
-
-Goalw [FreeUltrafilter_def]
-     "X: FreeUltrafilter S ==> ALL y: X. ~finite(y)";
-by (Blast_tac 1);
-qed "mem_FreeUltrafilter_not_finite";
-
-Goal "[| X: FreeUltrafilter S; x: X |] ==> ~ finite x";
-by (blast_tac (claset() addSDs [mem_FreeUltrafilter_not_finite]) 1);
-qed "mem_FreeUltrafiltersetD1";
-
-bind_thm ("mem_FreeUltrafiltersetE1", (mem_FreeUltrafiltersetD1 RS notE));
-
-Goal "[| X: FreeUltrafilter S; finite x|] ==> x ~: X";
-by (blast_tac (claset() addSDs [mem_FreeUltrafilter_not_finite]) 1);
-qed "mem_FreeUltrafiltersetD2";
-
-Goalw [FreeUltrafilter_def] 
-      "[| X: Ultrafilter S; \
-\              ALL x. ~(x: X & finite x) |] ==> X: FreeUltrafilter S";
-by (Blast_tac 1);
-qed "mem_FreeUltrafiltersetI1";
-
-Goalw [FreeUltrafilter_def]
-      "[| X: Ultrafilter S; \
-\              ALL x. (x ~: X | ~ finite x) |] ==> X: FreeUltrafilter S";
-by (Blast_tac 1);
-qed "mem_FreeUltrafiltersetI2";
-
-Goalw [FreeUltrafilter_def,Freefilter_def,Ultrafilter_def]
-     "(X: FreeUltrafilter S) = (X: Freefilter S & (ALL x:Pow(S). x: X | S - x: X))";
-by (Blast_tac 1);
-qed "FreeUltrafilter_iff";
-
-(*-------------------------------------------------------------------
-   A Filter F on S is an ultrafilter iff it is a maximal filter 
-   i.e. whenever G is a filter on I and F <= F then F = G
- --------------------------------------------------------------------*)
-(*---------------------------------------------------------------------
-  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 
- ---------------------------------------------------------------------*)
-Goal "[| F ~= {}; A <= S |] ==> \
-\        EX x. x: {X. X <= S & (EX f:F. A Int f <= X)}";
-by (Blast_tac 1);
-qed "lemma_set_extend";
-
-Goal "a: X ==> X ~= {}";
-by (Step_tac 1);
-qed "lemma_set_not_empty";
-
-Goal "x Int F <= {} ==> F <= - x";
-by (Blast_tac 1);
-qed "lemma_empty_Int_subset_Compl";
-
-Goalw [Filter_def,is_Filter_def]
-      "[| F: Filter S; A ~: F; A <= S|] \
-\          ==> ALL B. B ~: F | ~ B <= A";
-by (Blast_tac 1);
-qed "mem_Filterset_disjI";
-
-Goal "F : Ultrafilter S ==> \
-\         (F: Filter S & (ALL G: Filter S. F <= G --> F = G))";
-by (auto_tac (claset(),simpset() addsimps [Ultrafilter_def]));
-by (dres_inst_tac [("x","x")] bspec 1);
-by (etac mem_FiltersetD3 1 THEN assume_tac 1);
-by (Step_tac 1);
-by (dtac subsetD 1 THEN assume_tac 1);
-by (blast_tac (claset() addSDs [Filter_Int_not_empty]) 1);
-qed "Ultrafilter_max_Filter";
-
-
-(*--------------------------------------------------------------------------------
-     This is a very long and tedious proof; need to break it into parts.
-     Have proof that {X. X <= S & (EX f: F. A Int f <= X)} is a filter as 
-     a lemma
---------------------------------------------------------------------------------*)
-Goalw [Ultrafilter_def] 
-      "[| F: Filter S; \
-\              ALL G: Filter S. F <= G --> F = G |] ==> F : Ultrafilter S";
-by (Step_tac 1);
-by (rtac ccontr 1);
-by (forward_tac [mem_FiltersetD RS is_FilterD2] 1);
-by (forw_inst_tac [("x","{X. X <= S & (EX f: F. A Int f <= X)}")] bspec 1);
-by (EVERY1[rtac mem_FiltersetI2, Blast_tac, Asm_full_simp_tac]);
-by (blast_tac (claset() addDs [mem_FiltersetD3]) 1);
-by (etac (lemma_set_extend RS exE) 1);
-by (assume_tac 1 THEN etac lemma_set_not_empty 1);
-by (REPEAT(rtac ballI 2) THEN Asm_full_simp_tac 2);
-by (rtac conjI 2 THEN Blast_tac 2);
-by (REPEAT(etac conjE 2) THEN REPEAT(etac bexE 2));
-by (res_inst_tac [("x","f Int fa")] bexI 2);
-by (etac mem_FiltersetD1 3);
-by (assume_tac 3 THEN assume_tac 3);
-by (Fast_tac 2);
-by (EVERY[REPEAT(rtac allI 2), rtac impI 2,Asm_full_simp_tac 2]);
-by (EVERY[REPEAT(etac conjE 2), etac bexE 2]);
-by (res_inst_tac [("x","f")] bexI 2);
-by (rtac subsetI 2);
-by (Fast_tac 2 THEN assume_tac 2);
-by (Step_tac 2);
-by (Blast_tac 3);
-by (eres_inst_tac [("c","A")] equalityCE 3);
-by (REPEAT(Blast_tac 3));
-by (dres_inst_tac [("A","xa")] mem_FiltersetD3 2 THEN assume_tac 2);
-by (Blast_tac 2);
-by (dtac lemma_empty_Int_subset_Compl 1);
-by (EVERY1[ftac mem_Filterset_disjI , assume_tac, Fast_tac]);
-by (dtac mem_FiltersetD3 1 THEN assume_tac 1);
-by (dres_inst_tac [("x","f")] spec 1);
-by (Blast_tac 1);
-qed "max_Filter_Ultrafilter";
-
-Goal "(F : Ultrafilter S) = (F: Filter S & (ALL G: Filter S. F <= G --> F = G))";
-by (blast_tac (claset() addSIs [Ultrafilter_max_Filter,max_Filter_Ultrafilter]) 1);
-qed "Ultrafilter_iff";
-
-(*--------------------------------------------------------------------
-             A few properties of freefilters
- -------------------------------------------------------------------*)
-
-Goal "F1 Int F2 = ((F1 Int Y) Int F2) Un ((F2 Int (- Y)) Int F1)";
-by (Auto_tac);
-qed "lemma_Compl_cancel_eq";
-
-Goal "finite X ==> finite (X Int Y)";
-by (etac (Int_lower1 RS finite_subset) 1);
-qed "finite_IntI1";
-
-Goal "finite Y ==> finite (X Int Y)";
-by (etac (Int_lower2 RS finite_subset) 1);
-qed "finite_IntI2";
-
-Goal "[| finite (F1 Int Y); \
-\                 finite (F2 Int (- Y)) \
-\              |] ==> finite (F1 Int F2)";
-by (res_inst_tac [("Y1","Y")] (lemma_Compl_cancel_eq RS ssubst) 1);
-by (rtac finite_UnI 1);
-by (auto_tac (claset() addSIs [finite_IntI1,finite_IntI2],simpset()));
-qed "finite_Int_Compl_cancel";
-
-Goal "U: Freefilter S  ==> \
-\         ~ (EX f1: U. EX f2: U. finite (f1 Int x) \
-\                            & finite (f2 Int (- x)))";
-by (Step_tac 1);
-by (forw_inst_tac [("A","f1"),("B","f2")] 
-    (Freefilter_Filter RS mem_FiltersetD1) 1);
-by (dres_inst_tac [("x","f1 Int f2")] mem_FreefiltersetD1 3);
-by (dtac finite_Int_Compl_cancel 4);
-by (Auto_tac);
-qed "Freefilter_lemma_not_finite";
-
-(* the lemmas below follow *)
-Goal "U: Freefilter S ==> \
-\          ALL f: U. ~ finite (f Int x) | ~finite (f Int (- x))";
-by (blast_tac (claset() addSDs [Freefilter_lemma_not_finite,bspec]) 1);
-qed "Freefilter_Compl_not_finite_disjI";
-
-Goal "U: Freefilter S ==> \
-\          (ALL f: U. ~ finite (f Int x)) | (ALL f:U. ~finite (f Int (- x)))";
-by (blast_tac (claset() addSDs [Freefilter_lemma_not_finite,bspec]) 1);
-qed "Freefilter_Compl_not_finite_disjI2";
-
-Goal "- UNIV = {}";
-by (Auto_tac );
-qed "Compl_UNIV_eq";
-
-Addsimps [Compl_UNIV_eq];
-
-Goal "- {} = UNIV";
-by (Auto_tac );
-qed "Compl_empty_eq";
-
-Addsimps [Compl_empty_eq];
-
-val [prem] = goal (the_context ()) "~ finite (UNIV:: 'a set) ==> \
-\            {A:: 'a set. finite (- A)} : Filter UNIV";
-by (cut_facts_tac [prem] 1);
-by (rtac mem_FiltersetI2 1);
-by (auto_tac (claset(), simpset() delsimps [Collect_empty_eq]));
-by (eres_inst_tac [("c","UNIV")] equalityCE 1);
-by (Auto_tac);
-by (etac (Compl_anti_mono RS finite_subset) 1);
-by (assume_tac 1);
-qed "cofinite_Filter";
-
-Goal "~finite(UNIV :: 'a set) ==> ~finite (X :: 'a set) | ~finite (- X)";
-by (dres_inst_tac [("A1","X")] (Compl_partition RS ssubst) 1);
-by (Asm_full_simp_tac 1); 
-qed "not_finite_UNIV_disjI";
-
-Goal "[| ~finite(UNIV :: 'a set); \
-\                 finite (X :: 'a set) \
-\              |] ==>  ~finite (- X)";
-by (dres_inst_tac [("X","X")] not_finite_UNIV_disjI 1);
-by (Blast_tac 1);
-qed "not_finite_UNIV_Compl";
-
-val [prem] = goal (the_context ()) "~ finite (UNIV:: 'a set) ==> \
-\            !X: {A:: 'a set. finite (- A)}. ~ finite X";
-by (cut_facts_tac [prem] 1);
-by (auto_tac (claset() addDs [not_finite_UNIV_disjI],simpset()));
-qed "mem_cofinite_Filter_not_finite";
-
-val [prem] = goal (the_context ()) "~ finite (UNIV:: 'a set) ==> \
-\            {A:: 'a set. finite (- A)} : Freefilter UNIV";
-by (cut_facts_tac [prem] 1);
-by (rtac mem_FreefiltersetI2 1);
-by (rtac cofinite_Filter 1 THEN assume_tac 1);
-by (blast_tac (claset() addSDs [mem_cofinite_Filter_not_finite]) 1);
-qed "cofinite_Freefilter";
-
-Goal "UNIV - x = - x";
-by (Auto_tac);
-qed "UNIV_diff_Compl";
-Addsimps [UNIV_diff_Compl];
-
-Goalw [Ultrafilter_def,FreeUltrafilter_def]
-     "[| ~finite(UNIV :: 'a set); (U :: 'a set set): FreeUltrafilter UNIV\
-\         |] ==> {X. finite(- X)} <= U";
-by (ftac cofinite_Filter 1);
-by (Step_tac 1);
-by (forw_inst_tac [("X","- x :: 'a set")] not_finite_UNIV_Compl 1);
-by (assume_tac 1);
-by (Step_tac 1 THEN Fast_tac 1);
-by (dres_inst_tac [("x","x")] bspec 1);
-by (Blast_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [UNIV_diff_Compl]) 1);
-qed "FreeUltrafilter_contains_cofinite_set";
-
-(*--------------------------------------------------------------------
-   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
- -------------------------------------------------------------------*)
-Open_locale "UFT"; 
-
-Goalw [chain_def, thm "superfrechet_def", thm "frechet_def"]
-   "!!(c :: 'a set set set). c : chain (superfrechet S) ==>  Union c <= Pow S";
-by (Step_tac 1);
-by (dtac subsetD 1 THEN assume_tac 1);
-by (Step_tac 1);
-by (dres_inst_tac [("X","X")] mem_FiltersetD3 1);
-by (Auto_tac);
-qed "chain_Un_subset_Pow";
-
-Goalw [chain_def,Filter_def,is_Filter_def,
-           thm "superfrechet_def", thm "frechet_def"] 
-          "!!(c :: 'a set set set). c: chain (superfrechet S) \
-\         ==> !x: c. {} < x";
-by (blast_tac (claset() addSIs [psubsetI]) 1);
-qed "mem_chain_psubset_empty";
-
-Goal "!!(c :: 'a set set set). \
-\            [| c: chain (superfrechet S);\
-\               c ~= {} \
-\            |]\
-\            ==> Union(c) ~= {}";
-by (dtac mem_chain_psubset_empty 1);
-by (Step_tac 1);
-by (dtac bspec 1 THEN assume_tac 1);
-by (auto_tac (claset() addDs [Union_upper,bspec],
-    simpset() addsimps [psubset_def]));
-qed "chain_Un_not_empty";
-
-Goalw [is_Filter_def,Filter_def,chain_def,thm "superfrechet_def"] 
-           "!!(c :: 'a set set set). \
-\           c : chain (superfrechet S)  \
-\           ==> {} ~: Union(c)";
-by (Blast_tac 1);
-qed "Filter_empty_not_mem_Un";
-
-Goal "c: chain (superfrechet S) \
-\         ==> ALL u : Union(c). ALL v: Union(c). u Int v : Union(c)";
-by (Step_tac 1);
-by (forw_inst_tac [("x","X"),("y","Xa")] chainD 1);
-by (REPEAT(assume_tac 1));
-by (dtac chainD2 1);
-by (etac disjE 1);
-by (res_inst_tac [("X","Xa")] UnionI 1 THEN assume_tac 1);
-by (dres_inst_tac [("A","X")] subsetD 1 THEN assume_tac 1);
-by (dres_inst_tac [("c","Xa")] subsetD 1 THEN assume_tac 1);
-by (res_inst_tac [("X","X")] UnionI 2 THEN assume_tac 2);
-by (dres_inst_tac [("A","Xa")] subsetD 2 THEN assume_tac 2);
-by (dres_inst_tac [("c","X")] subsetD 2 THEN assume_tac 2);
-by (auto_tac (claset() addIs [mem_FiltersetD1], 
-     simpset() addsimps [thm "superfrechet_def"]));
-qed "Filter_Un_Int";
-
-Goal "c: chain (superfrechet S) \
-\         ==> ALL u v. u: Union(c) & \
-\                 (u :: 'a set) <= v & v <= S --> v: Union(c)";
-by (Step_tac 1);
-by (dtac chainD2 1);
-by (dtac subsetD 1 THEN assume_tac 1);
-by (rtac UnionI 1 THEN assume_tac 1);
-by (auto_tac (claset() addIs [mem_FiltersetD2], 
-     simpset() addsimps [thm "superfrechet_def"]));
-qed "Filter_Un_subset";
-
-Goalw [chain_def,thm "superfrechet_def"]
-      "!!(c :: 'a set set set). \
-\            [| c: chain (superfrechet S);\
-\               x: c \
-\            |] ==> x : Filter S";
-by (Blast_tac 1);
-qed "lemma_mem_chain_Filter";
-
-Goalw [chain_def,thm "superfrechet_def"]
-     "!!(c :: 'a set set set). \
-\            [| c: chain (superfrechet S);\
-\               x: c \
-\            |] ==> frechet S <= x";
-by (Blast_tac 1);
-qed "lemma_mem_chain_frechet_subset";
-
-Goal "!!(c :: 'a set set set). \
-\         [| c ~= {}; \
-\            c : chain (superfrechet (UNIV :: 'a set))\
-\         |] ==> Union c : superfrechet (UNIV)";
-by (simp_tac (simpset() addsimps 
-    [thm "superfrechet_def",thm "frechet_def"]) 1);
-by (Step_tac 1);
-by (rtac mem_FiltersetI2 1);
-by (etac chain_Un_subset_Pow 1);
-by (rtac UnionI 1 THEN assume_tac 1);
-by (etac (lemma_mem_chain_Filter RS mem_FiltersetD4) 1 THEN assume_tac 1);
-by (etac chain_Un_not_empty 1);
-by (etac Filter_empty_not_mem_Un 2);
-by (etac Filter_Un_Int 2);
-by (etac Filter_Un_subset 2);
-by (subgoal_tac "xa : frechet (UNIV)" 2);
-by (rtac UnionI 2 THEN assume_tac 2);
-by (rtac (lemma_mem_chain_frechet_subset RS subsetD) 2);
-by (auto_tac (claset(),simpset() addsimps [thm "frechet_def"]));
-qed "Un_chain_mem_cofinite_Filter_set";
-
-Goal "EX U: superfrechet (UNIV). \
-\               ALL G: superfrechet (UNIV). U <= G --> U = G";
-by (rtac Zorn_Lemma2 1);
-by (cut_facts_tac [thm "not_finite_UNIV" RS cofinite_Filter] 1);
-by (Step_tac 1);
-by (res_inst_tac [("Q","c={}")] (excluded_middle RS disjE) 1);
-by (res_inst_tac [("x","Union c")] bexI 1 THEN Blast_tac 1);
-by (rtac Un_chain_mem_cofinite_Filter_set 1 THEN REPEAT(assume_tac 1));
-by (res_inst_tac [("x","frechet (UNIV)")] bexI 1 THEN Blast_tac 1);
-by (auto_tac (claset(),
-	      simpset() addsimps 
-	      [thm "superfrechet_def", thm "frechet_def"]));
-qed "max_cofinite_Filter_Ex";
-
-Goal "EX U: superfrechet UNIV. (\
-\               ALL G: superfrechet UNIV. U <= G --> U = G) \ 
-\                             & (ALL x: U. ~finite x)";
-by (cut_facts_tac [thm "not_finite_UNIV" RS 
-         (export max_cofinite_Filter_Ex)] 1);
-by (Step_tac 1);
-by (res_inst_tac [("x","U")] bexI 1);
-by (auto_tac (claset(),simpset() addsimps 
-        [thm "superfrechet_def", thm "frechet_def"]));
-by (dres_inst_tac [("c","- x")] subsetD 1);
-by (Asm_simp_tac 1);
-by (forw_inst_tac [("A","x"),("B","- x")] mem_FiltersetD1 1);
-by (dtac Filter_empty_not_mem 3);
-by (ALLGOALS(Asm_full_simp_tac ));
-qed "max_cofinite_Freefilter_Ex";
-
-(*--------------------------------------------------------------------------------
-               There exists a free ultrafilter on any infinite set
- --------------------------------------------------------------------------------*)
-
-Goalw [FreeUltrafilter_def] "EX U. U: FreeUltrafilter (UNIV :: 'a set)";
-by (cut_facts_tac [thm "not_finite_UNIV" RS (export max_cofinite_Freefilter_Ex)] 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [thm "superfrechet_def", Ultrafilter_iff, thm "frechet_def"]) 1);
-by (Step_tac 1);
-by (res_inst_tac [("x","U")] exI 1);
-by (Step_tac 1);
-by (Blast_tac 1);
-qed "FreeUltrafilter_ex";
-
-bind_thm ("FreeUltrafilter_Ex", export FreeUltrafilter_ex);
-
-Close_locale "UFT";
--- a/src/HOL/Real/Hyperreal/Filter.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,45 +0,0 @@
-(*  Title       : Filter.thy
-    ID          : $Id$
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Filters and Ultrafilters
-*) 
-
-Filter = Zorn + 
-
-constdefs
-
-  is_Filter       :: ['a set set,'a set] => bool
-  "is_Filter F S == (F <= Pow(S) & S : F & {} ~: F &
-                   (ALL u: F. ALL v: F. u Int v : F) &
-                   (ALL u v. u: F & u <= v & v <= S --> v: F))" 
-
-  Filter          :: 'a set => 'a set set set
-  "Filter S == {X. is_Filter X S}"
- 
-  (* free filter does not contain any finite set *)
-
-  Freefilter      :: 'a set => 'a set set set
-  "Freefilter S == {X. X: Filter S & (ALL x: X. ~ finite x)}"
-
-  Ultrafilter     :: 'a set => 'a set set set
-  "Ultrafilter S == {X. X: Filter S & (ALL A: Pow(S). A: X | S - A : X)}"
-
-  FreeUltrafilter :: 'a set => 'a set set set
-  "FreeUltrafilter S == {X. X: Ultrafilter S & (ALL x: X. ~ finite x)}" 
-
-  (* A locale makes proof of Ultrafilter Theorem more modular *)
-locale UFT = 
-       fixes     frechet :: "'a set => 'a set set"
-                 superfrechet :: "'a set => 'a set set set"
-
-       assumes   not_finite_UNIV "~finite (UNIV :: 'a set)"
-
-       defines   frechet_def "frechet S == {A. finite (S - A)}"
-                 superfrechet_def "superfrechet S == 
-                                   {G.  G: Filter S & frechet S <= G}"
-end
-
-
-
-
--- a/src/HOL/Real/Hyperreal/HRealAbs.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,226 +0,0 @@
-(*  Title       : HRealAbs.ML
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Absolute value function for the hyperreals
-                  Similar to RealAbs.thy
-*) 
-
-(*------------------------------------------------------------
-  absolute value on hyperreals as pointwise operation on 
-  equivalence class representative
- ------------------------------------------------------------*)
-
-Goalw [hrabs_def]
-     "abs (Abs_hypreal (hyprel ^^ {X})) = \
-\     Abs_hypreal(hyprel ^^ {%n. abs (X n)})";
-by (auto_tac (claset(),
-           simpset() addsimps [hypreal_zero_def, hypreal_le,hypreal_minus]));
-by (ALLGOALS(Ultra_tac THEN' arith_tac ));
-qed "hypreal_hrabs";
-
-(*------------------------------------------------------------
-   Properties of the absolute value function over the reals
-   (adapted version of previously proved theorems about abs)
- ------------------------------------------------------------*)
-
-Goal "abs (0::hypreal) = 0";
-by (simp_tac (simpset() addsimps [hrabs_def]) 1); 
-qed "hrabs_zero";
-Addsimps [hrabs_zero];
-
-Goal "(0::hypreal)<=x ==> abs x = x";
-by (asm_simp_tac (simpset() addsimps [hrabs_def]) 1); 
-qed "hrabs_eqI1";
-
-Goal "(0::hypreal)<x ==> abs x = x";
-by (asm_simp_tac (simpset() addsimps [hypreal_less_imp_le, hrabs_eqI1]) 1);
-qed "hrabs_eqI2";
-
-Goal "x<(0::hypreal) ==> abs x = -x";
-by (asm_simp_tac (simpset() addsimps [hypreal_le_def, hrabs_def]) 1); 
-qed "hrabs_minus_eqI2";
-
-Goal "x<=(0::hypreal) ==> abs x = -x";
-by (auto_tac (claset() addDs [order_antisym], simpset() addsimps [hrabs_def])); 
-qed "hrabs_minus_eqI1";
-
-Goal "(0::hypreal)<= abs x";
-by (auto_tac (claset() addDs [hypreal_minus_zero_less_iff RS iffD2, 
-                              hypreal_less_asym], 
-              simpset() addsimps [hypreal_le_def, hrabs_def]));
-qed "hrabs_ge_zero";
-
-Goal "abs(abs x) = abs (x::hypreal)";
-by (auto_tac (claset() addDs [hypreal_minus_zero_less_iff RS iffD2, 
-                              hypreal_less_asym], 
-              simpset() addsimps [hypreal_le_def, hrabs_def]));
-qed "hrabs_idempotent";
-Addsimps [hrabs_idempotent];
-
-Goalw [hrabs_def] "(abs x = (0::hypreal)) = (x=0)";
-by (Simp_tac 1);
-qed "hrabs_zero_iff";
-AddIffs [hrabs_zero_iff];
-
-Goalw [hrabs_def] "(x::hypreal) <= abs x";
-by (auto_tac (claset() addDs [not_hypreal_leE RS hypreal_less_imp_le],
-              simpset() addsimps [hypreal_le_zero_iff]));
-qed "hrabs_ge_self";
-
-Goalw [hrabs_def] "-(x::hypreal) <= abs x";
-by (full_simp_tac (simpset() addsimps [hypreal_ge_zero_iff]) 1);
-qed "hrabs_ge_minus_self";
-
-(* very short proof by "transfer" *)
-Goal "abs(x*(y::hypreal)) = (abs x)*(abs y)";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(), 
-              simpset() addsimps [hypreal_hrabs, hypreal_mult,abs_mult]));
-qed "hrabs_mult";
-
-Goal "abs(inverse(x)) = inverse(abs(x::hypreal))";
-by (hypreal_div_undefined_case_tac "x=0" 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-       simpset() addsimps [hypreal_hrabs,
-                           hypreal_inverse,hypreal_zero_def]));
-by (ultra_tac (claset(), simpset() addsimps [abs_inverse]) 1);
-qed "hrabs_inverse";
-
-Goal "abs(x+(y::hypreal)) <= abs x + abs y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(), 
-      simpset() addsimps [hypreal_hrabs, hypreal_add,hypreal_le,
-                        abs_triangle_ineq]));
-qed "hrabs_triangle_ineq";
-
-Goal "abs((w::hypreal) + x + y) <= abs(w) + abs(x) + abs(y)";
-by (auto_tac (claset() addSIs [hrabs_triangle_ineq RS hypreal_le_trans,
-                               hypreal_add_left_le_mono1],
-    simpset() addsimps [hypreal_add_assoc]));
-qed "hrabs_triangle_ineq_three";
-
-Goalw [hrabs_def] "abs(-x)=abs((x::hypreal))";
-by (auto_tac (claset() addSDs [not_hypreal_leE, hypreal_less_asym] 
-                       addIs [hypreal_le_anti_sym],
-              simpset() addsimps [hypreal_ge_zero_iff]));
-qed "hrabs_minus_cancel";
-Addsimps [hrabs_minus_cancel];
-
-val prem1::prem2::rest = goal thy 
-    "[| abs x < r; abs y < s |] ==> abs(x+y) < r + (s::hypreal)";
-by (rtac hypreal_le_less_trans 1);
-by (rtac hrabs_triangle_ineq 1);
-by (rtac ([prem1,prem2] MRS hypreal_add_less_mono) 1);
-qed "hrabs_add_less";
-
-val prem1::prem2::rest = 
-    goal thy "[| abs x<r; abs y<s |] ==> abs(x*y)<r*(s::hypreal)";
-by (simp_tac (simpset() addsimps [hrabs_mult]) 1);
-by (rtac hypreal_mult_le_less_trans 1);
-by (rtac hrabs_ge_zero 1);
-by (rtac prem2 1);
-by (rtac hypreal_mult_less_mono1 1);
-by (rtac (prem2 RS (hrabs_ge_zero RS hypreal_le_less_trans)) 1);
-by (rtac prem1 1);
-by (rtac ([prem1 RS (hrabs_ge_zero RS hypreal_le_less_trans),
-           prem2 RS (hrabs_ge_zero RS hypreal_le_less_trans)] 
-          MRS hypreal_mult_order) 1);
-qed "hrabs_mult_less";
-
-Goal "1hr < abs x ==> abs y <= abs(x*y)";
-by (cut_inst_tac [("x1","y")] (hrabs_ge_zero RS hypreal_le_imp_less_or_eq) 1);
-by (EVERY1[etac disjE,rtac hypreal_less_imp_le]);
-by (dres_inst_tac [("x1","1hr")]  (hypreal_less_minus_iff RS iffD1) 1);
-by (forw_inst_tac [("y","abs x +-1hr")] hypreal_mult_order 1);
-by (assume_tac 1);
-by (rtac (hypreal_less_minus_iff RS iffD2) 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_mult_distrib2,
-    hrabs_mult, hypreal_mult_commute,hypreal_minus_mult_eq2 RS sym]) 1);
-by (dtac sym 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_le_refl,hrabs_mult]) 1);
-qed "hrabs_mult_le";
-
-Goal "[| 1hr < abs x; r < abs y|] ==> r < abs(x*y)";
-by (fast_tac (HOL_cs addIs [hrabs_mult_le, hypreal_less_le_trans]) 1);
-qed "hrabs_mult_gt";
-
-Goal "abs x < r ==> (0::hypreal) < r";
-by (blast_tac (claset() addSIs [hypreal_le_less_trans,
-    hrabs_ge_zero]) 1);
-qed "hrabs_less_gt_zero";
-
-Goalw [hrabs_def] "abs 1hr = 1hr";
-by (auto_tac (claset() addSDs [not_hypreal_leE RS hypreal_less_asym], 
-      simpset() addsimps [hypreal_zero_less_one]));
-qed "hrabs_one";
-
-Goal "abs x = (x::hypreal) | abs x = -x";
-by (cut_inst_tac [("x","0"),("y","x")] hypreal_linear 1);
-by (fast_tac (claset() addIs [hrabs_eqI2,hrabs_minus_eqI2,
-                            hrabs_zero]) 1);
-qed "hrabs_disj";
-
-Goal "abs x = (y::hypreal) ==> x = y | -x = y";
-by (dtac sym 1);
-by (hyp_subst_tac 1);
-by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1);
-by (REPEAT(Asm_simp_tac 1));
-qed "hrabs_eq_disj";
-
-Goal "(abs x < (r::hypreal)) = (-r < x & x < r)";
-by (Step_tac 1);
-by (rtac (hypreal_less_swap_iff RS iffD2) 1);
-by (asm_simp_tac (simpset() addsimps [(hrabs_ge_minus_self 
-    RS hypreal_le_less_trans)]) 1);
-by (asm_simp_tac (simpset() addsimps [(hrabs_ge_self 
-    RS hypreal_le_less_trans)]) 1);
-by (EVERY1 [dtac (hypreal_less_swap_iff RS iffD1), rotate_tac 1, 
-            dtac (hypreal_minus_minus RS subst), 
-            cut_inst_tac [("x","x")] hrabs_disj, dtac disjE ]);
-by (assume_tac 3 THEN Auto_tac);
-qed "hrabs_interval_iff";
-
-Goal "(abs x < (r::hypreal)) = (- x < r & x < r)";
-by (auto_tac (claset(),  simpset() addsimps [hrabs_interval_iff]));
-by (dtac (hypreal_less_swap_iff RS iffD1) 1);
-by (dtac (hypreal_less_swap_iff RS iffD1) 2);
-by (Auto_tac);
-qed "hrabs_interval_iff2";
-
-(* Needed in Geom.ML *)
-Goal "(y::hypreal) + - x + (y + - z) = abs (x + - z) ==> y = z | x = y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (auto_tac (claset(), 
-              simpset() addsimps [hypreal_hrabs, hypreal_minus,hypreal_add]));
-by (Ultra_tac 1 THEN arith_tac 1);
-qed "hrabs_add_lemma_disj";
-
-(***SHOULD BE TRIVIAL GIVEN SIMPROCS
-Goal "abs((x::hypreal) + -y) = abs (y + -x)";
-by (rtac (hrabs_minus_cancel RS subst) 1);
-by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
-qed "hrabs_minus_add_cancel";
-
-(* Needed in Geom.ML *)
-Goal "(x::hypreal) + - y + (z + - y) = abs (x + - z) \
-\     ==> y = z | x = y";
-by (rtac (hypreal_minus_eq_cancel RS subst) 1);
-by (res_inst_tac [("b1","y")] (hypreal_minus_eq_cancel RS subst) 1);
-by (rtac hrabs_add_lemma_disj 1);
-by (asm_full_simp_tac (simpset() addsimps [hrabs_minus_add_cancel] 
-         @ hypreal_add_ac) 1);
-qed "hrabs_add_lemma_disj2";
-***)
- 
-(*----------------------------------------------------------
-    Relating hrabs to abs through embedding of IR into IR*
- ----------------------------------------------------------*)
-Goalw [hypreal_of_real_def] 
-    "abs (hypreal_of_real r) = hypreal_of_real (abs r)";
-by (auto_tac (claset(), simpset() addsimps [hypreal_hrabs]));
-qed "hypreal_of_real_hrabs";
--- a/src/HOL/Real/Hyperreal/HRealAbs.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,12 +0,0 @@
-(*  Title       : HRealAbs.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Absolute value function for the hyperreals
-*) 
-
-HRealAbs = HyperOrd + RealAbs + 
-
-defs
-    hrabs_def "abs r  == if (0::hypreal) <=r then r else -r" 
-
-end
\ No newline at end of file
--- a/src/HOL/Real/Hyperreal/HyperDef.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1365 +0,0 @@
-(*  Title       : HOL/Real/Hyperreal/Hyper.ML
-    ID          : $Id$
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Ultrapower construction of hyperreals
-*) 
-
-(*------------------------------------------------------------------------
-             Proof that the set of naturals is not finite
- ------------------------------------------------------------------------*)
-
-(*** based on James' proof that the set of naturals is not finite ***)
-Goal "finite (A::nat set) --> (EX n. ALL m. Suc (n + m) ~: A)";
-by (rtac impI 1);
-by (eres_inst_tac [("F","A")] finite_induct 1);
-by (Blast_tac 1 THEN etac exE 1);
-by (res_inst_tac [("x","n + x")] exI 1);
-by (rtac allI 1 THEN eres_inst_tac [("x","x + m")] allE 1);
-by (auto_tac (claset(), simpset() addsimps add_ac));
-by (auto_tac (claset(),
-	      simpset() addsimps [add_assoc RS sym,
-				  less_add_Suc2 RS less_not_refl2]));
-qed_spec_mp "finite_exhausts";
-
-Goal "finite (A :: nat set) --> (EX n. n ~:A)";
-by (rtac impI 1 THEN dtac finite_exhausts 1);
-by (Blast_tac 1);
-qed_spec_mp "finite_not_covers";
-
-Goal "~ finite(UNIV:: nat set)";
-by (fast_tac (claset() addSDs [finite_exhausts]) 1);
-qed "not_finite_nat";
-
-(*------------------------------------------------------------------------
-   Existence of free ultrafilter over the naturals and proof of various 
-   properties of the FreeUltrafilterNat- an arbitrary free ultrafilter
- ------------------------------------------------------------------------*)
-
-Goal "EX U. U: FreeUltrafilter (UNIV::nat set)";
-by (rtac (not_finite_nat RS FreeUltrafilter_Ex) 1);
-qed "FreeUltrafilterNat_Ex";
-
-Goalw [FreeUltrafilterNat_def] 
-     "FreeUltrafilterNat: FreeUltrafilter(UNIV:: nat set)";
-by (rtac (FreeUltrafilterNat_Ex RS exE) 1);
-by (rtac someI2 1 THEN ALLGOALS(assume_tac));
-qed "FreeUltrafilterNat_mem";
-Addsimps [FreeUltrafilterNat_mem];
-
-Goalw [FreeUltrafilterNat_def] "finite x ==> x ~: FreeUltrafilterNat";
-by (rtac (FreeUltrafilterNat_Ex RS exE) 1);
-by (rtac someI2 1 THEN assume_tac 1);
-by (blast_tac (claset() addDs [mem_FreeUltrafiltersetD1]) 1);
-qed "FreeUltrafilterNat_finite";
-
-Goal "x: FreeUltrafilterNat ==> ~ finite x";
-by (blast_tac (claset() addDs [FreeUltrafilterNat_finite]) 1);
-qed "FreeUltrafilterNat_not_finite";
-
-Goalw [FreeUltrafilterNat_def] "{} ~: FreeUltrafilterNat";
-by (rtac (FreeUltrafilterNat_Ex RS exE) 1);
-by (rtac someI2 1 THEN assume_tac 1);
-by (blast_tac (claset() addDs [FreeUltrafilter_Ultrafilter,
-			       Ultrafilter_Filter,Filter_empty_not_mem]) 1);
-qed "FreeUltrafilterNat_empty";
-Addsimps [FreeUltrafilterNat_empty];
-
-Goal "[| X: FreeUltrafilterNat;  Y: FreeUltrafilterNat |]  \
-\     ==> X Int Y : FreeUltrafilterNat";
-by (cut_facts_tac [FreeUltrafilterNat_mem] 1);
-by (blast_tac (claset() addDs [FreeUltrafilter_Ultrafilter,
-			       Ultrafilter_Filter,mem_FiltersetD1]) 1);
-qed "FreeUltrafilterNat_Int";
-
-Goal "[| X: FreeUltrafilterNat;  X <= Y |] \
-\     ==> Y : FreeUltrafilterNat";
-by (cut_facts_tac [FreeUltrafilterNat_mem] 1);
-by (blast_tac (claset() addDs [FreeUltrafilter_Ultrafilter,
-			       Ultrafilter_Filter,mem_FiltersetD2]) 1);
-qed "FreeUltrafilterNat_subset";
-
-Goal "X: FreeUltrafilterNat ==> -X ~: FreeUltrafilterNat";
-by (Step_tac 1);
-by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
-by Auto_tac;
-qed "FreeUltrafilterNat_Compl";
-
-Goal "X~: FreeUltrafilterNat ==> -X : FreeUltrafilterNat";
-by (cut_facts_tac [FreeUltrafilterNat_mem RS (FreeUltrafilter_iff RS iffD1)] 1);
-by (Step_tac 1 THEN dres_inst_tac [("x","X")] bspec 1);
-by (auto_tac (claset(), simpset() addsimps [UNIV_diff_Compl]));
-qed "FreeUltrafilterNat_Compl_mem";
-
-Goal "(X ~: FreeUltrafilterNat) = (-X: FreeUltrafilterNat)";
-by (blast_tac (claset() addDs [FreeUltrafilterNat_Compl,
-			       FreeUltrafilterNat_Compl_mem]) 1);
-qed "FreeUltrafilterNat_Compl_iff1";
-
-Goal "(X: FreeUltrafilterNat) = (-X ~: FreeUltrafilterNat)";
-by (auto_tac (claset(),
-	      simpset() addsimps [FreeUltrafilterNat_Compl_iff1 RS sym]));
-qed "FreeUltrafilterNat_Compl_iff2";
-
-Goal "(UNIV::nat set) : FreeUltrafilterNat";
-by (rtac (FreeUltrafilterNat_mem RS FreeUltrafilter_Ultrafilter RS 
-          Ultrafilter_Filter RS mem_FiltersetD4) 1);
-qed "FreeUltrafilterNat_UNIV";
-Addsimps [FreeUltrafilterNat_UNIV];
-
-Goal "UNIV : FreeUltrafilterNat";
-by Auto_tac;
-qed "FreeUltrafilterNat_Nat_set";
-Addsimps [FreeUltrafilterNat_Nat_set];
-
-Goal "{n. P(n) = P(n)} : FreeUltrafilterNat";
-by (Simp_tac 1);
-qed "FreeUltrafilterNat_Nat_set_refl";
-AddIs [FreeUltrafilterNat_Nat_set_refl];
-
-Goal "{n::nat. P} : FreeUltrafilterNat ==> P";
-by (rtac ccontr 1);
-by (rotate_tac 1 1);
-by (Asm_full_simp_tac 1);
-qed "FreeUltrafilterNat_P";
-
-Goal "{n. P(n)} : FreeUltrafilterNat ==> EX n. P(n)";
-by (rtac ccontr 1 THEN rotate_tac 1 1);
-by (Asm_full_simp_tac 1);
-qed "FreeUltrafilterNat_Ex_P";
-
-Goal "ALL n. P(n) ==> {n. P(n)} : FreeUltrafilterNat";
-by (auto_tac (claset() addIs [FreeUltrafilterNat_Nat_set], simpset()));
-qed "FreeUltrafilterNat_all";
-
-(*-------------------------------------------------------
-     Define and use Ultrafilter tactics
- -------------------------------------------------------*)
-use "fuf.ML";
-
-(*-------------------------------------------------------
-  Now prove one further property of our free ultrafilter
- -------------------------------------------------------*)
-Goal "X Un Y: FreeUltrafilterNat \
-\     ==> X: FreeUltrafilterNat | Y: FreeUltrafilterNat";
-by Auto_tac;
-by (Ultra_tac 1);
-qed "FreeUltrafilterNat_Un";
-
-(*-------------------------------------------------------
-   Properties of hyprel
- -------------------------------------------------------*)
-
-(** Proving that hyprel is an equivalence relation **)
-(** Natural deduction for hyprel **)
-
-Goalw [hyprel_def]
-   "((X,Y): hyprel) = ({n. X n = Y n}: FreeUltrafilterNat)";
-by (Fast_tac 1);
-qed "hyprel_iff";
-
-Goalw [hyprel_def] 
-     "{n. X n = Y n}: FreeUltrafilterNat  ==> (X,Y): hyprel";
-by (Fast_tac 1);
-qed "hyprelI";
-
-Goalw [hyprel_def]
-  "p: hyprel --> (EX X Y. \
-\                 p = (X,Y) & {n. X n = Y n} : FreeUltrafilterNat)";
-by (Fast_tac 1);
-qed "hyprelE_lemma";
-
-val [major,minor] = goal (the_context ())
-  "[| p: hyprel;  \
-\     !!X Y. [| p = (X,Y); {n. X n = Y n}: FreeUltrafilterNat\
-\                    |] ==> Q |] ==> Q";
-by (cut_facts_tac [major RS (hyprelE_lemma RS mp)] 1);
-by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
-qed "hyprelE";
-
-AddSIs [hyprelI];
-AddSEs [hyprelE];
-
-Goalw [hyprel_def] "(x,x): hyprel";
-by (auto_tac (claset(),
-              simpset() addsimps [FreeUltrafilterNat_Nat_set]));
-qed "hyprel_refl";
-
-Goal "{n. X n = Y n} = {n. Y n = X n}";
-by Auto_tac;
-qed "lemma_perm";
-
-Goalw [hyprel_def] "(x,y): hyprel --> (y,x):hyprel";
-by (auto_tac (claset() addIs [lemma_perm RS subst], simpset()));
-qed_spec_mp "hyprel_sym";
-
-Goalw [hyprel_def]
-      "(x,y): hyprel --> (y,z):hyprel --> (x,z):hyprel";
-by Auto_tac;
-by (Ultra_tac 1);
-qed_spec_mp "hyprel_trans";
-
-Goalw [equiv_def, refl_def, sym_def, trans_def] "equiv UNIV hyprel";
-by (auto_tac (claset() addSIs [hyprel_refl] 
-                       addSEs [hyprel_sym,hyprel_trans] 
-                       delrules [hyprelI,hyprelE],
-	      simpset() addsimps [FreeUltrafilterNat_Nat_set]));
-qed "equiv_hyprel";
-
-(* (hyprel ^^ {x} = hyprel ^^ {y}) = ((x,y) : hyprel) *)
-bind_thm ("equiv_hyprel_iff",
-    	  [equiv_hyprel, UNIV_I, UNIV_I] MRS eq_equiv_class_iff);
-
-Goalw  [hypreal_def,hyprel_def,quotient_def] "hyprel^^{x}:hypreal";
-by (Blast_tac 1);
-qed "hyprel_in_hypreal";
-
-Goal "inj_on Abs_hypreal hypreal";
-by (rtac inj_on_inverseI 1);
-by (etac Abs_hypreal_inverse 1);
-qed "inj_on_Abs_hypreal";
-
-Addsimps [equiv_hyprel_iff,inj_on_Abs_hypreal RS inj_on_iff,
-          hyprel_iff, hyprel_in_hypreal, Abs_hypreal_inverse];
-
-Addsimps [equiv_hyprel RS eq_equiv_class_iff];
-bind_thm ("eq_hyprelD", equiv_hyprel RSN (2,eq_equiv_class));
-
-Goal "inj(Rep_hypreal)";
-by (rtac inj_inverseI 1);
-by (rtac Rep_hypreal_inverse 1);
-qed "inj_Rep_hypreal";
-
-Goalw [hyprel_def] "x: hyprel ^^ {x}";
-by (Step_tac 1);
-by (auto_tac (claset() addSIs [FreeUltrafilterNat_Nat_set], simpset()));
-qed "lemma_hyprel_refl";
-
-Addsimps [lemma_hyprel_refl];
-
-Goalw [hypreal_def] "{} ~: hypreal";
-by (auto_tac (claset() addSEs [quotientE], simpset()));
-qed "hypreal_empty_not_mem";
-
-Addsimps [hypreal_empty_not_mem];
-
-Goal "Rep_hypreal x ~= {}";
-by (cut_inst_tac [("x","x")] Rep_hypreal 1);
-by Auto_tac;
-qed "Rep_hypreal_nonempty";
-
-Addsimps [Rep_hypreal_nonempty];
-
-(*------------------------------------------------------------------------
-   hypreal_of_real: the injection from real to hypreal
- ------------------------------------------------------------------------*)
-
-Goal "inj(hypreal_of_real)";
-by (rtac injI 1);
-by (rewtac hypreal_of_real_def);
-by (dtac (inj_on_Abs_hypreal RS inj_onD) 1);
-by (REPEAT (rtac hyprel_in_hypreal 1));
-by (dtac eq_equiv_class 1);
-by (rtac equiv_hyprel 1);
-by (Fast_tac 1);
-by (rtac ccontr 1 THEN rotate_tac 1 1);
-by Auto_tac;
-qed "inj_hypreal_of_real";
-
-val [prem] = goal (the_context ())
-    "(!!x y. z = Abs_hypreal(hyprel^^{x}) ==> P) ==> P";
-by (res_inst_tac [("x1","z")] 
-    (rewrite_rule [hypreal_def] Rep_hypreal RS quotientE) 1);
-by (dres_inst_tac [("f","Abs_hypreal")] arg_cong 1);
-by (res_inst_tac [("x","x")] prem 1);
-by (asm_full_simp_tac (simpset() addsimps [Rep_hypreal_inverse]) 1);
-qed "eq_Abs_hypreal";
-
-(**** hypreal_minus: additive inverse on hypreal ****)
-
-Goalw [congruent_def]
-  "congruent hyprel (%X. hyprel^^{%n. - (X n)})";
-by Safe_tac;
-by (ALLGOALS Ultra_tac);
-qed "hypreal_minus_congruent";
-
-Goalw [hypreal_minus_def]
-   "- (Abs_hypreal(hyprel^^{%n. X n})) = Abs_hypreal(hyprel ^^ {%n. -(X n)})";
-by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
-by (simp_tac (simpset() addsimps 
-      [hyprel_in_hypreal RS Abs_hypreal_inverse,
-       [equiv_hyprel, hypreal_minus_congruent] MRS UN_equiv_class]) 1);
-qed "hypreal_minus";
-
-Goal "- (- z) = (z::hypreal)";
-by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (asm_simp_tac (simpset() addsimps [hypreal_minus]) 1);
-qed "hypreal_minus_minus";
-
-Addsimps [hypreal_minus_minus];
-
-Goal "inj(%r::hypreal. -r)";
-by (rtac injI 1);
-by (dres_inst_tac [("f","uminus")] arg_cong 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_minus_minus]) 1);
-qed "inj_hypreal_minus";
-
-Goalw [hypreal_zero_def] "-0 = (0::hypreal)";
-by (simp_tac (simpset() addsimps [hypreal_minus]) 1);
-qed "hypreal_minus_zero";
-Addsimps [hypreal_minus_zero];
-
-Goal "(-x = 0) = (x = (0::hypreal))"; 
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-       simpset() addsimps [hypreal_zero_def, hypreal_minus, eq_commute] @ 
-                          real_add_ac));
-qed "hypreal_minus_zero_iff";
-
-Addsimps [hypreal_minus_zero_iff];
-
-
-(**** hyperreal addition: hypreal_add  ****)
-
-Goalw [congruent2_def]
-    "congruent2 hyprel (%X Y. hyprel^^{%n. X n + Y n})";
-by Safe_tac;
-by (ALLGOALS(Ultra_tac));
-qed "hypreal_add_congruent2";
-
-Goalw [hypreal_add_def]
-  "Abs_hypreal(hyprel^^{%n. X n}) + Abs_hypreal(hyprel^^{%n. Y n}) = \
-\  Abs_hypreal(hyprel^^{%n. X n + Y n})";
-by (simp_tac (simpset() addsimps 
-         [[equiv_hyprel, hypreal_add_congruent2] MRS UN_equiv_class2]) 1);
-qed "hypreal_add";
-
-Goal "(z::hypreal) + w = w + z";
-by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","w")] eq_Abs_hypreal 1);
-by (asm_simp_tac (simpset() addsimps (real_add_ac @ [hypreal_add])) 1);
-qed "hypreal_add_commute";
-
-Goal "((z1::hypreal) + z2) + z3 = z1 + (z2 + z3)";
-by (res_inst_tac [("z","z1")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","z2")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","z3")] eq_Abs_hypreal 1);
-by (asm_simp_tac (simpset() addsimps [hypreal_add, real_add_assoc]) 1);
-qed "hypreal_add_assoc";
-
-(*For AC rewriting*)
-Goal "(x::hypreal)+(y+z)=y+(x+z)";
-by (rtac (hypreal_add_commute RS trans) 1);
-by (rtac (hypreal_add_assoc RS trans) 1);
-by (rtac (hypreal_add_commute RS arg_cong) 1);
-qed "hypreal_add_left_commute";
-
-(* hypreal addition is an AC operator *)
-bind_thms ("hypreal_add_ac", [hypreal_add_assoc,hypreal_add_commute,
-                      hypreal_add_left_commute]);
-
-Goalw [hypreal_zero_def] "(0::hypreal) + z = z";
-by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypreal_add]) 1);
-qed "hypreal_add_zero_left";
-
-Goal "z + (0::hypreal) = z";
-by (simp_tac (simpset() addsimps 
-    [hypreal_add_zero_left,hypreal_add_commute]) 1);
-qed "hypreal_add_zero_right";
-
-Goalw [hypreal_zero_def] "z + -z = (0::hypreal)";
-by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_minus, hypreal_add]) 1);
-qed "hypreal_add_minus";
-
-Goal "-z + z = (0::hypreal)";
-by (simp_tac (simpset() addsimps [hypreal_add_commute, hypreal_add_minus]) 1);
-qed "hypreal_add_minus_left";
-
-Addsimps [hypreal_add_minus,hypreal_add_minus_left,
-          hypreal_add_zero_left,hypreal_add_zero_right];
-
-Goal "EX y. (x::hypreal) + y = 0";
-by (fast_tac (claset() addIs [hypreal_add_minus]) 1);
-qed "hypreal_minus_ex";
-
-Goal "EX! y. (x::hypreal) + y = 0";
-by (auto_tac (claset() addIs [hypreal_add_minus], simpset()));
-by (dres_inst_tac [("f","%x. ya+x")] arg_cong 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
-qed "hypreal_minus_ex1";
-
-Goal "EX! y. y + (x::hypreal) = 0";
-by (auto_tac (claset() addIs [hypreal_add_minus_left], simpset()));
-by (dres_inst_tac [("f","%x. x+ya")] arg_cong 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc]) 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
-qed "hypreal_minus_left_ex1";
-
-Goal "x + y = (0::hypreal) ==> x = -y";
-by (cut_inst_tac [("z","y")] hypreal_add_minus_left 1);
-by (res_inst_tac [("x1","y")] (hypreal_minus_left_ex1 RS ex1E) 1);
-by (Blast_tac 1);
-qed "hypreal_add_minus_eq_minus";
-
-Goal "EX y::hypreal. x = -y";
-by (cut_inst_tac [("x","x")] hypreal_minus_ex 1);
-by (etac exE 1 THEN dtac hypreal_add_minus_eq_minus 1);
-by (Fast_tac 1);
-qed "hypreal_as_add_inverse_ex";
-
-Goal "-(x + (y::hypreal)) = -x + -y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-              simpset() addsimps [hypreal_minus, hypreal_add,
-                                  real_minus_add_distrib]));
-qed "hypreal_minus_add_distrib";
-Addsimps [hypreal_minus_add_distrib];
-
-Goal "-(y + -(x::hypreal)) = x + -y";
-by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
-qed "hypreal_minus_distrib1";
-
-Goal "(x + - (y::hypreal)) + (y + - z) = x + -z";
-by (res_inst_tac [("w1","y")] (hypreal_add_commute RS subst) 1);
-by (simp_tac (simpset() addsimps [hypreal_add_left_commute,
-                                  hypreal_add_assoc]) 1);
-by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
-qed "hypreal_add_minus_cancel1";
-
-Goal "((x::hypreal) + y = x + z) = (y = z)";
-by (Step_tac 1);
-by (dres_inst_tac [("f","%t.-x + t")] arg_cong 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
-qed "hypreal_add_left_cancel";
-
-Goal "z + (x + (y + -z)) = x + (y::hypreal)";
-by (simp_tac (simpset() addsimps hypreal_add_ac) 1);
-qed "hypreal_add_minus_cancel2";
-Addsimps [hypreal_add_minus_cancel2];
-
-Goal "y + -(x + y) = -(x::hypreal)";
-by (Full_simp_tac 1);
-by (rtac (hypreal_add_left_commute RS subst) 1);
-by (Full_simp_tac 1);
-qed "hypreal_add_minus_cancel";
-Addsimps [hypreal_add_minus_cancel];
-
-Goal "y + -(y + x) = -(x::hypreal)";
-by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
-qed "hypreal_add_minus_cancelc";
-Addsimps [hypreal_add_minus_cancelc];
-
-Goal "(z + -x) + (y + -z) = (y + -(x::hypreal))";
-by (full_simp_tac
-    (simpset() addsimps [hypreal_minus_add_distrib RS sym, 
-                         hypreal_add_left_cancel] @ hypreal_add_ac 
-               delsimps [hypreal_minus_add_distrib]) 1); 
-qed "hypreal_add_minus_cancel3";
-Addsimps [hypreal_add_minus_cancel3];
-
-Goal "(y + (x::hypreal)= z + x) = (y = z)";
-by (simp_tac (simpset() addsimps [hypreal_add_commute,
-                                  hypreal_add_left_cancel]) 1);
-qed "hypreal_add_right_cancel";
-
-Goal "z + (y + -z) = (y::hypreal)";
-by (simp_tac (simpset() addsimps hypreal_add_ac) 1);
-qed "hypreal_add_minus_cancel4";
-Addsimps [hypreal_add_minus_cancel4];
-
-Goal "z + (w + (x + (-z + y))) = w + x + (y::hypreal)";
-by (simp_tac (simpset() addsimps hypreal_add_ac) 1);
-qed "hypreal_add_minus_cancel5";
-Addsimps [hypreal_add_minus_cancel5];
-
-Goal "z + ((- z) + w) = (w::hypreal)";
-by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
-qed "hypreal_add_minus_cancelA";
-
-Goal "(-z) + (z + w) = (w::hypreal)";
-by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
-qed "hypreal_minus_add_cancelA";
-
-Addsimps [hypreal_add_minus_cancelA, hypreal_minus_add_cancelA];
-
-(**** hyperreal multiplication: hypreal_mult  ****)
-
-Goalw [congruent2_def]
-    "congruent2 hyprel (%X Y. hyprel^^{%n. X n * Y n})";
-by Safe_tac;
-by (ALLGOALS(Ultra_tac));
-qed "hypreal_mult_congruent2";
-
-Goalw [hypreal_mult_def]
-  "Abs_hypreal(hyprel^^{%n. X n}) * Abs_hypreal(hyprel^^{%n. Y n}) = \
-\  Abs_hypreal(hyprel^^{%n. X n * Y n})";
-by (simp_tac (simpset() addsimps 
-      [[equiv_hyprel, hypreal_mult_congruent2] MRS UN_equiv_class2]) 1);
-qed "hypreal_mult";
-
-Goal "(z::hypreal) * w = w * z";
-by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","w")] eq_Abs_hypreal 1);
-by (asm_simp_tac (simpset() addsimps ([hypreal_mult] @ real_mult_ac)) 1);
-qed "hypreal_mult_commute";
-
-Goal "((z1::hypreal) * z2) * z3 = z1 * (z2 * z3)";
-by (res_inst_tac [("z","z1")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","z2")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","z3")] eq_Abs_hypreal 1);
-by (asm_simp_tac (simpset() addsimps [hypreal_mult,real_mult_assoc]) 1);
-qed "hypreal_mult_assoc";
-
-qed_goal "hypreal_mult_left_commute" (the_context ())
-    "(z1::hypreal) * (z2 * z3) = z2 * (z1 * z3)"
- (fn _ => [rtac (hypreal_mult_commute RS trans) 1, 
-           rtac (hypreal_mult_assoc RS trans) 1,
-           rtac (hypreal_mult_commute RS arg_cong) 1]);
-
-(* hypreal multiplication is an AC operator *)
-bind_thms ("hypreal_mult_ac", [hypreal_mult_assoc, hypreal_mult_commute, 
-                       hypreal_mult_left_commute]);
-
-Goalw [hypreal_one_def] "1hr * z = z";
-by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_mult]) 1);
-qed "hypreal_mult_1";
-
-Goal "z * 1hr = z";
-by (simp_tac (simpset() addsimps [hypreal_mult_commute,
-    hypreal_mult_1]) 1);
-qed "hypreal_mult_1_right";
-
-Goalw [hypreal_zero_def] "0 * z = (0::hypreal)";
-by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_mult,real_mult_0]) 1);
-qed "hypreal_mult_0";
-
-Goal "z * 0 = (0::hypreal)";
-by (simp_tac (simpset() addsimps [hypreal_mult_commute, hypreal_mult_0]) 1);
-qed "hypreal_mult_0_right";
-
-Addsimps [hypreal_mult_0,hypreal_mult_0_right];
-
-Goal "-(x * y) = -x * (y::hypreal)";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-	      simpset() addsimps [hypreal_minus, hypreal_mult] 
-                                 @ real_mult_ac @ real_add_ac));
-qed "hypreal_minus_mult_eq1";
-
-Goal "-(x * y) = (x::hypreal) * -y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(), simpset() addsimps [hypreal_minus, hypreal_mult] 
-                                           @ real_mult_ac @ real_add_ac));
-qed "hypreal_minus_mult_eq2";
-
-(*Pull negations out*)
-Addsimps [hypreal_minus_mult_eq2 RS sym, hypreal_minus_mult_eq1 RS sym];
-
-Goal "-x*y = (x::hypreal)*-y";
-by Auto_tac;
-qed "hypreal_minus_mult_commute";
-
-(*-----------------------------------------------------------------------------
-    A few more theorems
- ----------------------------------------------------------------------------*)
-Goal "(z::hypreal) + v = z' + v' ==> z + (v + w) = z' + (v' + w)";
-by (asm_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
-qed "hypreal_add_assoc_cong";
-
-Goal "(z::hypreal) + (v + w) = v + (z + w)";
-by (REPEAT (ares_tac [hypreal_add_commute RS hypreal_add_assoc_cong] 1));
-qed "hypreal_add_assoc_swap";
-
-Goal "((z1::hypreal) + z2) * w = (z1 * w) + (z2 * w)";
-by (res_inst_tac [("z","z1")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","z2")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","w")] eq_Abs_hypreal 1);
-by (asm_simp_tac (simpset() addsimps [hypreal_mult,hypreal_add,
-     real_add_mult_distrib]) 1);
-qed "hypreal_add_mult_distrib";
-
-val hypreal_mult_commute'= read_instantiate [("z","w")] hypreal_mult_commute;
-
-Goal "(w::hypreal) * (z1 + z2) = (w * z1) + (w * z2)";
-by (simp_tac (simpset() addsimps [hypreal_mult_commute',hypreal_add_mult_distrib]) 1);
-qed "hypreal_add_mult_distrib2";
-
-bind_thms ("hypreal_mult_simps", [hypreal_mult_1, hypreal_mult_1_right]);
-Addsimps hypreal_mult_simps;
-
-(* 07/00 *)
-
-Goalw [hypreal_diff_def] "((z1::hypreal) - z2) * w = (z1 * w) - (z2 * w)";
-by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib]) 1);
-qed "hypreal_diff_mult_distrib";
-
-Goal "(w::hypreal) * (z1 - z2) = (w * z1) - (w * z2)";
-by (simp_tac (simpset() addsimps [hypreal_mult_commute', 
-				  hypreal_diff_mult_distrib]) 1);
-qed "hypreal_diff_mult_distrib2";
-
-(*** one and zero are distinct ***)
-Goalw [hypreal_zero_def,hypreal_one_def] "0 ~= 1hr";
-by (auto_tac (claset(), simpset() addsimps [real_zero_not_eq_one]));
-qed "hypreal_zero_not_eq_one";
-
-
-(**** multiplicative inverse on hypreal ****)
-
-Goalw [congruent_def]
-  "congruent hyprel (%X. hyprel^^{%n. if X n = #0 then #0 else inverse(X n)})";
-by (Auto_tac THEN Ultra_tac 1);
-qed "hypreal_inverse_congruent";
-
-Goalw [hypreal_inverse_def]
-      "inverse (Abs_hypreal(hyprel^^{%n. X n})) = \
-\      Abs_hypreal(hyprel ^^ {%n. if X n = #0 then #0 else inverse(X n)})";
-by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
-by (simp_tac (simpset() addsimps 
-   [hyprel_in_hypreal RS Abs_hypreal_inverse,
-    [equiv_hyprel, hypreal_inverse_congruent] MRS UN_equiv_class]) 1);
-qed "hypreal_inverse";
-
-Goal "inverse 0 = (0::hypreal)";
-by (simp_tac (simpset() addsimps [hypreal_inverse, hypreal_zero_def]) 1);
-qed "HYPREAL_INVERSE_ZERO";
-
-Goal "a / (0::hypreal) = 0";
-by (simp_tac
-    (simpset() addsimps [hypreal_divide_def, HYPREAL_INVERSE_ZERO]) 1);
-qed "HYPREAL_DIVISION_BY_ZERO";  (*NOT for adding to default simpset*)
-
-fun hypreal_div_undefined_case_tac s i =
-  case_tac s i THEN 
-  asm_simp_tac 
-       (simpset() addsimps [HYPREAL_DIVISION_BY_ZERO, HYPREAL_INVERSE_ZERO]) i;
-
-Goal "inverse (inverse (z::hypreal)) = z";
-by (hypreal_div_undefined_case_tac "z=0" 1);
-by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (asm_full_simp_tac (simpset() addsimps 
-                       [hypreal_inverse, hypreal_zero_def]) 1);
-qed "hypreal_inverse_inverse";
-Addsimps [hypreal_inverse_inverse];
-
-Goalw [hypreal_one_def] "inverse(1hr) = 1hr";
-by (full_simp_tac (simpset() addsimps [hypreal_inverse,
-                                       real_zero_not_eq_one RS not_sym]) 1);
-qed "hypreal_inverse_1";
-Addsimps [hypreal_inverse_1];
-
-
-(*** existence of inverse ***)
-
-Goalw [hypreal_one_def,hypreal_zero_def] 
-     "x ~= 0 ==> x*inverse(x) = 1hr";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (rotate_tac 1 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_inverse, hypreal_mult]) 1);
-by (dtac FreeUltrafilterNat_Compl_mem 1);
-by (blast_tac (claset() addSIs [real_mult_inv_right,
-    FreeUltrafilterNat_subset]) 1);
-qed "hypreal_mult_inverse";
-
-Goal "x ~= 0 ==> inverse(x)*x = 1hr";
-by (asm_simp_tac (simpset() addsimps [hypreal_mult_inverse,
-				      hypreal_mult_commute]) 1);
-qed "hypreal_mult_inverse_left";
-
-Goal "(c::hypreal) ~= 0 ==> (c*a=c*b) = (a=b)";
-by Auto_tac;
-by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_mult_inverse] @ hypreal_mult_ac)  1);
-qed "hypreal_mult_left_cancel";
-    
-Goal "(c::hypreal) ~= 0 ==> (a*c=b*c) = (a=b)";
-by (Step_tac 1);
-by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_mult_inverse] @ hypreal_mult_ac)  1);
-qed "hypreal_mult_right_cancel";
-
-Goalw [hypreal_zero_def] "x ~= 0 ==> inverse (x::hypreal) ~= 0";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_inverse, hypreal_mult]) 1);
-qed "hypreal_inverse_not_zero";
-
-Addsimps [hypreal_mult_inverse,hypreal_mult_inverse_left];
-
-Goal "[| x ~= 0; y ~= 0 |] ==> x * y ~= (0::hypreal)";
-by (Step_tac 1);
-by (dres_inst_tac [("f","%z. inverse x*z")] arg_cong 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym]) 1);
-qed "hypreal_mult_not_0";
-
-bind_thm ("hypreal_mult_not_0E",hypreal_mult_not_0 RS notE);
-
-Goal "x*y = (0::hypreal) ==> x = 0 | y = 0";
-by (auto_tac (claset() addIs [ccontr] addDs [hypreal_mult_not_0],
-              simpset()));
-qed "hypreal_mult_zero_disj";
-
-Goal "inverse(-x) = -inverse(x::hypreal)";
-by (hypreal_div_undefined_case_tac "x=0" 1);
-by (rtac (hypreal_mult_right_cancel RS iffD1) 1);
-by (stac hypreal_mult_inverse_left 2);
-by Auto_tac;
-qed "hypreal_minus_inverse";
-
-Goal "inverse(x*y) = inverse(x)*inverse(y::hypreal)";
-by (hypreal_div_undefined_case_tac "x=0" 1);
-by (hypreal_div_undefined_case_tac "y=0" 1);
-by (forw_inst_tac [("y","y")] hypreal_mult_not_0 1 THEN assume_tac 1);
-by (res_inst_tac [("c1","x")] (hypreal_mult_left_cancel RS iffD1) 1);
-by (auto_tac (claset(), simpset() addsimps [hypreal_mult_assoc RS sym]));
-by (res_inst_tac [("c1","y")] (hypreal_mult_left_cancel RS iffD1) 1);
-by (auto_tac (claset(), simpset() addsimps [hypreal_mult_left_commute]));
-by (asm_simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym]) 1);
-qed "hypreal_inverse_distrib";
-
-(*------------------------------------------------------------------
-                   Theorems for ordering 
- ------------------------------------------------------------------*)
-
-(* prove introduction and elimination rules for hypreal_less *)
-
-Goalw [hypreal_less_def]
- "P < (Q::hypreal) = (EX X Y. X : Rep_hypreal(P) & \
-\                             Y : Rep_hypreal(Q) & \
-\                             {n. X n < Y n} : FreeUltrafilterNat)";
-by (Fast_tac 1);
-qed "hypreal_less_iff";
-
-Goalw [hypreal_less_def]
- "[| {n. X n < Y n} : FreeUltrafilterNat; \
-\         X : Rep_hypreal(P); \
-\         Y : Rep_hypreal(Q) |] ==> P < (Q::hypreal)";
-by (Fast_tac 1);
-qed "hypreal_lessI";
-
-
-Goalw [hypreal_less_def]
-     "!! R1. [| R1 < (R2::hypreal); \
-\         !!X Y. {n. X n < Y n} : FreeUltrafilterNat ==> P; \
-\         !!X. X : Rep_hypreal(R1) ==> P; \ 
-\         !!Y. Y : Rep_hypreal(R2) ==> P |] \
-\     ==> P";
-by Auto_tac;
-qed "hypreal_lessE";
-
-Goalw [hypreal_less_def]
- "R1 < (R2::hypreal) ==> (EX X Y. {n. X n < Y n} : FreeUltrafilterNat & \
-\                                  X : Rep_hypreal(R1) & \
-\                                  Y : Rep_hypreal(R2))";
-by (Fast_tac 1);
-qed "hypreal_lessD";
-
-Goal "~ (R::hypreal) < R";
-by (res_inst_tac [("z","R")] eq_Abs_hypreal 1);
-by (auto_tac (claset(), simpset() addsimps [hypreal_less_def]));
-by (Ultra_tac 1);
-qed "hypreal_less_not_refl";
-
-(*** y < y ==> P ***)
-bind_thm("hypreal_less_irrefl",hypreal_less_not_refl RS notE);
-AddSEs [hypreal_less_irrefl];
-
-Goal "!!(x::hypreal). x < y ==> x ~= y";
-by (auto_tac (claset(), simpset() addsimps [hypreal_less_not_refl]));
-qed "hypreal_not_refl2";
-
-Goal "!!(R1::hypreal). [| R1 < R2; R2 < R3 |] ==> R1 < R3";
-by (res_inst_tac [("z","R1")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","R2")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","R3")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSIs [exI], simpset() addsimps [hypreal_less_def]));
-by (ultra_tac (claset() addIs [real_less_trans], simpset()) 1);
-qed "hypreal_less_trans";
-
-Goal "!! (R1::hypreal). [| R1 < R2; R2 < R1 |] ==> P";
-by (dtac hypreal_less_trans 1 THEN assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypreal_less_not_refl]) 1);
-qed "hypreal_less_asym";
-
-(*-------------------------------------------------------
-  TODO: The following theorem should have been proved 
-  first and then used througout the proofs as it probably 
-  makes many of them more straightforward. 
- -------------------------------------------------------*)
-Goalw [hypreal_less_def]
-      "(Abs_hypreal(hyprel^^{%n. X n}) < \
-\           Abs_hypreal(hyprel^^{%n. Y n})) = \
-\      ({n. X n < Y n} : FreeUltrafilterNat)";
-by (auto_tac (claset() addSIs [lemma_hyprel_refl], simpset()));
-by (Ultra_tac 1);
-qed "hypreal_less";
-
-(*---------------------------------------------------------------------------------
-             Hyperreals as a linearly ordered field
- ---------------------------------------------------------------------------------*)
-(*** sum order 
-Goalw [hypreal_zero_def] 
-      "[| 0 < x; 0 < y |] ==> (0::hypreal) < x + y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps
-    [hypreal_less_def,hypreal_add]));
-by (auto_tac (claset() addSIs [exI],simpset() addsimps
-    [hypreal_less_def,hypreal_add]));
-by (ultra_tac (claset() addIs [real_add_order],simpset()) 1);
-qed "hypreal_add_order";
-***)
-
-(*** mult order 
-Goalw [hypreal_zero_def] 
-          "[| 0 < x; 0 < y |] ==> (0::hypreal) < x * y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSIs [exI],simpset() addsimps
-    [hypreal_less_def,hypreal_mult]));
-by (ultra_tac (claset() addIs [rename_numerals real_mult_order],
-	       simpset()) 1);
-qed "hypreal_mult_order";
-****)
-
-
-(*---------------------------------------------------------------------------------
-                         Trichotomy of the hyperreals
-  --------------------------------------------------------------------------------*)
-
-Goalw [hyprel_def] "EX x. x: hyprel ^^ {%n. #0}";
-by (res_inst_tac [("x","%n. #0")] exI 1);
-by (Step_tac 1);
-by (auto_tac (claset() addSIs [FreeUltrafilterNat_Nat_set], simpset()));
-qed "lemma_hyprel_0r_mem";
-
-Goalw [hypreal_zero_def]"0 <  x | x = 0 | x < (0::hypreal)";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_less_def]));
-by (cut_facts_tac [lemma_hyprel_0r_mem] 1 THEN etac exE 1);
-by (dres_inst_tac [("x","xa")] spec 1);
-by (dres_inst_tac [("x","x")] spec 1);
-by (cut_inst_tac [("x","x")] lemma_hyprel_refl 1);
-by Auto_tac;
-by (dres_inst_tac [("x","x")] spec 1);
-by (dres_inst_tac [("x","xa")] spec 1);
-by Auto_tac;
-by (Ultra_tac 1);
-by (auto_tac (claset() addIs [real_linear_less2],simpset()));
-qed "hypreal_trichotomy";
-
-val prems = Goal "[| (0::hypreal) < x ==> P; \
-\                 x = 0 ==> P; \
-\                 x < 0 ==> P |] ==> P";
-by (cut_inst_tac [("x","x")] hypreal_trichotomy 1);
-by (REPEAT (eresolve_tac (disjE::prems) 1));
-qed "hypreal_trichotomyE";
-
-(*----------------------------------------------------------------------------
-            More properties of <
- ----------------------------------------------------------------------------*)
-
-Goal "((x::hypreal) < y) = (0 < y + -x)";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_add,
-    hypreal_zero_def,hypreal_minus,hypreal_less]));
-by (ALLGOALS(Ultra_tac));
-qed "hypreal_less_minus_iff"; 
-
-Goal "((x::hypreal) < y) = (x + -y < 0)";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_add,
-    hypreal_zero_def,hypreal_minus,hypreal_less]));
-by (ALLGOALS(Ultra_tac));
-qed "hypreal_less_minus_iff2";
-
-Goal "((x::hypreal) = y) = (0 = x + - y)";
-by Auto_tac;
-by (res_inst_tac [("x1","-y")] (hypreal_add_right_cancel RS iffD1) 1);
-by Auto_tac;
-qed "hypreal_eq_minus_iff"; 
-
-Goal "((x::hypreal) = y) = (0 = y + - x)";
-by Auto_tac;
-by (res_inst_tac [("x1","-x")] (hypreal_add_right_cancel RS iffD1) 1);
-by Auto_tac;
-qed "hypreal_eq_minus_iff2"; 
-
-(* 07/00 *)
-Goal "(0::hypreal) - x = -x";
-by (simp_tac (simpset() addsimps [hypreal_diff_def]) 1);
-qed "hypreal_diff_zero";
-
-Goal "x - (0::hypreal) = x";
-by (simp_tac (simpset() addsimps [hypreal_diff_def]) 1);
-qed "hypreal_diff_zero_right";
-
-Goal "x - x = (0::hypreal)";
-by (simp_tac (simpset() addsimps [hypreal_diff_def]) 1);
-qed "hypreal_diff_self";
-
-Addsimps [hypreal_diff_zero, hypreal_diff_zero_right, hypreal_diff_self];
-
-Goal "(x = y + z) = (x + -z = (y::hypreal))";
-by (auto_tac (claset(),simpset() addsimps [hypreal_add_assoc]));
-qed "hypreal_eq_minus_iff3";
-
-Goal "(x ~= a) = (x + -a ~= (0::hypreal))";
-by (auto_tac (claset() addDs [sym RS (hypreal_eq_minus_iff RS iffD2)],
-              simpset())); 
-qed "hypreal_not_eq_minus_iff";
-
-(*** linearity ***)
-Goal "(x::hypreal) < y | x = y | y < x";
-by (stac hypreal_eq_minus_iff2 1);
-by (res_inst_tac [("x1","x")] (hypreal_less_minus_iff RS ssubst) 1);
-by (res_inst_tac [("x1","y")] (hypreal_less_minus_iff2 RS ssubst) 1);
-by (rtac hypreal_trichotomyE 1);
-by Auto_tac;
-qed "hypreal_linear";
-
-Goal "((w::hypreal) ~= z) = (w<z | z<w)";
-by (cut_facts_tac [hypreal_linear] 1);
-by (Blast_tac 1);
-qed "hypreal_neq_iff";
-
-Goal "!!(x::hypreal). [| x < y ==> P;  x = y ==> P; \
-\          y < x ==> P |] ==> P";
-by (cut_inst_tac [("x","x"),("y","y")] hypreal_linear 1);
-by Auto_tac;
-qed "hypreal_linear_less2";
-
-(*------------------------------------------------------------------------------
-                            Properties of <=
- ------------------------------------------------------------------------------*)
-(*------ hypreal le iff reals le a.e ------*)
-
-Goalw [hypreal_le_def,real_le_def]
-      "(Abs_hypreal(hyprel^^{%n. X n}) <= \
-\           Abs_hypreal(hyprel^^{%n. Y n})) = \
-\      ({n. X n <= Y n} : FreeUltrafilterNat)";
-by (auto_tac (claset(),simpset() addsimps [hypreal_less]));
-by (ALLGOALS(Ultra_tac));
-qed "hypreal_le";
-
-(*---------------------------------------------------------*)
-(*---------------------------------------------------------*)
-Goalw [hypreal_le_def] 
-     "~(w < z) ==> z <= (w::hypreal)";
-by (assume_tac 1);
-qed "hypreal_leI";
-
-Goalw [hypreal_le_def] 
-      "z<=w ==> ~(w<(z::hypreal))";
-by (assume_tac 1);
-qed "hypreal_leD";
-
-bind_thm ("hypreal_leE", make_elim hypreal_leD);
-
-Goal "(~(w < z)) = (z <= (w::hypreal))";
-by (fast_tac (claset() addSIs [hypreal_leI,hypreal_leD]) 1);
-qed "hypreal_less_le_iff";
-
-Goalw [hypreal_le_def] "~ z <= w ==> w<(z::hypreal)";
-by (Fast_tac 1);
-qed "not_hypreal_leE";
-
-Goalw [hypreal_le_def] "z < w ==> z <= (w::hypreal)";
-by (fast_tac (claset() addEs [hypreal_less_asym]) 1);
-qed "hypreal_less_imp_le";
-
-Goalw [hypreal_le_def] "!!(x::hypreal). x <= y ==> x < y | x = y";
-by (cut_facts_tac [hypreal_linear] 1);
-by (fast_tac (claset() addEs [hypreal_less_irrefl,hypreal_less_asym]) 1);
-qed "hypreal_le_imp_less_or_eq";
-
-Goalw [hypreal_le_def] "z<w | z=w ==> z <=(w::hypreal)";
-by (cut_facts_tac [hypreal_linear] 1);
-by (fast_tac (claset() addEs [hypreal_less_irrefl,hypreal_less_asym]) 1);
-qed "hypreal_less_or_eq_imp_le";
-
-Goal "(x <= (y::hypreal)) = (x < y | x=y)";
-by (REPEAT(ares_tac [iffI, hypreal_less_or_eq_imp_le, hypreal_le_imp_less_or_eq] 1));
-qed "hypreal_le_eq_less_or_eq";
-val hypreal_le_less = hypreal_le_eq_less_or_eq;
-
-Goal "w <= (w::hypreal)";
-by (simp_tac (simpset() addsimps [hypreal_le_eq_less_or_eq]) 1);
-qed "hypreal_le_refl";
-Addsimps [hypreal_le_refl];
-
-(* Axiom 'linorder_linear' of class 'linorder': *)
-Goal "(z::hypreal) <= w | w <= z";
-by (simp_tac (simpset() addsimps [hypreal_le_less]) 1);
-by (cut_facts_tac [hypreal_linear] 1);
-by (Blast_tac 1);
-qed "hypreal_le_linear";
-
-Goal "[| i <= j; j < k |] ==> i < (k::hypreal)";
-by (dtac hypreal_le_imp_less_or_eq 1);
-by (fast_tac (claset() addIs [hypreal_less_trans]) 1);
-qed "hypreal_le_less_trans";
-
-Goal "!! (i::hypreal). [| i < j; j <= k |] ==> i < k";
-by (dtac hypreal_le_imp_less_or_eq 1);
-by (fast_tac (claset() addIs [hypreal_less_trans]) 1);
-qed "hypreal_less_le_trans";
-
-Goal "[| i <= j; j <= k |] ==> i <= (k::hypreal)";
-by (EVERY1 [dtac hypreal_le_imp_less_or_eq, dtac hypreal_le_imp_less_or_eq,
-            rtac hypreal_less_or_eq_imp_le, fast_tac (claset() addIs [hypreal_less_trans])]);
-qed "hypreal_le_trans";
-
-Goal "[| z <= w; w <= z |] ==> z = (w::hypreal)";
-by (EVERY1 [dtac hypreal_le_imp_less_or_eq, dtac hypreal_le_imp_less_or_eq,
-            fast_tac (claset() addEs [hypreal_less_irrefl,hypreal_less_asym])]);
-qed "hypreal_le_anti_sym";
-
-Goal "[| ~ y < x; y ~= x |] ==> x < (y::hypreal)";
-by (rtac not_hypreal_leE 1);
-by (fast_tac (claset() addDs [hypreal_le_imp_less_or_eq]) 1);
-qed "not_less_not_eq_hypreal_less";
-
-(* Axiom 'order_less_le' of class 'order': *)
-Goal "(w::hypreal) < z = (w <= z & w ~= z)";
-by (simp_tac (simpset() addsimps [hypreal_le_def, hypreal_neq_iff]) 1);
-by (blast_tac (claset() addIs [hypreal_less_asym]) 1);
-qed "hypreal_less_le";
-
-Goal "(0 < -R) = (R < (0::hypreal))";
-by (res_inst_tac [("z","R")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-       simpset() addsimps [hypreal_zero_def, hypreal_less,hypreal_minus]));
-qed "hypreal_minus_zero_less_iff";
-Addsimps [hypreal_minus_zero_less_iff];
-
-Goal "(-R < 0) = ((0::hypreal) < R)";
-by (res_inst_tac [("z","R")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-         simpset() addsimps [hypreal_zero_def, hypreal_less,hypreal_minus]));
-by (ALLGOALS(Ultra_tac));
-qed "hypreal_minus_zero_less_iff2";
-
-Goalw [hypreal_le_def] "((0::hypreal) <= -r) = (r <= (0::hypreal))";
-by (simp_tac (simpset() addsimps [hypreal_minus_zero_less_iff2]) 1);
-qed "hypreal_minus_zero_le_iff";
-Addsimps [hypreal_minus_zero_le_iff];
-
-(*----------------------------------------------------------
-  hypreal_of_real preserves field and order properties
- -----------------------------------------------------------*)
-Goalw [hypreal_of_real_def] 
-     "hypreal_of_real (z1 + z2) = hypreal_of_real z1 + hypreal_of_real z2";
-by (simp_tac (simpset() addsimps [hypreal_add, hypreal_add_mult_distrib]) 1);
-qed "hypreal_of_real_add";
-Addsimps [hypreal_of_real_add];
-
-Goalw [hypreal_of_real_def] 
-     "hypreal_of_real (z1 * z2) = hypreal_of_real z1 * hypreal_of_real z2";
-by (simp_tac (simpset() addsimps [hypreal_mult, hypreal_add_mult_distrib2]) 1);
-qed "hypreal_of_real_mult";
-Addsimps [hypreal_of_real_mult];
-
-Goalw [hypreal_less_def,hypreal_of_real_def] 
-     "(z1 < z2) = (hypreal_of_real z1 <  hypreal_of_real z2)";
-by Auto_tac;
-by (res_inst_tac [("x","%n. z1")] exI 1);
-by (Step_tac 1); 
-by (res_inst_tac [("x","%n. z2")] exI 2);
-by Auto_tac;
-by (rtac FreeUltrafilterNat_P 1);
-by (Ultra_tac 1);
-qed "hypreal_of_real_less_iff";
-
-Addsimps [hypreal_of_real_less_iff RS sym];
-
-Goalw [hypreal_le_def,real_le_def] 
-            "(z1 <= z2) = (hypreal_of_real z1 <=  hypreal_of_real z2)";
-by Auto_tac;
-qed "hypreal_of_real_le_iff";
-
-Goalw [hypreal_of_real_def] "hypreal_of_real (-r) = - hypreal_of_real  r";
-by (auto_tac (claset(),simpset() addsimps [hypreal_minus]));
-qed "hypreal_of_real_minus";
-Addsimps [hypreal_of_real_minus];
-
-(*DON'T insert this or the next one as default simprules.
-  They are used in both orientations and anyway aren't the ones we finally
-  need, which would use binary literals.*)
-Goalw [hypreal_of_real_def,hypreal_one_def] "hypreal_of_real  #1 = 1hr";
-by (Step_tac 1);
-qed "hypreal_of_real_one";
-
-Goalw [hypreal_of_real_def,hypreal_zero_def] "hypreal_of_real #0 = 0";
-by (Step_tac 1);
-qed "hypreal_of_real_zero";
-
-Goal "(hypreal_of_real r = 0) = (r = #0)";
-by (auto_tac (claset() addIs [FreeUltrafilterNat_P],
-    simpset() addsimps [hypreal_of_real_def,
-                        hypreal_zero_def,FreeUltrafilterNat_Nat_set]));
-qed "hypreal_of_real_zero_iff";
-AddIffs [hypreal_of_real_zero_iff];
-
-
-Goal "hypreal_of_real (inverse r) = inverse (hypreal_of_real r)";
-by (case_tac "r=#0" 1);
-by (asm_simp_tac (simpset() addsimps [REAL_DIVIDE_ZERO, INVERSE_ZERO, 
-                              HYPREAL_INVERSE_ZERO, hypreal_of_real_zero]) 1);
-by (res_inst_tac [("c1","hypreal_of_real r")] 
-    (hypreal_mult_left_cancel RS iffD1) 1);
-by (stac (hypreal_of_real_mult RS sym) 2); 
-by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_one]));
-qed "hypreal_of_real_inverse";
-Addsimps [hypreal_of_real_inverse];
-
-Goal "hypreal_of_real (z1 / z2) = hypreal_of_real z1 / hypreal_of_real z2";
-by (simp_tac (simpset() addsimps [hypreal_divide_def, real_divide_def]) 1);
-qed "hypreal_of_real_divide"; 
-Addsimps [hypreal_of_real_divide];
-
-
-(*** Division lemmas ***)
-
-Goal "(0::hypreal)/x = 0";
-by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); 
-qed "hypreal_zero_divide";
-
-Goal "x/1hr = x";
-by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); 
-qed "hypreal_divide_one";
-Addsimps [hypreal_zero_divide, hypreal_divide_one];
-
-Goal "(x::hypreal) * (y/z) = (x*y)/z";
-by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 1); 
-qed "hypreal_times_divide1_eq";
-
-Goal "(y/z) * (x::hypreal) = (y*x)/z";
-by (simp_tac (simpset() addsimps [hypreal_divide_def]@hypreal_mult_ac) 1); 
-qed "hypreal_times_divide2_eq";
-
-Addsimps [hypreal_times_divide1_eq, hypreal_times_divide2_eq];
-
-Goal "(x::hypreal) / (y/z) = (x*z)/y";
-by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_inverse_distrib]@
-                                 hypreal_mult_ac) 1); 
-qed "hypreal_divide_divide1_eq";
-
-Goal "((x::hypreal) / y) / z = x/(y*z)";
-by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_inverse_distrib, 
-                                  hypreal_mult_assoc]) 1); 
-qed "hypreal_divide_divide2_eq";
-
-Addsimps [hypreal_divide_divide1_eq, hypreal_divide_divide2_eq];
-
-(** As with multiplication, pull minus signs OUT of the / operator **)
-
-Goal "(-x) / (y::hypreal) = - (x/y)";
-by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); 
-qed "hypreal_minus_divide_eq";
-Addsimps [hypreal_minus_divide_eq];
-
-Goal "(x / -(y::hypreal)) = - (x/y)";
-by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_minus_inverse]) 1); 
-qed "hypreal_divide_minus_eq";
-Addsimps [hypreal_divide_minus_eq];
-
-Goal "(x+y)/(z::hypreal) = x/z + y/z";
-by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_add_mult_distrib]) 1); 
-qed "hypreal_add_divide_distrib";
-
-
-Goal "x+x=x*(1hr+1hr)";
-by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib2]) 1);
-qed "hypreal_add_self";
-
-(*FIXME: DELETE (used in Lim.ML) *)
-Goal "(z::hypreal) ~= 0 ==> x*y = (x*inverse(z))*(z*y)";
-by (asm_simp_tac (simpset() addsimps hypreal_mult_ac)  1);
-qed "lemma_chain";
-
-Goal "[|(x::hypreal) ~= 0; y ~= 0 |] ==> \
-\                   inverse(x) + inverse(y) = (x + y)*inverse(x*y)";
-by (asm_full_simp_tac (simpset() addsimps [hypreal_inverse_distrib,
-             hypreal_add_mult_distrib,hypreal_mult_assoc RS sym]) 1);
-by (stac hypreal_mult_assoc 1);
-by (rtac (hypreal_mult_left_commute RS subst) 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
-qed "hypreal_inverse_add";
-
-Goal "x = -x ==> x = (0::hypreal)";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_minus,
-    hypreal_zero_def]));
-by (Ultra_tac 1);
-qed "hypreal_self_eq_minus_self_zero";
-
-Goal "(x + x = 0) = (x = (0::hypreal))";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_add,
-    hypreal_zero_def]));
-qed "hypreal_add_self_zero_cancel";
-Addsimps [hypreal_add_self_zero_cancel];
-
-Goal "(x + x + y = y) = (x = (0::hypreal))";
-by Auto_tac;
-by (dtac (hypreal_eq_minus_iff RS iffD1) 1 THEN dtac sym 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_add_assoc]));
-qed "hypreal_add_self_zero_cancel2";
-Addsimps [hypreal_add_self_zero_cancel2];
-
-Goal "(x + (x + y) = y) = (x = (0::hypreal))";
-by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
-qed "hypreal_add_self_zero_cancel2a";
-Addsimps [hypreal_add_self_zero_cancel2a];
-
-Goal "(b = -a) = (-b = (a::hypreal))";
-by Auto_tac;
-qed "hypreal_minus_eq_swap";
-
-Goal "(-b = -a) = (b = (a::hypreal))";
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypreal_minus_eq_swap]) 1);
-qed "hypreal_minus_eq_cancel";
-Addsimps [hypreal_minus_eq_cancel];
-
-Goal "x < x + 1hr";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_add,
-    hypreal_one_def,hypreal_less]));
-qed "hypreal_less_self_add_one";
-Addsimps [hypreal_less_self_add_one];
-
-Goal "((x::hypreal) + x = y + y) = (x = y)";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_add]));
-by (ALLGOALS(Ultra_tac));
-qed "hypreal_add_self_cancel";
-Addsimps [hypreal_add_self_cancel];
-
-Goal "(y = x + - y + x) = (y = (x::hypreal))";
-by Auto_tac;
-by (dres_inst_tac [("x1","y")] 
-    (hypreal_add_right_cancel RS iffD2) 1);
-by (auto_tac (claset(),simpset() addsimps hypreal_add_ac));
-qed "hypreal_add_self_minus_cancel";
-Addsimps [hypreal_add_self_minus_cancel];
-
-Goal "(y = x + (- y + x)) = (y = (x::hypreal))";
-by (asm_full_simp_tac (simpset() addsimps 
-         [hypreal_add_assoc RS sym])1);
-qed "hypreal_add_self_minus_cancel2";
-Addsimps [hypreal_add_self_minus_cancel2];
-
-(* of course, can prove this by "transfer" as well *)
-Goal "z + -x = y + (y + (-x + -z)) = (y = (z::hypreal))";
-by Auto_tac;
-by (dres_inst_tac [("x1","z")] 
-    (hypreal_add_right_cancel RS iffD2) 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypreal_minus_add_distrib RS sym] @ hypreal_add_ac
-    delsimps [hypreal_minus_add_distrib]) 1);
-by (asm_full_simp_tac (simpset() addsimps 
-     [hypreal_add_assoc RS sym,hypreal_add_right_cancel]) 1);
-qed "hypreal_add_self_minus_cancel3";
-Addsimps [hypreal_add_self_minus_cancel3];
-
-Goal "(x * x = 0) = (x = (0::hypreal))";
-by Auto_tac;
-by (blast_tac (claset() addIs [hypreal_mult_not_0E]) 1);
-qed "hypreal_mult_self_eq_zero_iff";
-Addsimps [hypreal_mult_self_eq_zero_iff];
-
-Goalw [hypreal_diff_def] "(x<y) = (x-y < (0::hypreal))";
-by (rtac hypreal_less_minus_iff2 1);
-qed "hypreal_less_eq_diff";
-
-(*** Subtraction laws ***)
-
-Goal "x + (y - z) = (x + y) - (z::hypreal)";
-by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
-qed "hypreal_add_diff_eq";
-
-Goal "(x - y) + z = (x + z) - (y::hypreal)";
-by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
-qed "hypreal_diff_add_eq";
-
-Goal "(x - y) - z = x - (y + (z::hypreal))";
-by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
-qed "hypreal_diff_diff_eq";
-
-Goal "x - (y - z) = (x + z) - (y::hypreal)";
-by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
-qed "hypreal_diff_diff_eq2";
-
-Goal "(x-y < z) = (x < z + (y::hypreal))";
-by (stac hypreal_less_eq_diff 1);
-by (res_inst_tac [("y1", "z")] (hypreal_less_eq_diff RS ssubst) 1);
-by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
-qed "hypreal_diff_less_eq";
-
-Goal "(x < z-y) = (x + (y::hypreal) < z)";
-by (stac hypreal_less_eq_diff 1);
-by (res_inst_tac [("y1", "z-y")] (hypreal_less_eq_diff RS ssubst) 1);
-by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
-qed "hypreal_less_diff_eq";
-
-Goalw [hypreal_le_def] "(x-y <= z) = (x <= z + (y::hypreal))";
-by (simp_tac (simpset() addsimps [hypreal_less_diff_eq]) 1);
-qed "hypreal_diff_le_eq";
-
-Goalw [hypreal_le_def] "(x <= z-y) = (x + (y::hypreal) <= z)";
-by (simp_tac (simpset() addsimps [hypreal_diff_less_eq]) 1);
-qed "hypreal_le_diff_eq";
-
-Goalw [hypreal_diff_def] "(x-y = z) = (x = z + (y::hypreal))";
-by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc]));
-qed "hypreal_diff_eq_eq";
-
-Goalw [hypreal_diff_def] "(x = z-y) = (x + (y::hypreal) = z)";
-by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc]));
-qed "hypreal_eq_diff_eq";
-
-(*This list of rewrites simplifies (in)equalities by bringing subtractions
-  to the top and then moving negative terms to the other side.  
-  Use with hypreal_add_ac*)
-val hypreal_compare_rls = 
-  [symmetric hypreal_diff_def,
-   hypreal_add_diff_eq, hypreal_diff_add_eq, hypreal_diff_diff_eq, hypreal_diff_diff_eq2, 
-   hypreal_diff_less_eq, hypreal_less_diff_eq, hypreal_diff_le_eq, hypreal_le_diff_eq, 
-   hypreal_diff_eq_eq, hypreal_eq_diff_eq];
-
-
-(** For the cancellation simproc.
-    The idea is to cancel like terms on opposite sides by subtraction **)
-
-Goal "(x::hypreal) - y = x' - y' ==> (x<y) = (x'<y')";
-by (stac hypreal_less_eq_diff 1);
-by (res_inst_tac [("y1", "y")] (hypreal_less_eq_diff RS ssubst) 1);
-by (Asm_simp_tac 1);
-qed "hypreal_less_eqI";
-
-Goal "(x::hypreal) - y = x' - y' ==> (y<=x) = (y'<=x')";
-by (dtac hypreal_less_eqI 1);
-by (asm_simp_tac (simpset() addsimps [hypreal_le_def]) 1);
-qed "hypreal_le_eqI";
-
-Goal "(x::hypreal) - y = x' - y' ==> (x=y) = (x'=y')";
-by Safe_tac;
-by (ALLGOALS
-    (asm_full_simp_tac
-     (simpset() addsimps [hypreal_eq_diff_eq, hypreal_diff_eq_eq])));
-qed "hypreal_eq_eqI";
-
--- a/src/HOL/Real/Hyperreal/HyperDef.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,98 +0,0 @@
-(*  Title       : HOL/Real/Hyperreal/HyperDef.thy
-    ID          : $Id$
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Construction of hyperreals using ultrafilters
-*) 
-
-HyperDef = Filter + Real +
-
-consts 
- 
-    FreeUltrafilterNat   :: nat set set    ("\\<U>")
-
-defs
-
-    FreeUltrafilterNat_def
-    "FreeUltrafilterNat    ==   (@U. U : FreeUltrafilter (UNIV:: nat set))"
-
-
-constdefs
-    hyprel :: "((nat=>real)*(nat=>real)) set"
-    "hyprel == {p. ? X Y. p = ((X::nat=>real),Y) & 
-                   {n::nat. X(n) = Y(n)}: FreeUltrafilterNat}"
-
-typedef hypreal = "UNIV//hyprel"   (Equiv.quotient_def)
-
-instance
-   hypreal  :: {ord, zero, plus, times, minus, inverse}
-
-consts 
-
-  "1hr"       :: hypreal               ("1hr")  
-  "whr"       :: hypreal               ("whr")  
-  "ehr"       :: hypreal               ("ehr")  
-
-
-defs
-
-  hypreal_zero_def
-  "0 == Abs_hypreal(hyprel^^{%n::nat. (#0::real)})"
-
-  hypreal_one_def
-  "1hr == Abs_hypreal(hyprel^^{%n::nat. (#1::real)})"
-
-  (* an infinite number = [<1,2,3,...>] *)
-  omega_def
-  "whr == Abs_hypreal(hyprel^^{%n::nat. real_of_posnat n})"
-    
-  (* an infinitesimal number = [<1,1/2,1/3,...>] *)
-  epsilon_def
-  "ehr == Abs_hypreal(hyprel^^{%n::nat. inverse (real_of_posnat n)})"
-
-  hypreal_minus_def
-  "- P == Abs_hypreal(UN X: Rep_hypreal(P). hyprel^^{%n::nat. - (X n)})"
-
-  hypreal_diff_def 
-  "x - y == x + -(y::hypreal)"
-
-  hypreal_inverse_def
-  "inverse P == Abs_hypreal(UN X: Rep_hypreal(P). 
-                    hyprel^^{%n. if X n = #0 then #0 else inverse (X n)})"
-
-  hypreal_divide_def
-  "P / Q::hypreal == P * inverse Q"
-  
-constdefs
-
-  hypreal_of_real  :: real => hypreal                 
-  "hypreal_of_real r         == Abs_hypreal(hyprel^^{%n::nat. r})"
-  
-  (* n::nat --> (n+1)::hypreal *)
-  hypreal_of_posnat :: nat => hypreal                
-  "hypreal_of_posnat n  == (hypreal_of_real(real_of_preal
-                            (preal_of_prat(prat_of_pnat(pnat_of_nat n)))))"
-
-  hypreal_of_nat :: nat => hypreal                   
-  "hypreal_of_nat n      == hypreal_of_posnat n + -1hr"
-
-defs 
-
-  hypreal_add_def  
-  "P + Q == Abs_hypreal(UN X:Rep_hypreal(P). UN Y:Rep_hypreal(Q).
-                hyprel^^{%n::nat. X n + Y n})"
-
-  hypreal_mult_def  
-  "P * Q == Abs_hypreal(UN X:Rep_hypreal(P). UN Y:Rep_hypreal(Q).
-                hyprel^^{%n::nat. X n * Y n})"
-
-  hypreal_less_def
-  "P < (Q::hypreal) == EX X Y. X: Rep_hypreal(P) & 
-                               Y: Rep_hypreal(Q) & 
-                               {n::nat. X n < Y n} : FreeUltrafilterNat"
-  hypreal_le_def
-  "P <= (Q::hypreal) == ~(Q < P)" 
-
-end
- 
-
--- a/src/HOL/Real/Hyperreal/HyperNat.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1338 +0,0 @@
-(*  Title       : HyperNat.ML
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Explicit construction of hypernaturals using 
-                  ultrafilters
-*) 
-       
-(*------------------------------------------------------------------------
-                       Properties of hypnatrel
- ------------------------------------------------------------------------*)
-
-(** Proving that hyprel is an equivalence relation       **)
-(** Natural deduction for hypnatrel - similar to hyprel! **)
-
-Goalw [hypnatrel_def]
-   "((X,Y): hypnatrel) = ({n. X n = Y n}: FreeUltrafilterNat)";
-by (Fast_tac 1);
-qed "hypnatrel_iff";
-
-Goalw [hypnatrel_def] 
-     "{n. X n = Y n}: FreeUltrafilterNat ==> (X,Y): hypnatrel";
-by (Fast_tac 1);
-qed "hypnatrelI";
-
-Goalw [hypnatrel_def]
-  "p: hypnatrel --> (EX X Y. \
-\                 p = (X,Y) & {n. X n = Y n} : FreeUltrafilterNat)";
-by (Fast_tac 1);
-qed "hypnatrelE_lemma";
-
-val [major,minor] = goal thy
-  "[| p: hypnatrel;  \
-\     !!X Y. [| p = (X,Y); {n. X n = Y n}: FreeUltrafilterNat\
-\            |] ==> Q |] ==> Q";
-by (cut_facts_tac [major RS (hypnatrelE_lemma RS mp)] 1);
-by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
-qed "hypnatrelE";
-
-AddSIs [hypnatrelI];
-AddSEs [hypnatrelE];
-
-Goalw [hypnatrel_def] "(x,x): hypnatrel";
-by (Auto_tac);
-qed "hypnatrel_refl";
-
-Goalw [hypnatrel_def] "(x,y): hypnatrel --> (y,x):hypnatrel";
-by (auto_tac (claset() addIs [lemma_perm RS subst],simpset()));
-qed_spec_mp "hypnatrel_sym";
-
-Goalw [hypnatrel_def]
-      "(x,y): hypnatrel --> (y,z):hypnatrel --> (x,z):hypnatrel";
-by (Auto_tac);
-by (Fuf_tac 1);
-qed_spec_mp "hypnatrel_trans";
-
-Goalw [equiv_def, refl_def, sym_def, trans_def]
-    "equiv {x::nat=>nat. True} hypnatrel";
-by (auto_tac (claset() addSIs [hypnatrel_refl] addSEs 
-    [hypnatrel_sym,hypnatrel_trans] delrules [hypnatrelI,hypnatrelE],
-    simpset()));
-qed "equiv_hypnatrel";
-
-val equiv_hypnatrel_iff =
-    [TrueI, TrueI] MRS 
-    ([CollectI, CollectI] MRS 
-    (equiv_hypnatrel RS eq_equiv_class_iff));
-
-Goalw  [hypnat_def,hypnatrel_def,quotient_def] "hypnatrel^^{x}:hypnat";
-by (Blast_tac 1);
-qed "hypnatrel_in_hypnat";
-
-Goal "inj_on Abs_hypnat hypnat";
-by (rtac inj_on_inverseI 1);
-by (etac Abs_hypnat_inverse 1);
-qed "inj_on_Abs_hypnat";
-
-Addsimps [equiv_hypnatrel_iff,inj_on_Abs_hypnat RS inj_on_iff,
-          hypnatrel_iff, hypnatrel_in_hypnat, Abs_hypnat_inverse];
-
-Addsimps [equiv_hypnatrel RS eq_equiv_class_iff];
-val eq_hypnatrelD = equiv_hypnatrel RSN (2,eq_equiv_class);
-
-Goal "inj(Rep_hypnat)";
-by (rtac inj_inverseI 1);
-by (rtac Rep_hypnat_inverse 1);
-qed "inj_Rep_hypnat";
-
-Goalw [hypnatrel_def] "x: hypnatrel ^^ {x}";
-by (Step_tac 1);
-by (Auto_tac);
-qed "lemma_hypnatrel_refl";
-
-Addsimps [lemma_hypnatrel_refl];
-
-Goalw [hypnat_def] "{} ~: hypnat";
-by (auto_tac (claset() addSEs [quotientE],simpset()));
-qed "hypnat_empty_not_mem";
-
-Addsimps [hypnat_empty_not_mem];
-
-Goal "Rep_hypnat x ~= {}";
-by (cut_inst_tac [("x","x")] Rep_hypnat 1);
-by (Auto_tac);
-qed "Rep_hypnat_nonempty";
-
-Addsimps [Rep_hypnat_nonempty];
-
-(*------------------------------------------------------------------------
-   hypnat_of_nat: the injection from nat to hypnat
- ------------------------------------------------------------------------*)
-Goal "inj(hypnat_of_nat)";
-by (rtac injI 1);
-by (rewtac hypnat_of_nat_def);
-by (dtac (inj_on_Abs_hypnat RS inj_onD) 1);
-by (REPEAT (rtac hypnatrel_in_hypnat 1));
-by (dtac eq_equiv_class 1);
-by (rtac equiv_hypnatrel 1);
-by (Fast_tac 1);
-by (rtac ccontr 1 THEN rotate_tac 1 1);
-by (Auto_tac);
-qed "inj_hypnat_of_nat";
-
-val [prem] = goal thy
-    "(!!x. z = Abs_hypnat(hypnatrel^^{x}) ==> P) ==> P";
-by (res_inst_tac [("x1","z")] 
-    (rewrite_rule [hypnat_def] Rep_hypnat RS quotientE) 1);
-by (dres_inst_tac [("f","Abs_hypnat")] arg_cong 1);
-by (res_inst_tac [("x","x")] prem 1);
-by (asm_full_simp_tac (simpset() addsimps [Rep_hypnat_inverse]) 1);
-qed "eq_Abs_hypnat";
-
-(*-----------------------------------------------------------
-   Addition for hyper naturals: hypnat_add 
- -----------------------------------------------------------*)
-Goalw [congruent2_def]
-    "congruent2 hypnatrel (%X Y. hypnatrel^^{%n. X n + Y n})";
-by Safe_tac;
-by (ALLGOALS(Fuf_tac));
-qed "hypnat_add_congruent2";
-
-Goalw [hypnat_add_def]
-  "Abs_hypnat(hypnatrel^^{%n. X n}) + Abs_hypnat(hypnatrel^^{%n. Y n}) = \
-\  Abs_hypnat(hypnatrel^^{%n. X n + Y n})";
-by (asm_simp_tac
-    (simpset() addsimps [[equiv_hypnatrel, hypnat_add_congruent2] 
-     MRS UN_equiv_class2]) 1);
-qed "hypnat_add";
-
-Goal "(z::hypnat) + w = w + z";
-by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","w")] eq_Abs_hypnat 1);
-by (asm_simp_tac (simpset() addsimps (add_ac @ [hypnat_add])) 1);
-qed "hypnat_add_commute";
-
-Goal "((z1::hypnat) + z2) + z3 = z1 + (z2 + z3)";
-by (res_inst_tac [("z","z1")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","z2")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","z3")] eq_Abs_hypnat 1);
-by (asm_simp_tac (simpset() addsimps [hypnat_add,add_assoc]) 1);
-qed "hypnat_add_assoc";
-
-(*For AC rewriting*)
-Goal "(x::hypnat)+(y+z)=y+(x+z)";
-by (rtac (hypnat_add_commute RS trans) 1);
-by (rtac (hypnat_add_assoc RS trans) 1);
-by (rtac (hypnat_add_commute RS arg_cong) 1);
-qed "hypnat_add_left_commute";
-
-(* hypnat addition is an AC operator *)
-val hypnat_add_ac = [hypnat_add_assoc,hypnat_add_commute,
-                      hypnat_add_left_commute];
-
-Goalw [hypnat_zero_def] "(0::hypnat) + z = z";
-by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypnat_add]) 1);
-qed "hypnat_add_zero_left";
-
-Goal "z + (0::hypnat) = z";
-by (simp_tac (simpset() addsimps 
-    [hypnat_add_zero_left,hypnat_add_commute]) 1);
-qed "hypnat_add_zero_right";
-
-Addsimps [hypnat_add_zero_left,hypnat_add_zero_right];
-
-(*-----------------------------------------------------------
-   Subtraction for hyper naturals: hypnat_minus
- -----------------------------------------------------------*)
-Goalw [congruent2_def]
-    "congruent2 hypnatrel (%X Y. hypnatrel^^{%n. X n - Y n})";
-by Safe_tac;
-by (ALLGOALS(Fuf_tac));
-qed "hypnat_minus_congruent2";
- 
-Goalw [hypnat_minus_def]
-  "Abs_hypnat(hypnatrel^^{%n. X n}) - Abs_hypnat(hypnatrel^^{%n. Y n}) = \
-\  Abs_hypnat(hypnatrel^^{%n. X n - Y n})";
-by (asm_simp_tac
-    (simpset() addsimps [[equiv_hypnatrel, hypnat_minus_congruent2] 
-     MRS UN_equiv_class2]) 1);
-qed "hypnat_minus";
-
-Goalw [hypnat_zero_def] "z - z = (0::hypnat)";
-by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
-by (asm_full_simp_tac (simpset() addsimps [hypnat_minus]) 1);
-qed "hypnat_minus_zero";
-
-Goalw [hypnat_zero_def] "(0::hypnat) - n = 0";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (asm_full_simp_tac (simpset() addsimps [hypnat_minus]) 1);
-qed "hypnat_diff_0_eq_0";
-
-Addsimps [hypnat_minus_zero,hypnat_diff_0_eq_0];
-
-Goalw [hypnat_zero_def] "(m+n = (0::hypnat)) = (m=0 & n=0)";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
-by (auto_tac (claset() addIs [FreeUltrafilterNat_subset]
-    addDs [FreeUltrafilterNat_Int],
-    simpset() addsimps [hypnat_add] ));
-qed "hypnat_add_is_0";
-
-AddIffs [hypnat_add_is_0];
-
-Goal "!!i::hypnat. i-j-k = i - (j+k)";
-by (res_inst_tac [("z","i")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","j")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","k")] eq_Abs_hypnat 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypnat_minus,hypnat_add,diff_diff_left]) 1);
-qed "hypnat_diff_diff_left";
-
-Goal "!!i::hypnat. i-j-k = i-k-j";
-by (simp_tac (simpset() addsimps 
-    [hypnat_diff_diff_left, hypnat_add_commute]) 1);
-qed "hypnat_diff_commute";
-
-Goal "!!n::hypnat. (n+m) - n = m";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypnat_minus,hypnat_add]) 1);
-qed "hypnat_diff_add_inverse";
-Addsimps [hypnat_diff_add_inverse];
-
-Goal "!!n::hypnat.(m+n) - n = m";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypnat_minus,hypnat_add]) 1);
-qed "hypnat_diff_add_inverse2";
-Addsimps [hypnat_diff_add_inverse2];
-
-Goal "!!k::hypnat. (k+m) - (k+n) = m - n";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","k")] eq_Abs_hypnat 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypnat_minus,hypnat_add]) 1);
-qed "hypnat_diff_cancel";
-Addsimps [hypnat_diff_cancel];
-
-Goal "!!m::hypnat. (m+k) - (n+k) = m - n";
-val hypnat_add_commute_k = read_instantiate [("w","k")] hypnat_add_commute;
-by (asm_simp_tac (simpset() addsimps ([hypnat_add_commute_k])) 1);
-qed "hypnat_diff_cancel2";
-Addsimps [hypnat_diff_cancel2];
-
-Goalw [hypnat_zero_def] "!!n::hypnat. n - (n+m) = (0::hypnat)";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypnat_minus,hypnat_add]) 1);
-qed "hypnat_diff_add_0";
-Addsimps [hypnat_diff_add_0];
-
-(*-----------------------------------------------------------
-   Multiplication for hyper naturals: hypnat_mult
- -----------------------------------------------------------*)
-Goalw [congruent2_def]
-    "congruent2 hypnatrel (%X Y. hypnatrel^^{%n. X n * Y n})";
-by Safe_tac;
-by (ALLGOALS(Fuf_tac));
-qed "hypnat_mult_congruent2";
-
-Goalw [hypnat_mult_def]
-  "Abs_hypnat(hypnatrel^^{%n. X n}) * Abs_hypnat(hypnatrel^^{%n. Y n}) = \
-\  Abs_hypnat(hypnatrel^^{%n. X n * Y n})";
-by (asm_simp_tac
-    (simpset() addsimps [[equiv_hypnatrel,hypnat_mult_congruent2] MRS
-     UN_equiv_class2]) 1);
-qed "hypnat_mult";
-
-Goal "(z::hypnat) * w = w * z";
-by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","w")] eq_Abs_hypnat 1);
-by (asm_simp_tac (simpset() addsimps ([hypnat_mult] @ mult_ac)) 1);
-qed "hypnat_mult_commute";
-
-Goal "((z1::hypnat) * z2) * z3 = z1 * (z2 * z3)";
-by (res_inst_tac [("z","z1")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","z2")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","z3")] eq_Abs_hypnat 1);
-by (asm_simp_tac (simpset() addsimps [hypnat_mult,mult_assoc]) 1);
-qed "hypnat_mult_assoc";
-
-qed_goal "hypnat_mult_left_commute" thy
-    "(z1::hypnat) * (z2 * z3) = z2 * (z1 * z3)"
- (fn _ => [rtac (hypnat_mult_commute RS trans) 1, rtac (hypnat_mult_assoc RS trans) 1,
-           rtac (hypnat_mult_commute RS arg_cong) 1]);
-
-(* hypnat multiplication is an AC operator *)
-val hypnat_mult_ac = [hypnat_mult_assoc, hypnat_mult_commute, 
-                       hypnat_mult_left_commute];
-
-Goalw [hypnat_one_def] "1hn * z = z";
-by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
-by (asm_full_simp_tac (simpset() addsimps [hypnat_mult]) 1);
-qed "hypnat_mult_1";
-Addsimps [hypnat_mult_1];
-
-Goal "z * 1hn = z";
-by (simp_tac (simpset() addsimps [hypnat_mult_commute]) 1);
-qed "hypnat_mult_1_right";
-Addsimps [hypnat_mult_1_right];
-
-Goalw [hypnat_zero_def] "(0::hypnat) * z = 0";
-by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
-by (asm_full_simp_tac (simpset() addsimps [hypnat_mult]) 1);
-qed "hypnat_mult_0";
-Addsimps [hypnat_mult_0];
-
-Goal "z * (0::hypnat) = 0";
-by (simp_tac (simpset() addsimps [hypnat_mult_commute]) 1);
-qed "hypnat_mult_0_right";
-Addsimps [hypnat_mult_0_right];
-
-Goal "!!m::hypnat. (m - n) * k = (m * k) - (n * k)";
-by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","k")] eq_Abs_hypnat 1);
-by (asm_simp_tac (simpset() addsimps [hypnat_mult,
-    hypnat_minus,diff_mult_distrib]) 1);
-qed "hypnat_diff_mult_distrib" ;
-
-Goal "!!m::hypnat. k * (m - n) = (k * m) - (k * n)";
-val hypnat_mult_commute_k = read_instantiate [("z","k")] hypnat_mult_commute;
-by (simp_tac (simpset() addsimps [hypnat_diff_mult_distrib, 
-    hypnat_mult_commute_k]) 1);
-qed "hypnat_diff_mult_distrib2" ;
-
-(*--------------------------
-    A Few more theorems
- -------------------------*)
-qed_goal "hypnat_add_assoc_cong" thy
-    "!!z. (z::hypnat) + v = z' + v' ==> z + (v + w) = z' + (v' + w)"
- (fn _ => [(asm_simp_tac (simpset() addsimps [hypnat_add_assoc RS sym]) 1)]);
-
-qed_goal "hypnat_add_assoc_swap" thy 
-          "(z::hypnat) + (v + w) = v + (z + w)"
- (fn _ => [(REPEAT (ares_tac [hypnat_add_commute RS hypnat_add_assoc_cong] 1))]);
-
-Goal "((z1::hypnat) + z2) * w = (z1 * w) + (z2 * w)";
-by (res_inst_tac [("z","z1")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","z2")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","w")] eq_Abs_hypnat 1);
-by (asm_simp_tac (simpset() addsimps [hypnat_mult,hypnat_add,
-     add_mult_distrib]) 1);
-qed "hypnat_add_mult_distrib";
-Addsimps [hypnat_add_mult_distrib];
-
-val hypnat_mult_commute'= read_instantiate [("z","w")] hypnat_mult_commute;
-
-Goal "(w::hypnat) * (z1 + z2) = (w * z1) + (w * z2)";
-by (simp_tac (simpset() addsimps [hypnat_mult_commute']) 1);
-qed "hypnat_add_mult_distrib2";
-Addsimps [hypnat_add_mult_distrib2];
-
-(*** (hypnat) one and zero are distinct ***)
-Goalw [hypnat_zero_def,hypnat_one_def] "(0::hypnat) ~= 1hn";
-by (Auto_tac);
-qed "hypnat_zero_not_eq_one";
-Addsimps [hypnat_zero_not_eq_one];
-Addsimps [hypnat_zero_not_eq_one RS not_sym];
-
-Goalw [hypnat_zero_def] "(m*n = (0::hypnat)) = (m=0 | n=0)";
-by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem],
-    simpset() addsimps [hypnat_mult]));
-by (ALLGOALS(Fuf_tac));
-qed "hypnat_mult_is_0";
-Addsimps [hypnat_mult_is_0];
-
-(*------------------------------------------------------------------
-                   Theorems for ordering 
- ------------------------------------------------------------------*)
-
-(* prove introduction and elimination rules for hypnat_less *)
-
-Goalw [hypnat_less_def]
- "P < (Q::hypnat) = (EX X Y. X : Rep_hypnat(P) & \
-\                             Y : Rep_hypnat(Q) & \
-\                             {n. X n < Y n} : FreeUltrafilterNat)";
-by (Fast_tac 1);
-qed "hypnat_less_iff";
-
-Goalw [hypnat_less_def]
- "!!P. [| {n. X n < Y n} : FreeUltrafilterNat; \
-\         X : Rep_hypnat(P); \
-\         Y : Rep_hypnat(Q) |] ==> P < (Q::hypnat)";
-by (Fast_tac 1);
-qed "hypnat_lessI";
-
-Goalw [hypnat_less_def]
-     "!! R1. [| R1 < (R2::hypnat); \
-\         !!X Y. {n. X n < Y n} : FreeUltrafilterNat ==> P; \
-\         !!X. X : Rep_hypnat(R1) ==> P; \ 
-\         !!Y. Y : Rep_hypnat(R2) ==> P |] \
-\     ==> P";
-by (Auto_tac);
-qed "hypnat_lessE";
-
-Goalw [hypnat_less_def]
- "!!R1. R1 < (R2::hypnat) ==> (EX X Y. {n. X n < Y n} : FreeUltrafilterNat & \
-\                                  X : Rep_hypnat(R1) & \
-\                                  Y : Rep_hypnat(R2))";
-by (Fast_tac 1);
-qed "hypnat_lessD";
-
-Goal "~ (R::hypnat) < R";
-by (res_inst_tac [("z","R")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),simpset() addsimps [hypnat_less_def]));
-by (Fuf_empty_tac 1);
-qed "hypnat_less_not_refl";
-Addsimps [hypnat_less_not_refl];
-
-bind_thm("hypnat_less_irrefl",hypnat_less_not_refl RS notE);
-
-qed_goal "hypnat_not_refl2" thy 
-    "!!(x::hypnat). x < y ==> x ~= y" (fn _ => [Auto_tac]);
-
-Goalw [hypnat_less_def,hypnat_zero_def] "~ n<(0::hypnat)";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (Auto_tac);
-by (Fuf_empty_tac 1);
-qed "hypnat_not_less0";
-AddIffs [hypnat_not_less0];
-
-(* n<(0::hypnat) ==> R *)
-bind_thm ("hypnat_less_zeroE", hypnat_not_less0 RS notE);
-
-Goalw [hypnat_less_def,hypnat_zero_def,hypnat_one_def]
-      "(n<1hn) = (n=0)";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (auto_tac (claset() addSIs [exI] addEs 
-    [FreeUltrafilterNat_subset],simpset()));
-by (Fuf_tac 1);
-qed "hypnat_less_one";
-AddIffs [hypnat_less_one];
-
-Goal "!!(R1::hypnat). [| R1 < R2; R2 < R3 |] ==> R1 < R3";
-by (res_inst_tac [("z","R1")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","R2")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","R3")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),simpset() addsimps [hypnat_less_def]));
-by (res_inst_tac [("x","X")] exI 1);
-by (Auto_tac);
-by (res_inst_tac [("x","Ya")] exI 1);
-by (Auto_tac);
-by (Fuf_tac 1);
-qed "hypnat_less_trans";
-
-Goal "!! (R1::hypnat). [| R1 < R2; R2 < R1 |] ==> P";
-by (dtac hypnat_less_trans 1 THEN assume_tac 1);
-by (Asm_full_simp_tac 1);
-qed "hypnat_less_asym";
-
-(*----- hypnat less iff less a.e -----*)
-(* See comments in HYPER for corresponding thm *)
-
-Goalw [hypnat_less_def]
-      "(Abs_hypnat(hypnatrel^^{%n. X n}) < \
-\           Abs_hypnat(hypnatrel^^{%n. Y n})) = \
-\      ({n. X n < Y n} : FreeUltrafilterNat)";
-by (auto_tac (claset() addSIs [lemma_hypnatrel_refl],simpset()));
-by (Fuf_tac 1);
-qed "hypnat_less";
-
-Goal "~ m<n --> n+(m-n) = (m::hypnat)";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [hypnat_minus,hypnat_add,hypnat_less]));
-by (dtac FreeUltrafilterNat_Compl_mem 1);
-by (Fuf_tac 1);
-qed_spec_mp "hypnat_add_diff_inverse";
-
-Goal "n<=m ==> n+(m-n) = (m::hypnat)";
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypnat_add_diff_inverse, hypnat_le_def]) 1);
-qed "hypnat_le_add_diff_inverse";
-
-Goal "n<=m ==> (m-n)+n = (m::hypnat)";
-by (asm_simp_tac (simpset() addsimps [hypnat_le_add_diff_inverse, 
-    hypnat_add_commute]) 1);
-qed "hypnat_le_add_diff_inverse2";
-
-(*---------------------------------------------------------------------------------
-                    Hyper naturals as a linearly ordered field
- ---------------------------------------------------------------------------------*)
-Goalw [hypnat_zero_def] 
-     "[| (0::hypnat) < x; 0 < y |] ==> 0 < x + y";
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),simpset() addsimps
-   [hypnat_less_def,hypnat_add]));
-by (REPEAT(Step_tac 1));
-by (Fuf_tac 1);
-qed "hypnat_add_order";
-
-Goalw [hypnat_zero_def] 
-      "!!(x::hypnat). [| (0::hypnat) < x; 0 < y |] ==> 0 < x * y";
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [hypnat_less_def,hypnat_mult]));
-by (REPEAT(Step_tac 1));
-by (Fuf_tac 1);
-qed "hypnat_mult_order";
-
-(*---------------------------------------------------------------------------------
-                   Trichotomy of the hyper naturals
-  --------------------------------------------------------------------------------*)
-Goalw [hypnatrel_def] "EX x. x: hypnatrel ^^ {%n. 0}";
-by (res_inst_tac [("x","%n. 0")] exI 1);
-by (Step_tac 1);
-by (Auto_tac);
-qed "lemma_hypnatrel_0_mem";
-
-(* linearity is actually proved further down! *)
-Goalw [hypnat_zero_def,
-       hypnat_less_def]"(0::hypnat) <  x | x = 0 | x < 0";
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (Auto_tac );
-by (REPEAT(Step_tac 1));
-by (REPEAT(dtac FreeUltrafilterNat_Compl_mem 1));
-by (Fuf_tac 1);
-qed "hypnat_trichotomy";
-
-Goal "!!x. [| (0::hypnat) < x ==> P; \
-\                 x = 0 ==> P; \
-\                 x < 0 ==> P |] ==> P";
-by (cut_inst_tac [("x","x")] hypnat_trichotomy 1);
-by (Auto_tac);
-qed "hypnat_trichotomyE";
-
-(*------------------------------------------------------------------------------
-            More properties of <
- ------------------------------------------------------------------------------*)
-Goal "!!(A::hypnat). A < B ==> A + C < B + C";
-by (res_inst_tac [("z","A")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","B")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","C")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [hypnat_less_def,hypnat_add]));
-by (REPEAT(Step_tac 1));
-by (Fuf_tac 1);
-qed "hypnat_add_less_mono1";
-
-Goal "!!(A::hypnat). A < B ==> C + A < C + B";
-by (auto_tac (claset() addIs [hypnat_add_less_mono1],
-    simpset() addsimps [hypnat_add_commute]));
-qed "hypnat_add_less_mono2";
-
-Goal "!!k l::hypnat. [|i<j;  k<l |] ==> i + k < j + l";
-by (etac (hypnat_add_less_mono1 RS hypnat_less_trans) 1);
-by (simp_tac (simpset() addsimps [hypnat_add_commute]) 1);
-(*j moves to the end because it is free while k, l are bound*)
-by (etac hypnat_add_less_mono1 1);
-qed "hypnat_add_less_mono";
-
-(*---------------------------------------
-        hypnat_minus_less
- ---------------------------------------*)
-Goalw [hypnat_less_def,hypnat_zero_def] 
-      "((x::hypnat) < y) = ((0::hypnat) < y - x)";
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [hypnat_minus]));
-by (REPEAT(Step_tac 1));
-by (REPEAT(Step_tac 2));
-by (ALLGOALS(fuf_tac (claset() addDs [sym],simpset())));
-
-(*** linearity ***)
-Goalw [hypnat_less_def] "(x::hypnat) < y | x = y | y < x";
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
-by (Auto_tac );
-by (REPEAT(Step_tac 1));
-by (REPEAT(dtac FreeUltrafilterNat_Compl_mem 1));
-by (Fuf_tac 1);
-qed "hypnat_linear";
-
-Goal
-    "!!(x::hypnat). [| x < y ==> P;  x = y ==> P; \
-\          y < x ==> P |] ==> P";
-by (cut_inst_tac [("x","x"),("y","y")] hypnat_linear 1);
-by (Auto_tac);
-qed "hypnat_linear_less2";
-
-(*------------------------------------------------------------------------------
-                            Properties of <=
- ------------------------------------------------------------------------------*)
-(*------ hypnat le iff nat le a.e ------*)
-Goalw [hypnat_le_def,le_def]
-      "(Abs_hypnat(hypnatrel^^{%n. X n}) <= \
-\           Abs_hypnat(hypnatrel^^{%n. Y n})) = \
-\      ({n. X n <= Y n} : FreeUltrafilterNat)";
-by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem],
-    simpset() addsimps [hypnat_less]));
-by (Fuf_tac 1 THEN Fuf_empty_tac 1);
-qed "hypnat_le";
-
-(*---------------------------------------------------------*)
-(*---------------------------------------------------------*)
-Goalw [hypnat_le_def] "!!w. ~(w < z) ==> z <= (w::hypnat)";
-by (assume_tac 1);
-qed "hypnat_leI";
-
-Goalw [hypnat_le_def] "!!w. z<=w ==> ~(w<(z::hypnat))";
-by (assume_tac 1);
-qed "hypnat_leD";
-
-val hypnat_leE = make_elim hypnat_leD;
-
-Goal "!!w. (~(w < z)) = (z <= (w::hypnat))";
-by (fast_tac (claset() addSIs [hypnat_leI,hypnat_leD]) 1);
-qed "hypnat_less_le_iff";
-
-Goalw [hypnat_le_def] "!!z. ~ z <= w ==> w<(z::hypnat)";
-by (Fast_tac 1);
-qed "not_hypnat_leE";
-
-Goalw [hypnat_le_def] "!!z. z < w ==> z <= (w::hypnat)";
-by (fast_tac (claset() addEs [hypnat_less_asym]) 1);
-qed "hypnat_less_imp_le";
-
-Goalw [hypnat_le_def] "!!(x::hypnat). x <= y ==> x < y | x = y";
-by (cut_facts_tac [hypnat_linear] 1);
-by (fast_tac (claset() addEs [hypnat_less_irrefl,hypnat_less_asym]) 1);
-qed "hypnat_le_imp_less_or_eq";
-
-Goalw [hypnat_le_def] "!!z. z<w | z=w ==> z <=(w::hypnat)";
-by (cut_facts_tac [hypnat_linear] 1);
-by (fast_tac (claset() addEs [hypnat_less_irrefl,hypnat_less_asym]) 1);
-qed "hypnat_less_or_eq_imp_le";
-
-Goal "(x <= (y::hypnat)) = (x < y | x=y)";
-by (REPEAT(ares_tac [iffI, hypnat_less_or_eq_imp_le, hypnat_le_imp_less_or_eq] 1));
-qed "hypnat_le_eq_less_or_eq";
-
-Goal "w <= (w::hypnat)";
-by (simp_tac (simpset() addsimps [hypnat_le_eq_less_or_eq]) 1);
-qed "hypnat_le_refl";
-Addsimps [hypnat_le_refl];
-
-val prems = goal thy "!!i. [| i <= j; j < k |] ==> i < (k::hypnat)";
-by (dtac hypnat_le_imp_less_or_eq 1);
-by (fast_tac (claset() addIs [hypnat_less_trans]) 1);
-qed "hypnat_le_less_trans";
-
-Goal "!! (i::hypnat). [| i < j; j <= k |] ==> i < k";
-by (dtac hypnat_le_imp_less_or_eq 1);
-by (fast_tac (claset() addIs [hypnat_less_trans]) 1);
-qed "hypnat_less_le_trans";
-
-Goal "!!i. [| i <= j; j <= k |] ==> i <= (k::hypnat)";
-by (EVERY1 [dtac hypnat_le_imp_less_or_eq, dtac hypnat_le_imp_less_or_eq,
-            rtac hypnat_less_or_eq_imp_le, fast_tac (claset() addIs [hypnat_less_trans])]);
-qed "hypnat_le_trans";
-
-Goal "!!z. [| z <= w; w <= z |] ==> z = (w::hypnat)";
-by (EVERY1 [dtac hypnat_le_imp_less_or_eq, dtac hypnat_le_imp_less_or_eq,
-            fast_tac (claset() addEs [hypnat_less_irrefl,hypnat_less_asym])]);
-qed "hypnat_le_anti_sym";
-
-Goal "!!x. [| ~ y < x; y ~= x |] ==> x < (y::hypnat)";
-by (rtac not_hypnat_leE 1);
-by (fast_tac (claset() addDs [hypnat_le_imp_less_or_eq]) 1);
-qed "not_less_not_eq_hypnat_less";
-
-Goal "!!x. [| (0::hypnat) <= x; 0 <= y |] ==> 0 <= x * y";
-by (REPEAT(dtac hypnat_le_imp_less_or_eq 1));
-by (auto_tac (claset() addIs [hypnat_mult_order,
-    hypnat_less_imp_le],simpset() addsimps [hypnat_le_refl]));
-qed "hypnat_le_mult_order";
-
-Goalw [hypnat_one_def,hypnat_zero_def,hypnat_less_def] 
-      "(0::hypnat) < 1hn";
-by (res_inst_tac [("x","%n. 0")] exI 1);
-by (res_inst_tac [("x","%n. 1")] exI 1);
-by (Auto_tac);
-qed "hypnat_zero_less_one";
-
-Goal "!!x. [| (0::hypnat) <= x; 0 <= y |] ==> 0 <= x + y";
-by (REPEAT(dtac hypnat_le_imp_less_or_eq 1));
-by (auto_tac (claset() addIs [hypnat_add_order,
-    hypnat_less_imp_le],simpset() addsimps [hypnat_le_refl]));
-qed "hypnat_le_add_order";
-
-Goal "!!(q1::hypnat). q1 <= q2  ==> x + q1 <= x + q2";
-by (dtac hypnat_le_imp_less_or_eq 1);
-by (Step_tac 1);
-by (auto_tac (claset() addSIs [hypnat_le_refl,
-    hypnat_less_imp_le,hypnat_add_less_mono1],
-    simpset() addsimps [hypnat_add_commute]));
-qed "hypnat_add_le_mono2";
-
-Goal "!!(q1::hypnat). q1 <= q2  ==> q1 + x <= q2 + x";
-by (auto_tac (claset() addDs [hypnat_add_le_mono2],
-    simpset() addsimps [hypnat_add_commute]));
-qed "hypnat_add_le_mono1";
-
-Goal "!!k l::hypnat. [|i<=j;  k<=l |] ==> i + k <= j + l";
-by (etac (hypnat_add_le_mono1 RS hypnat_le_trans) 1);
-by (simp_tac (simpset() addsimps [hypnat_add_commute]) 1);
-(*j moves to the end because it is free while k, l are bound*)
-by (etac hypnat_add_le_mono1 1);
-qed "hypnat_add_le_mono";
-
-Goalw [hypnat_zero_def]
-     "!!x::hypnat. [| (0::hypnat) < z; x < y |] ==> x * z < y * z";
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","z")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [hypnat_less,hypnat_mult]));
-by (Fuf_tac 1);
-qed "hypnat_mult_less_mono1";
-
-Goal "!!x::hypnat. [| 0 < z; x < y |] ==> z * x < z * y";
-by (auto_tac (claset() addIs [hypnat_mult_less_mono1],
-              simpset() addsimps [hypnat_mult_commute]));
-qed "hypnat_mult_less_mono2";
-
-Addsimps [hypnat_mult_less_mono2,hypnat_mult_less_mono1];
-
-Goal "(x::hypnat) <= n + x";
-by (res_inst_tac [("x","n")] hypnat_trichotomyE 1);
-by (auto_tac (claset() addDs [(hypnat_less_imp_le RS
-    hypnat_add_le_mono1)],simpset() addsimps [hypnat_le_refl]));
-qed "hypnat_add_self_le";
-Addsimps [hypnat_add_self_le];
-
-Goal "(x::hypnat) < x + 1hn";
-by (cut_facts_tac [hypnat_zero_less_one 
-         RS hypnat_add_less_mono2] 1);
-by (Auto_tac);
-qed "hypnat_add_one_self_less";
-Addsimps [hypnat_add_one_self_less];
-
-Goalw [hypnat_le_def] "~ x + 1hn <= x";
-by (Simp_tac 1);
-qed "not_hypnat_add_one_le_self";
-Addsimps [not_hypnat_add_one_le_self];
-
-Goal "((0::hypnat) < n) = (1hn <= n)";
-by (res_inst_tac [("x","n")] hypnat_trichotomyE 1);
-by (auto_tac (claset(),simpset() addsimps [hypnat_le_def]));
-qed "hypnat_gt_zero_iff";
-
-Addsimps  [hypnat_le_add_diff_inverse, hypnat_le_add_diff_inverse2,
-           hypnat_less_imp_le RS hypnat_le_add_diff_inverse2];
-
-Goal "(0 < n) = (EX m. n = m + 1hn)";
-by (Step_tac 1);
-by (res_inst_tac [("x","m")] hypnat_trichotomyE 2);
-by (rtac hypnat_less_trans 2);
-by (res_inst_tac [("x","n - 1hn")] exI 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [hypnat_gt_zero_iff,hypnat_add_commute]));
-qed "hypnat_gt_zero_iff2";
-
-Goalw [hypnat_zero_def] "(0::hypnat) <= n";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (asm_simp_tac (simpset() addsimps [hypnat_le]) 1);
-qed "hypnat_le_zero";
-Addsimps [hypnat_le_zero];
-
-(*------------------------------------------------------------------
-     hypnat_of_nat: properties embedding of naturals in hypernaturals
- -----------------------------------------------------------------*)
-    (** hypnat_of_nat preserves field and order properties **)
-
-Goalw [hypnat_of_nat_def] 
-      "hypnat_of_nat ((z1::nat) + z2) = \
-\      hypnat_of_nat z1 + hypnat_of_nat z2";
-by (asm_simp_tac (simpset() addsimps [hypnat_add]) 1);
-qed "hypnat_of_nat_add";
-
-Goalw [hypnat_of_nat_def] 
-      "hypnat_of_nat ((z1::nat) - z2) = \
-\      hypnat_of_nat z1 - hypnat_of_nat z2";
-by (asm_simp_tac (simpset() addsimps [hypnat_minus]) 1);
-qed "hypnat_of_nat_minus";
-
-Goalw [hypnat_of_nat_def] 
-      "hypnat_of_nat (z1 * z2) = hypnat_of_nat z1 * hypnat_of_nat z2";
-by (full_simp_tac (simpset() addsimps [hypnat_mult]) 1);
-qed "hypnat_of_nat_mult";
-
-Goalw [hypnat_less_def,hypnat_of_nat_def] 
-      "(z1 < z2) = (hypnat_of_nat z1 < hypnat_of_nat z2)";
-by (auto_tac (claset() addSIs [exI] addIs 
-   [FreeUltrafilterNat_all],simpset()));
-by (rtac FreeUltrafilterNat_P 1 THEN Fuf_tac 1);
-qed "hypnat_of_nat_less_iff";
-Addsimps [hypnat_of_nat_less_iff RS sym];
-
-Goalw [hypnat_le_def,le_def] 
-      "(z1 <= z2) = (hypnat_of_nat z1 <= hypnat_of_nat z2)";
-by (Auto_tac);
-qed "hypnat_of_nat_le_iff";
-
-Goalw [hypnat_of_nat_def,hypnat_one_def] "hypnat_of_nat  1 = 1hn";
-by (Simp_tac 1);
-qed "hypnat_of_nat_one";
-
-Goalw [hypnat_of_nat_def,hypnat_zero_def] "hypnat_of_nat  0 = 0";
-by (Simp_tac 1);
-qed "hypnat_of_nat_zero";
-
-Goal "(hypnat_of_nat  n = 0) = (n = 0)";
-by (auto_tac (claset() addIs [FreeUltrafilterNat_P],
-    simpset() addsimps [hypnat_of_nat_def,
-    hypnat_zero_def]));
-qed "hypnat_of_nat_zero_iff";
-
-Goal "(hypnat_of_nat  n ~= 0) = (n ~= 0)";
-by (full_simp_tac (simpset() addsimps [hypnat_of_nat_zero_iff]) 1);
-qed "hypnat_of_nat_not_zero_iff";
-
-goalw HyperNat.thy [hypnat_of_nat_def,hypnat_one_def]
-      "hypnat_of_nat (Suc n) = hypnat_of_nat n + 1hn";
-by (auto_tac (claset(),simpset() addsimps [hypnat_add]));
-qed "hypnat_of_nat_Suc";
-
-(*---------------------------------------------------------------------------------
-              Existence of infinite hypernatural number
- ---------------------------------------------------------------------------------*)
-
-Goal "hypnatrel^^{%n::nat. n} : hypnat";
-by (Auto_tac);
-qed "hypnat_omega";
-
-Goalw [hypnat_omega_def] "Rep_hypnat(whn) : hypnat";
-by (rtac Rep_hypnat 1);
-qed "Rep_hypnat_omega";
-
-(* See Hyper.thy for similar argument*)
-(* existence of infinite number not corresponding to any natural number *)
-(* use assumption that member FreeUltrafilterNat is not finite       *)
-(* a few lemmas first *)
-
-Goalw [hypnat_omega_def,hypnat_of_nat_def]
-      "~ (EX x. hypnat_of_nat x = whn)";
-by (auto_tac (claset() addDs [FreeUltrafilterNat_not_finite], 
-              simpset()));
-qed "not_ex_hypnat_of_nat_eq_omega";
-
-Goal "hypnat_of_nat x ~= whn";
-by (cut_facts_tac [not_ex_hypnat_of_nat_eq_omega] 1);
-by (Auto_tac);
-qed "hypnat_of_nat_not_eq_omega";
-Addsimps [hypnat_of_nat_not_eq_omega RS not_sym];
-
-(*-----------------------------------------------------------
-    Properties of the set SHNat of embedded natural numbers
-              (cf. set SReal in NSA.thy/NSA.ML)
- ----------------------------------------------------------*)
-
-(* Infinite hypernatural not in embedded SHNat *)
-Goalw [SHNat_def] "whn ~: SHNat";
-by (Auto_tac);
-qed "SHNAT_omega_not_mem";
-Addsimps [SHNAT_omega_not_mem];
-
-(*-----------------------------------------------------------------------
-     Closure laws for members of (embedded) set standard naturals SHNat
- -----------------------------------------------------------------------*)
-Goalw [SHNat_def] 
-      "!!x. [| x: SHNat; y: SHNat |] ==> x + y: SHNat";
-by (Step_tac 1);
-by (res_inst_tac [("x","N + Na")] exI 1);
-by (simp_tac (simpset() addsimps [hypnat_of_nat_add]) 1);
-qed "SHNat_add";
-
-Goalw [SHNat_def] 
-      "!!x. [| x: SHNat; y: SHNat |] ==> x - y: SHNat";
-by (Step_tac 1);
-by (res_inst_tac [("x","N - Na")] exI 1);
-by (simp_tac (simpset() addsimps [hypnat_of_nat_minus]) 1);
-qed "SHNat_minus";
-
-Goalw [SHNat_def] 
-      "!!x. [| x: SHNat; y: SHNat |] ==> x * y: SHNat";
-by (Step_tac 1);
-by (res_inst_tac [("x","N * Na")] exI 1);
-by (simp_tac (simpset() addsimps [hypnat_of_nat_mult]) 1);
-qed "SHNat_mult";
-
-Goal "!!x. [| x + y : SHNat; y: SHNat |] ==> x: SHNat";
-by (dres_inst_tac [("x","x+y")] SHNat_minus 1);
-by (Auto_tac);
-qed "SHNat_add_cancel";
-
-Goalw [SHNat_def] "hypnat_of_nat x : SHNat";
-by (Blast_tac 1);
-qed "SHNat_hypnat_of_nat";
-Addsimps [SHNat_hypnat_of_nat];
-
-Goal "hypnat_of_nat 1 : SHNat";
-by (Simp_tac 1);
-qed "SHNat_hypnat_of_nat_one";
-
-Goal "hypnat_of_nat 0 : SHNat";
-by (Simp_tac 1);
-qed "SHNat_hypnat_of_nat_zero";
-
-Goal "1hn : SHNat";
-by (simp_tac (simpset() addsimps [SHNat_hypnat_of_nat_one,
-    hypnat_of_nat_one RS sym]) 1);
-qed "SHNat_one";
-
-Goal "0 : SHNat";
-by (simp_tac (simpset() addsimps [SHNat_hypnat_of_nat_zero,
-    hypnat_of_nat_zero RS sym]) 1);
-qed "SHNat_zero";
-
-Addsimps [SHNat_hypnat_of_nat_one,SHNat_hypnat_of_nat_zero,
-          SHNat_one,SHNat_zero];
-
-Goal "1hn + 1hn : SHNat";
-by (rtac ([SHNat_one,SHNat_one] MRS SHNat_add) 1);
-qed "SHNat_two";
-
-Goalw [SHNat_def] "{x. hypnat_of_nat x : SHNat} = (UNIV::nat set)";
-by (Auto_tac);
-qed "SHNat_UNIV_nat";
-
-Goalw [SHNat_def] "(x: SHNat) = (EX y. x = hypnat_of_nat  y)";
-by (Auto_tac);
-qed "SHNat_iff";
-
-Goalw [SHNat_def] "hypnat_of_nat ``(UNIV::nat set) = SHNat";
-by (Auto_tac);
-qed "hypnat_of_nat_image";
-
-Goalw [SHNat_def] "inv hypnat_of_nat ``SHNat = (UNIV::nat set)";
-by (Auto_tac);
-by (rtac (inj_hypnat_of_nat RS inv_f_f RS subst) 1);
-by (Blast_tac 1);
-qed "inv_hypnat_of_nat_image";
-
-Goalw [SHNat_def] 
-      "!!P. [| EX x. x: P; P <= SHNat |] ==> \
-\           EX Q. P = hypnat_of_nat `` Q";
-by (Best_tac 1); 
-qed "SHNat_hypnat_of_nat_image";
-
-Goalw [SHNat_def] 
-      "SHNat = hypnat_of_nat `` (UNIV::nat set)";
-by (Auto_tac);
-qed "SHNat_hypnat_of_nat_iff";
-
-Goalw [SHNat_def] "SHNat <= (UNIV::hypnat set)";
-by (Auto_tac);
-qed "SHNat_subset_UNIV";
-
-Goal "{n. n <= Suc m} = {n. n <= m} Un {n. n = Suc m}";
-by (auto_tac (claset(),simpset() addsimps [le_Suc_eq]));
-qed "leSuc_Un_eq";
-
-Goal "finite {n::nat. n <= m}";
-by (nat_ind_tac "m" 1);
-by (auto_tac (claset(),simpset() addsimps [leSuc_Un_eq]));
-qed "finite_nat_le_segment";
-
-Goal "{n::nat. m < n} : FreeUltrafilterNat";
-by (cut_inst_tac [("m2","m")] (finite_nat_le_segment RS 
-    FreeUltrafilterNat_finite RS FreeUltrafilterNat_Compl_mem) 1);
-by (Fuf_tac 1);
-qed "lemma_unbounded_set";
-Addsimps [lemma_unbounded_set];
-
-Goalw [SHNat_def,hypnat_of_nat_def,
-           hypnat_less_def,hypnat_omega_def] 
-           "ALL n: SHNat. n < whn";
-by (Clarify_tac 1);
-by (auto_tac (claset() addSIs [exI],simpset()));
-qed "hypnat_omega_gt_SHNat";
-
-Goal "hypnat_of_nat n < whn";
-by (cut_facts_tac [hypnat_omega_gt_SHNat] 1);
-by (dres_inst_tac [("x","hypnat_of_nat n")] bspec 1);
-by (Auto_tac);
-qed "hypnat_of_nat_less_whn";
-Addsimps [hypnat_of_nat_less_whn];
-
-Goal "hypnat_of_nat n <= whn";
-by (rtac (hypnat_of_nat_less_whn RS hypnat_less_imp_le) 1);
-qed "hypnat_of_nat_le_whn";
-Addsimps [hypnat_of_nat_le_whn];
-
-Goal "0 < whn";
-by (rtac (hypnat_omega_gt_SHNat RS ballE) 1);
-by (Auto_tac);
-qed "hypnat_zero_less_hypnat_omega";
-Addsimps [hypnat_zero_less_hypnat_omega];
-
-Goal "1hn < whn";
-by (rtac (hypnat_omega_gt_SHNat RS ballE) 1);
-by (Auto_tac);
-qed "hypnat_one_less_hypnat_omega";
-Addsimps [hypnat_one_less_hypnat_omega];
-
-(*--------------------------------------------------------------------------
-     Theorems about infinite hypernatural numbers -- HNatInfinite
- -------------------------------------------------------------------------*)
-Goalw [HNatInfinite_def,SHNat_def] "whn : HNatInfinite";
-by (Auto_tac);
-qed "HNatInfinite_whn";
-Addsimps [HNatInfinite_whn];
-
-Goalw [HNatInfinite_def] "!!x. x: SHNat ==> x ~: HNatInfinite";
-by (Simp_tac 1);
-qed "SHNat_not_HNatInfinite";
-
-Goalw [HNatInfinite_def] "!!x. x ~: HNatInfinite ==> x: SHNat";
-by (Asm_full_simp_tac 1);
-qed "not_HNatInfinite_SHNat";
-
-Goalw [HNatInfinite_def] "!!x. x ~: SHNat ==> x: HNatInfinite";
-by (Simp_tac 1);
-qed "not_SHNat_HNatInfinite";
-
-Goalw [HNatInfinite_def] "!!x. x: HNatInfinite ==> x ~: SHNat";
-by (Asm_full_simp_tac 1);
-qed "HNatInfinite_not_SHNat";
-
-Goal "(x: SHNat) = (x ~: HNatInfinite)";
-by (blast_tac (claset() addSIs [SHNat_not_HNatInfinite,
-    not_HNatInfinite_SHNat]) 1);
-qed "SHNat_not_HNatInfinite_iff";
-
-Goal "(x ~: SHNat) = (x: HNatInfinite)";
-by (blast_tac (claset() addSIs [not_SHNat_HNatInfinite,
-    HNatInfinite_not_SHNat]) 1);
-qed "not_SHNat_HNatInfinite_iff";
-
-Goal "x : SHNat | x : HNatInfinite";
-by (simp_tac (simpset() addsimps [SHNat_not_HNatInfinite_iff]) 1);
-qed "SHNat_HNatInfinite_disj";
-
-(*-------------------------------------------------------------------
-   Proof of alternative definition for set of Infinite hypernatural 
-   numbers --- HNatInfinite = {N. ALL n: SHNat. n < N}
- -------------------------------------------------------------------*)
-Goal "!!N (xa::nat=>nat). \
-\         (ALL N. {n. xa n ~= N} : FreeUltrafilterNat) ==> \
-\         ({n. N < xa n} : FreeUltrafilterNat)";
-by (nat_ind_tac "N" 1);
-by (dres_inst_tac [("x","0")] spec 1);
-by (rtac ccontr 1 THEN dtac FreeUltrafilterNat_Compl_mem 1
-    THEN dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
-by (Asm_full_simp_tac 1);
-by (dres_inst_tac [("x","Suc N")] spec 1);
-by (fuf_tac (claset() addSDs [Suc_leI],simpset() addsimps 
-    [le_eq_less_or_eq]) 1);
-qed "HNatInfinite_FreeUltrafilterNat_lemma";
-
-(*** alternative definition ***)
-Goalw [HNatInfinite_def,SHNat_def,hypnat_of_nat_def] 
-      "HNatInfinite = {N. ALL n:SHNat. n < N}";
-by (Step_tac 1);
-by (dres_inst_tac [("x","Abs_hypnat \
-\        (hypnatrel ^^ {%n. N})")] bspec 2);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),simpset() addsimps [hypnat_less_iff]));
-by (auto_tac (claset() addSIs [exI] addEs 
-    [HNatInfinite_FreeUltrafilterNat_lemma],
-    simpset() addsimps [FreeUltrafilterNat_Compl_iff1, 
-     CLAIM "- {n. xa n = N} = {n. xa n ~= N}"]));
-qed "HNatInfinite_iff";
-
-(*--------------------------------------------------------------------
-   Alternative definition for HNatInfinite using Free ultrafilter
- --------------------------------------------------------------------*)
-Goal "!!x. x : HNatInfinite ==> EX X: Rep_hypnat x. \
-\          ALL u. {n. u < X n}:  FreeUltrafilterNat";
-by (auto_tac (claset(),simpset() addsimps [hypnat_less_def,
-    HNatInfinite_iff,SHNat_iff,hypnat_of_nat_def]));
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (EVERY[Auto_tac, rtac bexI 1, 
-    rtac lemma_hypnatrel_refl 2, Step_tac 1]);
-by (dres_inst_tac [("x","hypnat_of_nat u")] bspec 1);
-by (Simp_tac 1);
-by (auto_tac (claset(),
-    simpset() addsimps [hypnat_of_nat_def]));
-by (Fuf_tac 1);
-qed "HNatInfinite_FreeUltrafilterNat";
-
-Goal "!!x. EX X: Rep_hypnat x. \
-\          ALL u. {n. u < X n}:  FreeUltrafilterNat \
-\          ==> x: HNatInfinite";
-by (auto_tac (claset(),simpset() addsimps [hypnat_less_def,
-    HNatInfinite_iff,SHNat_iff,hypnat_of_nat_def]));
-by (rtac exI 1 THEN Auto_tac);
-qed "FreeUltrafilterNat_HNatInfinite";
-
-Goal "!!x. (x : HNatInfinite) = (EX X: Rep_hypnat x. \
-\          ALL u. {n. u < X n}:  FreeUltrafilterNat)";
-by (blast_tac (claset() addIs [HNatInfinite_FreeUltrafilterNat,
-    FreeUltrafilterNat_HNatInfinite]) 1);
-qed "HNatInfinite_FreeUltrafilterNat_iff";
-
-Goal "!!x. x : HNatInfinite ==> 1hn < x";
-by (auto_tac (claset(),simpset() addsimps [HNatInfinite_iff]));
-qed "HNatInfinite_gt_one";
-Addsimps [HNatInfinite_gt_one];
-
-Goal "!!x. 0 ~: HNatInfinite";
-by (auto_tac (claset(),simpset() 
-    addsimps [HNatInfinite_iff]));
-by (dres_inst_tac [("a","1hn")] equals0D 1);
-by (Asm_full_simp_tac 1);
-qed "zero_not_mem_HNatInfinite";
-Addsimps [zero_not_mem_HNatInfinite];
-
-Goal "!!x. x : HNatInfinite ==> x ~= 0";
-by (Auto_tac);
-qed "HNatInfinite_not_eq_zero";
-
-Goal "!!x. x : HNatInfinite ==> 1hn <= x";
-by (blast_tac (claset() addIs [hypnat_less_imp_le,
-         HNatInfinite_gt_one]) 1);
-qed "HNatInfinite_ge_one";
-Addsimps [HNatInfinite_ge_one];
-
-(*--------------------------------------------------
-                   Closure Rules
- --------------------------------------------------*)
-Goal "!!x. [| x: HNatInfinite; y: HNatInfinite |] \
-\           ==> x + y: HNatInfinite";
-by (auto_tac (claset(),simpset() addsimps [HNatInfinite_iff]));
-by (dtac bspec 1 THEN assume_tac 1);
-by (dtac (SHNat_zero RSN (2,bspec)) 1);
-by (dtac hypnat_add_less_mono 1 THEN assume_tac 1);
-by (Asm_full_simp_tac 1);
-qed "HNatInfinite_add";
-
-Goal "!!x. [| x: HNatInfinite; y: SHNat |] \
-\                       ==> x + y: HNatInfinite";
-by (rtac ccontr 1 THEN dtac not_HNatInfinite_SHNat 1);
-by (dres_inst_tac [("x","x + y")] SHNat_minus 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [SHNat_not_HNatInfinite_iff]));
-qed "HNatInfinite_SHNat_add";
-
-goal HyperNat.thy "!!x. [| x: HNatInfinite; y: SHNat |] \
-\                       ==> x - y: HNatInfinite";
-by (rtac ccontr 1 THEN dtac not_HNatInfinite_SHNat 1);
-by (dres_inst_tac [("x","x - y")] SHNat_add 1);
-by (subgoal_tac "y <= x" 2);
-by (auto_tac (claset() addSDs [hypnat_le_add_diff_inverse2],
-    simpset() addsimps [not_SHNat_HNatInfinite_iff RS sym]));
-by (auto_tac (claset() addSIs [hypnat_less_imp_le],
-    simpset() addsimps [not_SHNat_HNatInfinite_iff,HNatInfinite_iff]));
-qed "HNatInfinite_SHNat_diff";
-
-Goal 
-     "!!x. x: HNatInfinite ==> x + 1hn: HNatInfinite";
-by (auto_tac (claset() addIs [HNatInfinite_SHNat_add],
-              simpset()));
-qed "HNatInfinite_add_one";
-
-Goal 
-     "!!x. x: HNatInfinite ==> x - 1hn: HNatInfinite";
-by (rtac ccontr 1 THEN dtac not_HNatInfinite_SHNat 1);
-by (dres_inst_tac [("x","x - 1hn"),("y","1hn")] SHNat_add 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [not_SHNat_HNatInfinite_iff RS sym]));
-qed "HNatInfinite_minus_one";
-
-Goal "!!x. x : HNatInfinite ==> EX y. x = y + 1hn";
-by (res_inst_tac [("x","x - 1hn")] exI 1);
-by (Auto_tac);
-qed "HNatInfinite_is_Suc";
-
-(*---------------------------------------------------------------
-       HNat : the hypernaturals embedded in the hyperreals
-       Obtained using the NS extension of the naturals
- --------------------------------------------------------------*)
- 
-Goalw [HNat_def,starset_def,
-         hypreal_of_posnat_def,hypreal_of_real_def] 
-         "hypreal_of_posnat N : HNat";
-by (Auto_tac);
-by (Ultra_tac 1);
-by (res_inst_tac [("x","Suc N")] exI 1);
-by (dtac sym 1 THEN auto_tac (claset(),simpset() 
-    addsimps [real_of_preal_real_of_nat_Suc]));
-qed "HNat_hypreal_of_posnat";
-Addsimps [HNat_hypreal_of_posnat];
-
-Goalw [HNat_def,starset_def]
-     "[| x: HNat; y: HNat |] ==> x + y: HNat";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSDs [bspec] addIs [lemma_hyprel_refl],
-    simpset() addsimps [hypreal_add]));
-by (Ultra_tac 1);
-by (dres_inst_tac [("t","Y xb")] sym 1);
-by (auto_tac (claset(),simpset() addsimps [real_of_nat_add RS sym]));
-qed "HNat_add";
-
-Goalw [HNat_def,starset_def]
-     "[| x: HNat; y: HNat |] ==> x * y: HNat";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSDs [bspec] addIs [lemma_hyprel_refl],
-    simpset() addsimps [hypreal_mult]));
-by (Ultra_tac 1);
-by (dres_inst_tac [("t","Y xb")] sym 1);
-by (auto_tac (claset(),simpset() addsimps [real_of_nat_mult RS sym]));
-qed "HNat_mult";
-
-(*---------------------------------------------------------------
-    Embedding of the hypernaturals into the hyperreal
- --------------------------------------------------------------*)
-(*** A lemma that should have been derived a long time ago! ***)
-Goal "(Ya : hyprel ^^{%n. f(n)}) = \
-\         ({n. f n = Ya n} : FreeUltrafilterNat)";
-by (Auto_tac);
-qed "lemma_hyprel_FUFN";
-
-Goalw [hypreal_of_hypnat_def]
-      "hypreal_of_hypnat (Abs_hypnat(hypnatrel^^{%n. X n})) = \
-\         Abs_hypreal(hyprel ^^ {%n. real_of_nat (X n)})";
-by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
-by (auto_tac (claset() addEs [FreeUltrafilterNat_Int RS
-    FreeUltrafilterNat_subset],simpset() addsimps 
-    [lemma_hyprel_FUFN]));
-qed "hypreal_of_hypnat";
-
-Goal "inj(hypreal_of_hypnat)";
-by (rtac injI 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [hypreal_of_hypnat,real_of_nat_eq_cancel]));
-qed "inj_hypreal_of_hypnat";
-
-Goal "(hypreal_of_hypnat n = hypreal_of_hypnat m) = (n = m)";
-by (auto_tac (claset(),simpset() addsimps [inj_hypreal_of_hypnat RS injD]));
-qed "hypreal_of_hypnat_eq_cancel";
-Addsimps [hypreal_of_hypnat_eq_cancel];
-
-Goal "(hypnat_of_nat n = hypnat_of_nat m) = (n = m)";
-by (auto_tac (claset() addDs [inj_hypnat_of_nat RS injD],
-              simpset()));
-qed "hypnat_of_nat_eq_cancel";
-Addsimps [hypnat_of_nat_eq_cancel];
-
-Goalw [hypreal_zero_def,hypnat_zero_def] 
-           "hypreal_of_hypnat  0 = 0";
-by (simp_tac (simpset() addsimps [hypreal_of_hypnat,
-    real_of_nat_zero]) 1);
-qed "hypreal_of_hypnat_zero";
-
-Goalw [hypreal_one_def,hypnat_one_def] 
-           "hypreal_of_hypnat  1hn = 1hr";
-by (simp_tac (simpset() addsimps [hypreal_of_hypnat,
-    real_of_nat_one]) 1);
-qed "hypreal_of_hypnat_one";
-
-Goal "hypreal_of_hypnat n1 + hypreal_of_hypnat n2 = hypreal_of_hypnat (n1 + n2)";
-by (res_inst_tac [("z","n1")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","n2")] eq_Abs_hypnat 1);
-by (asm_simp_tac (simpset() addsimps [hypreal_of_hypnat,
-        hypreal_add,hypnat_add,real_of_nat_add]) 1);
-qed "hypreal_of_hypnat_add";
-
-Goal "hypreal_of_hypnat n1 * hypreal_of_hypnat n2 = hypreal_of_hypnat (n1 * n2)";
-by (res_inst_tac [("z","n1")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","n2")] eq_Abs_hypnat 1);
-by (asm_simp_tac (simpset() addsimps [hypreal_of_hypnat,
-        hypreal_mult,hypnat_mult,real_of_nat_mult]) 1);
-qed "hypreal_of_hypnat_mult";
-
-Goal "(hypreal_of_hypnat n < hypreal_of_hypnat m) = (n < m)";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
-by (asm_simp_tac (simpset() addsimps 
-    [hypreal_of_hypnat,hypreal_less,hypnat_less]) 1);
-qed "hypreal_of_hypnat_less_iff";
-Addsimps [hypreal_of_hypnat_less_iff];
-
-Goal "(hypreal_of_hypnat N = 0) = (N = 0)";
-by (simp_tac (simpset() addsimps [hypreal_of_hypnat_zero RS sym]) 1);
-qed "hypreal_of_hypnat_eq_zero_iff";
-Addsimps [hypreal_of_hypnat_eq_zero_iff];
-
-Goal "ALL n. N <= n ==> N = (0::hypnat)";
-by (dres_inst_tac [("x","0")] spec 1);
-by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),simpset() addsimps [hypnat_le,hypnat_zero_def]));
-qed "hypnat_eq_zero";
-Addsimps [hypnat_eq_zero];
-
-Goal "~ (ALL n. n = (0::hypnat))";
-by Auto_tac;
-by (res_inst_tac [("x","1hn")] exI 1);
-by (Simp_tac 1);
-qed "hypnat_not_all_eq_zero";
-Addsimps [hypnat_not_all_eq_zero];
-
-Goal "n ~= 0 ==> (n <= 1hn) = (n = 1hn)";
-by (auto_tac (claset(),simpset() addsimps [hypnat_le_eq_less_or_eq]));
-qed "hypnat_le_one_eq_one";
-Addsimps [hypnat_le_one_eq_one];
-
-
-
-
--- a/src/HOL/Real/Hyperreal/HyperNat.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,84 +0,0 @@
-(*  Title       : HyperNat.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Explicit construction of hypernaturals using 
-                  ultrafilters
-*) 
-
-HyperNat = Star + 
-
-constdefs
-    hypnatrel :: "((nat=>nat)*(nat=>nat)) set"
-    "hypnatrel == {p. ? X Y. p = ((X::nat=>nat),Y) & 
-                   {n::nat. X(n) = Y(n)}: FreeUltrafilterNat}"
-
-typedef hypnat = "{x::nat=>nat. True}//hypnatrel"              (Equiv.quotient_def)
-
-instance
-   hypnat  :: {ord,zero,plus,times,minus}
-
-consts
-  "1hn"       :: hypnat               ("1hn")  
-  "whn"       :: hypnat               ("whn")  
-
-defs
-
-  hypnat_zero_def      "0 == Abs_hypnat(hypnatrel^^{%n::nat. 0})"
-  hypnat_one_def       "1hn == Abs_hypnat(hypnatrel^^{%n::nat. 1})"
-
-  (* omega is in fact an infinite hypernatural number = [<1,2,3,...>] *)
-  hypnat_omega_def     "whn == Abs_hypnat(hypnatrel^^{%n::nat. n})"
- 
-
-constdefs
-
-  (* embedding the naturals in the hypernaturals *)
-  hypnat_of_nat   :: nat => hypnat		  ("**# _" [80] 80)	 
-  "hypnat_of_nat m  == Abs_hypnat(hypnatrel^^{%n::nat. m})"
-
-  (* set of naturals embedded in the hypernaturals*)
-  SHNat         :: "hypnat set"
-  "SHNat        == {n. EX N. n = hypnat_of_nat N}"  
- 
-  (* embedding the naturals in the hyperreals*)
-  SNat         :: "hypreal set"
-  "SNat        == {n. EX N. n = hypreal_of_nat N}"  
-
-  (* hypernaturals as members of the hyperreals; the set is defined as  *)
-  (* the nonstandard extension of set of naturals embedded in the reals *)
-  HNat         :: "hypreal set"
-  "HNat         == *s* {n. EX no. n = real_of_nat no}"
-
-  (* the set of infinite hypernatural numbers *)
-  HNatInfinite :: "hypnat set"
-  "HNatInfinite == {n. n ~: SHNat}"
-
-  (* explicit embedding of the hypernaturals in the hyperreals *)    
-  hypreal_of_hypnat :: hypnat => hypreal    ("&H# _" [80] 80)
-  "hypreal_of_hypnat N  == Abs_hypreal(UN X: Rep_hypnat(N). 
-                            hyprel^^{%n::nat. real_of_nat (X n)})"
-  
-defs
-  hypnat_add_def  
-  "P + Q == Abs_hypnat(UN X:Rep_hypnat(P). UN Y:Rep_hypnat(Q).
-                hypnatrel^^{%n::nat. X n + Y n})"
-
-  hypnat_mult_def  
-  "P * Q == Abs_hypnat(UN X:Rep_hypnat(P). UN Y:Rep_hypnat(Q).
-                hypnatrel^^{%n::nat. X n * Y n})"
-
-  hypnat_minus_def  
-  "P - Q == Abs_hypnat(UN X:Rep_hypnat(P). UN Y:Rep_hypnat(Q).
-                hypnatrel^^{%n::nat. X n - Y n})"
-
-  hypnat_less_def
-  "P < (Q::hypnat) == EX X Y. X: Rep_hypnat(P) & 
-                               Y: Rep_hypnat(Q) & 
-                               {n::nat. X n < Y n} : FreeUltrafilterNat"
-  hypnat_le_def
-  "P <= (Q::hypnat) == ~(Q < P)" 
-
-end
-
-
-
--- a/src/HOL/Real/Hyperreal/HyperOrd.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,768 +0,0 @@
-(*  Title:	 Real/Hyperreal/HyperOrd.ML
-    Author:      Jacques D. Fleuriot
-    Copyright:   1998 University of Cambridge
-                 2000 University of Edinburgh
-    Description: Type "hypreal" is a linear order and also 
-                 satisfies plus_ac0: + is an AC-operator with zero
-*)
-
-(**** The simproc abel_cancel ****)
-
-(*** Two lemmas needed for the simprocs ***)
-
-(*Deletion of other terms in the formula, seeking the -x at the front of z*)
-Goal "((x::hypreal) + (y + z) = y + u) = ((x + z) = u)";
-by (stac hypreal_add_left_commute 1);
-by (rtac hypreal_add_left_cancel 1);
-qed "hypreal_add_cancel_21";
-
-(*A further rule to deal with the case that
-  everything gets cancelled on the right.*)
-Goal "((x::hypreal) + (y + z) = y) = (x = -z)";
-by (stac hypreal_add_left_commute 1);
-by (res_inst_tac [("t", "y")] (hypreal_add_zero_right RS subst) 1
-    THEN stac hypreal_add_left_cancel 1);
-by (simp_tac (simpset() addsimps [hypreal_eq_diff_eq RS sym]) 1);
-qed "hypreal_add_cancel_end";
-
-
-structure Hyperreal_Cancel_Data =
-struct
-  val ss		= HOL_ss
-  val eq_reflection	= eq_reflection
-
-  val sg_ref		= Sign.self_ref (Theory.sign_of (the_context ()))
-  val T			= Type("HyperDef.hypreal",[])
-  val zero		= Const ("0", T)
-  val restrict_to_left  = restrict_to_left
-  val add_cancel_21	= hypreal_add_cancel_21
-  val add_cancel_end	= hypreal_add_cancel_end
-  val add_left_cancel	= hypreal_add_left_cancel
-  val add_assoc		= hypreal_add_assoc
-  val add_commute	= hypreal_add_commute
-  val add_left_commute	= hypreal_add_left_commute
-  val add_0		= hypreal_add_zero_left
-  val add_0_right	= hypreal_add_zero_right
-
-  val eq_diff_eq	= hypreal_eq_diff_eq
-  val eqI_rules		= [hypreal_less_eqI, hypreal_eq_eqI, hypreal_le_eqI]
-  fun dest_eqI th = 
-      #1 (HOLogic.dest_bin "op =" HOLogic.boolT 
-	      (HOLogic.dest_Trueprop (concl_of th)))
-
-  val diff_def		= hypreal_diff_def
-  val minus_add_distrib	= hypreal_minus_add_distrib
-  val minus_minus	= hypreal_minus_minus
-  val minus_0		= hypreal_minus_zero
-  val add_inverses	= [hypreal_add_minus, hypreal_add_minus_left]
-  val cancel_simps	= [hypreal_add_minus_cancelA, hypreal_minus_add_cancelA]
-end;
-
-structure Hyperreal_Cancel = Abel_Cancel (Hyperreal_Cancel_Data);
-
-Addsimprocs [Hyperreal_Cancel.sum_conv, Hyperreal_Cancel.rel_conv];
-
-Goal "- (z - y) = y - (z::hypreal)";
-by (Simp_tac 1);
-qed "hypreal_minus_diff_eq";
-Addsimps [hypreal_minus_diff_eq];
-
-
-Goal "((x::hypreal) < y) = (-y < -x)";
-by (stac hypreal_less_minus_iff 1);
-by (res_inst_tac [("x1","x")] (hypreal_less_minus_iff RS ssubst) 1);
-by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
-qed "hypreal_less_swap_iff";
-
-Goalw [hypreal_zero_def] 
-      "((0::hypreal) < x) = (-x < x)";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(), simpset() addsimps [hypreal_less, hypreal_minus]));
-by (ALLGOALS(Ultra_tac));
-qed "hypreal_gt_zero_iff";
-
-Goal "(A::hypreal) < B ==> A + C < B + C";
-by (res_inst_tac [("z","A")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","B")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","C")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSIs [exI], 
-              simpset() addsimps [hypreal_less_def,hypreal_add]));
-by (Ultra_tac 1);
-qed "hypreal_add_less_mono1";
-
-Goal "!!(A::hypreal). A < B ==> C + A < C + B";
-by (auto_tac (claset() addIs [hypreal_add_less_mono1],
-    simpset() addsimps [hypreal_add_commute]));
-qed "hypreal_add_less_mono2";
-
-Goal "(x < (0::hypreal)) = (x < -x)";
-by (rtac (hypreal_minus_zero_less_iff RS subst) 1);
-by (stac hypreal_gt_zero_iff 1);
-by (Full_simp_tac 1);
-qed "hypreal_lt_zero_iff";
-
-Goalw [hypreal_le_def] "((0::hypreal) <= x) = (-x <= x)";
-by (auto_tac (claset(),simpset() addsimps [hypreal_lt_zero_iff RS sym]));
-qed "hypreal_ge_zero_iff";
-
-Goalw [hypreal_le_def] "(x <= (0::hypreal)) = (x <= -x)";
-by (auto_tac (claset(),simpset() addsimps [hypreal_gt_zero_iff RS sym]));
-qed "hypreal_le_zero_iff";
-
-Goalw [hypreal_zero_def] 
-      "[| 0 < x; 0 < y |] ==> (0::hypreal) < x + y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps
-    [hypreal_less_def,hypreal_add]));
-by (auto_tac (claset() addSIs [exI],simpset() addsimps
-    [hypreal_less_def,hypreal_add]));
-by (ultra_tac (claset() addIs [real_add_order],simpset()) 1);
-qed "hypreal_add_order";
-
-Goal "[| 0 < x; 0 <= y |] ==> (0::hypreal) < x + y";
-by (auto_tac (claset() addDs [sym,hypreal_le_imp_less_or_eq]
-              addIs [hypreal_add_order],simpset()));
-qed "hypreal_add_order_le";            
-
-Goalw [hypreal_zero_def] 
-          "[| 0 < x; 0 < y |] ==> (0::hypreal) < x * y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSIs [exI],simpset() addsimps
-    [hypreal_less_def,hypreal_mult]));
-by (ultra_tac (claset() addIs [rename_numerals real_mult_order],
-	       simpset()) 1);
-qed "hypreal_mult_order";
-
-Goal "[| x < 0; y < 0 |] ==> (0::hypreal) < x * y";
-by (REPEAT(dtac (hypreal_minus_zero_less_iff RS iffD2) 1));
-by (dtac hypreal_mult_order 1 THEN assume_tac 1);
-by (Asm_full_simp_tac 1);
-qed "hypreal_mult_less_zero1";
-
-Goal "[| 0 <= x; 0 <= y |] ==> (0::hypreal) <= x * y";
-by (REPEAT(dtac hypreal_le_imp_less_or_eq 1));
-by (auto_tac (claset() addIs [hypreal_mult_order,
-    hypreal_less_imp_le],simpset()));
-qed "hypreal_le_mult_order";
-
-
-Goal "[| x <= 0; y <= 0 |] ==> (0::hypreal) <= x * y";
-by (rtac hypreal_less_or_eq_imp_le 1);
-by (dtac hypreal_le_imp_less_or_eq 1 THEN etac disjE 1);
-by Auto_tac;
-by (dtac hypreal_le_imp_less_or_eq 1);
-by (auto_tac (claset() addDs [hypreal_mult_less_zero1],simpset()));
-qed "hypreal_mult_le_zero1";
-
-Goal "[| 0 <= x; y < 0 |] ==> x * y <= (0::hypreal)";
-by (rtac hypreal_less_or_eq_imp_le 1);
-by (dtac hypreal_le_imp_less_or_eq 1 THEN etac disjE 1);
-by Auto_tac;
-by (dtac (hypreal_minus_zero_less_iff RS iffD2) 1);
-by (rtac (hypreal_minus_zero_less_iff RS subst) 1);
-by (blast_tac (claset() addDs [hypreal_mult_order] 
-    addIs [hypreal_minus_mult_eq2 RS ssubst]) 1);
-qed "hypreal_mult_le_zero";
-
-Goal "[| 0 < x; y < 0 |] ==> x*y < (0::hypreal)";
-by (dtac (hypreal_minus_zero_less_iff RS iffD2) 1);
-by (dtac hypreal_mult_order 1 THEN assume_tac 1);
-by (rtac (hypreal_minus_zero_less_iff RS iffD1) 1);
-by (Asm_full_simp_tac 1);
-qed "hypreal_mult_less_zero";
-
-Goalw [hypreal_one_def,hypreal_zero_def,hypreal_less_def] "0 < 1hr";
-by (res_inst_tac [("x","%n. #0")] exI 1);
-by (res_inst_tac [("x","%n. #1")] exI 1);
-by (auto_tac (claset(),
-        simpset() addsimps [real_zero_less_one, FreeUltrafilterNat_Nat_set]));
-qed "hypreal_zero_less_one";
-
-Goal "1hr < 1hr + 1hr";
-by (rtac (hypreal_less_minus_iff RS iffD2) 1);
-by (full_simp_tac (simpset() addsimps [hypreal_zero_less_one,
-    hypreal_add_assoc]) 1);
-qed "hypreal_one_less_two";
-
-Goal "0 < 1hr + 1hr";
-by (rtac ([hypreal_zero_less_one,
-          hypreal_one_less_two] MRS hypreal_less_trans) 1);
-qed "hypreal_zero_less_two";
-
-Goal "1hr + 1hr ~= 0";
-by (rtac (hypreal_zero_less_two RS hypreal_not_refl2 RS not_sym) 1);
-qed "hypreal_two_not_zero";
-Addsimps [hypreal_two_not_zero];
-
-Goal "x/(1hr + 1hr) + x/(1hr + 1hr) = x";
-by (stac hypreal_add_self 1);
-by (full_simp_tac (simpset() addsimps [hypreal_mult_assoc, hypreal_divide_def,
-    hypreal_two_not_zero RS hypreal_mult_inverse_left]) 1);
-qed "hypreal_sum_of_halves";
-
-Goal "[| 0 <= x; 0 <= y |] ==> (0::hypreal) <= x + y";
-by (REPEAT(dtac hypreal_le_imp_less_or_eq 1));
-by (auto_tac (claset() addIs [hypreal_add_order,
-    hypreal_less_imp_le],simpset()));
-qed "hypreal_le_add_order";
-
-(*** Monotonicity results ***)
-
-Goal "(v+z < w+z) = (v < (w::hypreal))";
-by (Simp_tac 1);
-qed "hypreal_add_right_cancel_less";
-
-Goal "(z+v < z+w) = (v < (w::hypreal))";
-by (Simp_tac 1);
-qed "hypreal_add_left_cancel_less";
-
-Addsimps [hypreal_add_right_cancel_less, 
-          hypreal_add_left_cancel_less];
-
-Goal "(v+z <= w+z) = (v <= (w::hypreal))";
-by (Simp_tac 1);
-qed "hypreal_add_right_cancel_le";
-
-Goal "(z+v <= z+w) = (v <= (w::hypreal))";
-by (Simp_tac 1);
-qed "hypreal_add_left_cancel_le";
-
-Addsimps [hypreal_add_right_cancel_le, hypreal_add_left_cancel_le];
-
-Goal  "[| (z1::hypreal) < y1; z2 < y2 |] ==> z1 + z2 < y1 + y2";
-by (dtac (hypreal_less_minus_iff RS iffD1) 1);
-by (dtac (hypreal_less_minus_iff RS iffD1) 1);
-by (dtac hypreal_add_order 1 THEN assume_tac 1);
-by (thin_tac "0 < y2 + - z2" 1);
-by (dres_inst_tac [("C","z1 + z2")] hypreal_add_less_mono1 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [hypreal_minus_add_distrib RS sym] @ hypreal_add_ac
-    delsimps [hypreal_minus_add_distrib]));
-qed "hypreal_add_less_mono";
-
-Goal "(q1::hypreal) <= q2  ==> x + q1 <= x + q2";
-by (dtac hypreal_le_imp_less_or_eq 1);
-by (Step_tac 1);
-by (auto_tac (claset() addSIs [hypreal_le_refl,
-    hypreal_less_imp_le,hypreal_add_less_mono1],
-    simpset() addsimps [hypreal_add_commute]));
-qed "hypreal_add_left_le_mono1";
-
-Goal "(q1::hypreal) <= q2  ==> q1 + x <= q2 + x";
-by (auto_tac (claset() addDs [hypreal_add_left_le_mono1],
-    simpset() addsimps [hypreal_add_commute]));
-qed "hypreal_add_le_mono1";
-
-Goal "[|(i::hypreal)<=j;  k<=l |] ==> i + k <= j + l";
-by (etac (hypreal_add_le_mono1 RS hypreal_le_trans) 1);
-by (Simp_tac 1);
-qed "hypreal_add_le_mono";
-
-Goal "[|(i::hypreal)<j;  k<=l |] ==> i + k < j + l";
-by (auto_tac (claset() addSDs [hypreal_le_imp_less_or_eq] 
-    addIs [hypreal_add_less_mono1,hypreal_add_less_mono],
-    simpset()));
-qed "hypreal_add_less_le_mono";
-
-Goal "[|(i::hypreal)<=j;  k<l |] ==> i + k < j + l";
-by (auto_tac (claset() addSDs [hypreal_le_imp_less_or_eq] 
-    addIs [hypreal_add_less_mono2,hypreal_add_less_mono],simpset()));
-qed "hypreal_add_le_less_mono";
-
-Goal "(A::hypreal) + C < B + C ==> A < B";
-by (Full_simp_tac 1);
-qed "hypreal_less_add_right_cancel";
-
-Goal "(C::hypreal) + A < C + B ==> A < B";
-by (Full_simp_tac 1);
-qed "hypreal_less_add_left_cancel";
-
-Goal "[|r < x; (0::hypreal) <= y|] ==> r < x + y";
-by (auto_tac (claset() addDs [hypreal_add_less_le_mono],
-    simpset()));
-qed "hypreal_add_zero_less_le_mono";
-
-Goal "!!(A::hypreal). A + C <= B + C ==> A <= B";
-by (dres_inst_tac [("x","-C")] hypreal_add_le_mono1 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc]) 1);
-qed "hypreal_le_add_right_cancel";
-
-Goal "!!(A::hypreal). C + A <= C + B ==> A <= B";
-by (dres_inst_tac [("x","-C")] hypreal_add_left_le_mono1 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
-qed "hypreal_le_add_left_cancel";
-
-Goal "(0::hypreal) <= x*x";
-by (res_inst_tac [("x","0"),("y","x")] hypreal_linear_less2 1);
-by (auto_tac (claset() addIs [hypreal_mult_order,
-    hypreal_mult_less_zero1,hypreal_less_imp_le],
-    simpset()));
-qed "hypreal_le_square";
-Addsimps [hypreal_le_square];
-
-Goalw [hypreal_le_def] "- (x*x) <= (0::hypreal)";
-by (auto_tac (claset() addSDs [(hypreal_le_square RS 
-    hypreal_le_less_trans)],simpset() addsimps 
-    [hypreal_minus_zero_less_iff,hypreal_less_not_refl]));
-qed "hypreal_less_minus_square";
-Addsimps [hypreal_less_minus_square];
-
-Goal "(0*x<r)=((0::hypreal)<r)";
-by (Simp_tac 1);
-qed "hypreal_mult_0_less";
-
-Goal "[| (0::hypreal) < z; x < y |] ==> x*z < y*z";       
-by (rotate_tac 1 1);
-by (dtac (hypreal_less_minus_iff RS iffD1) 1);
-by (rtac (hypreal_less_minus_iff RS iffD2) 1);
-by (dtac hypreal_mult_order 1 THEN assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_mult_distrib2,
-					   hypreal_mult_commute ]) 1);
-qed "hypreal_mult_less_mono1";
-
-Goal "[| (0::hypreal)<z; x<y |] ==> z*x<z*y";       
-by (asm_simp_tac (simpset() addsimps [hypreal_mult_commute,hypreal_mult_less_mono1]) 1);
-qed "hypreal_mult_less_mono2";
-
-Goal "[| (0::hypreal)<=z; x<y |] ==> x*z<=y*z";
-by (EVERY1 [rtac hypreal_less_or_eq_imp_le, dtac hypreal_le_imp_less_or_eq]);
-by (auto_tac (claset() addIs [hypreal_mult_less_mono1],simpset()));
-qed "hypreal_mult_le_less_mono1";
-
-Goal "[| (0::hypreal)<=z; x<y |] ==> z*x<=z*y";
-by (asm_simp_tac (simpset() addsimps [hypreal_mult_commute,
-				      hypreal_mult_le_less_mono1]) 1);
-qed "hypreal_mult_le_less_mono2";
-
-Goal "[| (0::hypreal)<=z; x<=y |] ==> z*x<=z*y";
-by (dres_inst_tac [("x","x")] hypreal_le_imp_less_or_eq 1);
-by (auto_tac (claset() addIs [hypreal_mult_le_less_mono2,hypreal_le_refl],simpset()));
-qed "hypreal_mult_le_le_mono1";
-
-val prem1::prem2::prem3::rest = goal thy
-     "[| (0::hypreal)<y; x<r; y*r<t*s |] ==> y*x<t*s";
-by (rtac ([([prem1,prem2] MRS hypreal_mult_less_mono2),prem3] MRS hypreal_less_trans) 1);
-qed "hypreal_mult_less_trans";
-
-Goal "[| 0<=y; x<r; y*r<t*s; (0::hypreal)<t*s|] ==> y*x<t*s";
-by (dtac hypreal_le_imp_less_or_eq 1);
-by (fast_tac (HOL_cs addEs [(hypreal_mult_0_less RS iffD2),hypreal_mult_less_trans]) 1);
-qed "hypreal_mult_le_less_trans";
-
-Goal "[| 0 <= y; x <= r; y*r < t*s; (0::hypreal) < t*s|] ==> y*x < t*s";
-by (dres_inst_tac [("x","x")] hypreal_le_imp_less_or_eq 1);
-by (fast_tac (claset() addIs [hypreal_mult_le_less_trans]) 1);
-qed "hypreal_mult_le_le_trans";
-
-Goal "[| 0 < r1; r1 <r2; (0::hypreal) < x; x < y|] \
-\                     ==> r1 * x < r2 * y";
-by (dres_inst_tac [("x","x")] hypreal_mult_less_mono2 1);
-by (dres_inst_tac [("R1.0","0")] hypreal_less_trans 2);
-by (dres_inst_tac [("x","r1")] hypreal_mult_less_mono1 3);
-by Auto_tac;
-by (blast_tac (claset() addIs [hypreal_less_trans]) 1);
-qed "hypreal_mult_less_mono";
-
-Goal "[| 0 < r1; r1 <r2; 0 < y|] \
-\                           ==> (0::hypreal) < r2 * y";
-by (dres_inst_tac [("R1.0","0")] hypreal_less_trans 1);
-by (assume_tac 1);
-by (blast_tac (claset() addIs [hypreal_mult_order]) 1);
-qed "hypreal_mult_order_trans";
-
-Goal "[| 0 < r1; r1 <= r2; (0::hypreal) <= x; x <= y |] \
-\                  ==> r1 * x <= r2 * y";
-by (rtac hypreal_less_or_eq_imp_le 1);
-by (REPEAT(dtac hypreal_le_imp_less_or_eq 1));
-by (auto_tac (claset() addIs [hypreal_mult_less_mono,
-    hypreal_mult_less_mono1,hypreal_mult_less_mono2,
-    hypreal_mult_order_trans,hypreal_mult_order],simpset()));
-qed "hypreal_mult_le_mono";
-
-Goal "0 < x ==> 0 < inverse (x::hypreal)";
-by (EVERY1[rtac ccontr, dtac hypreal_leI]);
-by (forward_tac [hypreal_minus_zero_less_iff2 RS iffD2] 1);
-by (forward_tac [hypreal_not_refl2 RS not_sym] 1);
-by (dtac (hypreal_not_refl2 RS not_sym RS hypreal_inverse_not_zero) 1);
-by (EVERY1[dtac hypreal_le_imp_less_or_eq, Step_tac]); 
-by (dtac hypreal_mult_less_zero1 1 THEN assume_tac 1);
-by (auto_tac (claset() addIs [hypreal_zero_less_one RS hypreal_less_asym],
-    simpset() addsimps [hypreal_minus_zero_less_iff]));
-qed "hypreal_inverse_gt_zero";
-
-Goal "x < 0 ==> inverse (x::hypreal) < 0";
-by (ftac hypreal_not_refl2 1);
-by (dtac (hypreal_minus_zero_less_iff RS iffD2) 1);
-by (rtac (hypreal_minus_zero_less_iff RS iffD1) 1);
-by (stac (hypreal_minus_inverse RS sym) 1);
-by (auto_tac (claset() addIs [hypreal_inverse_gt_zero],  simpset()));
-qed "hypreal_inverse_less_zero";
-
-(* check why this does not work without 2nd substitution anymore! *)
-Goal "x < y ==> x < (x + y)*inverse(1hr + 1hr)";
-by (dres_inst_tac [("C","x")] hypreal_add_less_mono2 1);
-by (dtac (hypreal_add_self RS subst) 1);
-by (dtac (hypreal_zero_less_two RS hypreal_inverse_gt_zero RS 
-          hypreal_mult_less_mono1) 1);
-by (auto_tac (claset() addDs [hypreal_two_not_zero RS 
-          (hypreal_mult_inverse RS subst)],simpset() 
-          addsimps [hypreal_mult_assoc]));
-qed "hypreal_less_half_sum";
-
-(* check why this does not work without 2nd substitution anymore! *)
-Goal "x < y ==> (x + y)*inverse(1hr + 1hr) < y";
-by (dres_inst_tac [("C","y")] hypreal_add_less_mono1 1);
-by (dtac (hypreal_add_self RS subst) 1);
-by (dtac (hypreal_zero_less_two RS hypreal_inverse_gt_zero RS 
-          hypreal_mult_less_mono1) 1);
-by (auto_tac (claset() addDs [hypreal_two_not_zero RS 
-          (hypreal_mult_inverse RS subst)],simpset() 
-          addsimps [hypreal_mult_assoc]));
-qed "hypreal_gt_half_sum";
-
-Goal "!!(x::hypreal). x < y ==> EX r. x < r & r < y";
-by (blast_tac (claset() addSIs [hypreal_less_half_sum,
-                                hypreal_gt_half_sum]) 1);
-qed "hypreal_dense";
-
-
-(*?? more needed and earlier??*)
-Goal "(x+y = (0::hypreal)) = (x = -y)";
-by (stac hypreal_eq_minus_iff 1);
-by Auto_tac;
-qed "hypreal_add_eq_0_iff";
-AddIffs [hypreal_add_eq_0_iff];
-
-Goal "[| 0 <= x; 0 <= y |] ==> (x+y = 0) = (x = 0 & y = (0::hypreal))";
-by (auto_tac (claset() addIs [order_antisym], simpset()));
-qed "hypreal_add_nonneg_eq_0_iff";
-
-Goal "(x*x + y*y = 0) = (x = 0 & y = (0::hypreal))";
-by (simp_tac
-    (HOL_ss addsimps [hypreal_le_square, hypreal_add_nonneg_eq_0_iff]) 1);
-by Auto_tac;
-qed "hypreal_squares_add_zero_iff";
-Addsimps [hypreal_squares_add_zero_iff];
-
-Goal "(x*x + y*y + z*z = 0) = (x = 0 & y = 0 & z = (0::hypreal))";
-by (simp_tac (HOL_ss addsimps [hypreal_le_square, hypreal_le_add_order, 
-                               hypreal_add_nonneg_eq_0_iff]) 1);
-by Auto_tac;
-qed "hypreal_three_squares_add_zero_iff";
-Addsimps [hypreal_three_squares_add_zero_iff];
-
-Addsimps [rename_numerals real_le_square];
-
-Goal "(x::hypreal)*x <= x*x + y*y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-              simpset() addsimps [hypreal_mult,hypreal_add,hypreal_le]));
-qed "hypreal_self_le_add_pos";
-Addsimps [hypreal_self_le_add_pos];
-
-(*lcp: new lemma unfortunately needed...*)
-Goal "-(x*x) <= (y*y::real)";
-by (rtac order_trans 1);
-by (rtac real_le_square 2);
-by Auto_tac;
-qed "minus_square_le_square";
-
-Goal "(x::hypreal)*x <= x*x + y*y + z*z";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-	      simpset() addsimps [hypreal_mult, hypreal_add, hypreal_le,
-				  minus_square_le_square]));
-qed "hypreal_self_le_add_pos2";
-Addsimps [hypreal_self_le_add_pos2];
-
-
-(*----------------------------------------------------------------------------
-             Embedding of the naturals in the hyperreals
- ----------------------------------------------------------------------------*)
-Goalw [hypreal_of_posnat_def] "hypreal_of_posnat 0 = 1hr";
-by (full_simp_tac (simpset() addsimps 
-    [pnat_one_iff RS sym,real_of_preal_def]) 1);
-by (fold_tac [real_one_def]);
-by (simp_tac (simpset() addsimps [hypreal_of_real_one]) 1);
-qed "hypreal_of_posnat_one";
-
-Goalw [hypreal_of_posnat_def] "hypreal_of_posnat 1 = 1hr + 1hr";
-by (full_simp_tac (simpset() addsimps 
-		   [real_of_preal_def,
-		    rename_numerals (real_one_def RS meta_eq_to_obj_eq),
-		    hypreal_add,hypreal_of_real_def,pnat_two_eq,
-		    hypreal_one_def, real_add,
-		    prat_of_pnat_add RS sym, preal_of_prat_add RS sym] @ 
-		    pnat_add_ac) 1);
-qed "hypreal_of_posnat_two";
-
-(*FIXME: delete this and all posnat*)
-Goalw [hypreal_of_posnat_def]
-     "hypreal_of_posnat n1 + hypreal_of_posnat n2 = \
-\     hypreal_of_posnat (n1 + n2) + 1hr";
-by (full_simp_tac (HOL_ss addsimps [hypreal_of_posnat_one RS sym,
-    hypreal_of_real_add RS sym,hypreal_of_posnat_def,real_of_preal_add RS sym,
-    preal_of_prat_add RS sym,prat_of_pnat_add RS sym,pnat_of_nat_add]) 1);
-qed "hypreal_of_posnat_add";
-
-Goal "hypreal_of_posnat (n + 1) = hypreal_of_posnat n + 1hr";
-by (res_inst_tac [("x1","1hr")] (hypreal_add_right_cancel RS iffD1) 1);
-by (rtac (hypreal_of_posnat_add RS subst) 1);
-by (full_simp_tac (simpset() addsimps [hypreal_of_posnat_two,hypreal_add_assoc]) 1);
-qed "hypreal_of_posnat_add_one";
-
-Goalw [real_of_posnat_def,hypreal_of_posnat_def] 
-      "hypreal_of_posnat n = hypreal_of_real (real_of_posnat n)";
-by (rtac refl 1);
-qed "hypreal_of_real_of_posnat";
-
-Goalw [hypreal_of_posnat_def] 
-      "(n < m) = (hypreal_of_posnat n < hypreal_of_posnat m)";
-by Auto_tac;
-qed "hypreal_of_posnat_less_iff";
-
-Addsimps [hypreal_of_posnat_less_iff RS sym];
-(*---------------------------------------------------------------------------------
-               Existence of infinite hyperreal number
- ---------------------------------------------------------------------------------*)
-
-Goal "hyprel^^{%n::nat. real_of_posnat n} : hypreal";
-by Auto_tac;
-qed "hypreal_omega";
-
-Goalw [omega_def] "Rep_hypreal(whr) : hypreal";
-by (rtac Rep_hypreal 1);
-qed "Rep_hypreal_omega";
-
-(* existence of infinite number not corresponding to any real number *)
-(* use assumption that member FreeUltrafilterNat is not finite       *)
-(* a few lemmas first *)
-
-Goal "{n::nat. x = real_of_posnat n} = {} | \
-\     (EX y. {n::nat. x = real_of_posnat n} = {y})";
-by (auto_tac (claset() addDs [inj_real_of_posnat RS injD],simpset()));
-qed "lemma_omega_empty_singleton_disj";
-
-Goal "finite {n::nat. x = real_of_posnat n}";
-by (cut_inst_tac [("x","x")] lemma_omega_empty_singleton_disj 1);
-by Auto_tac;
-qed "lemma_finite_omega_set";
-
-Goalw [omega_def,hypreal_of_real_def] 
-      "~ (EX x. hypreal_of_real x = whr)";
-by (auto_tac (claset(),simpset() addsimps [lemma_finite_omega_set 
-    RS FreeUltrafilterNat_finite]));
-qed "not_ex_hypreal_of_real_eq_omega";
-
-Goal "hypreal_of_real x ~= whr";
-by (cut_facts_tac [not_ex_hypreal_of_real_eq_omega] 1);
-by Auto_tac;
-qed "hypreal_of_real_not_eq_omega";
-
-(* existence of infinitesimal number also not *)
-(* corresponding to any real number *)
-
-Goal "{n::nat. x = inverse(real_of_posnat n)} = {} | \
-\     (EX y. {n::nat. x = inverse(real_of_posnat n)} = {y})";
-by (Step_tac 1 THEN Step_tac 1);
-by (auto_tac (claset() addIs [real_of_posnat_inverse_inj],simpset()));
-qed "lemma_epsilon_empty_singleton_disj";
-
-Goal "finite {n::nat. x = inverse(real_of_posnat n)}";
-by (cut_inst_tac [("x","x")] lemma_epsilon_empty_singleton_disj 1);
-by Auto_tac;
-qed "lemma_finite_epsilon_set";
-
-Goalw [epsilon_def,hypreal_of_real_def] 
-      "~ (EX x. hypreal_of_real x = ehr)";
-by (auto_tac (claset(),simpset() addsimps [lemma_finite_epsilon_set 
-    RS FreeUltrafilterNat_finite]));
-qed "not_ex_hypreal_of_real_eq_epsilon";
-
-Goal "hypreal_of_real x ~= ehr";
-by (cut_facts_tac [not_ex_hypreal_of_real_eq_epsilon] 1);
-by Auto_tac;
-qed "hypreal_of_real_not_eq_epsilon";
-
-Goalw [epsilon_def,hypreal_zero_def] "ehr ~= 0";
-by (auto_tac (claset(),
-     simpset() addsimps [rename_numerals real_of_posnat_not_eq_zero]));
-qed "hypreal_epsilon_not_zero";
-
-Addsimps [rename_numerals real_of_posnat_not_eq_zero];
-Goalw [omega_def,hypreal_zero_def] "whr ~= 0";
-by (Simp_tac 1);
-qed "hypreal_omega_not_zero";
-
-Goal "ehr = inverse(whr)";
-by (asm_full_simp_tac (simpset() addsimps 
-                     [hypreal_inverse, omega_def, epsilon_def]) 1);
-qed "hypreal_epsilon_inverse_omega";
-
-(*----------------------------------------------------------------
-     Another embedding of the naturals in the 
-    hyperreals (see hypreal_of_posnat)
- ----------------------------------------------------------------*)
-Goalw [hypreal_of_nat_def] "hypreal_of_nat 0 = 0";
-by (full_simp_tac (simpset() addsimps [hypreal_of_posnat_one]) 1);
-qed "hypreal_of_nat_zero";
-
-Goalw [hypreal_of_nat_def] "hypreal_of_nat 1 = 1hr";
-by (full_simp_tac (simpset() addsimps [hypreal_of_posnat_two,
-                                       hypreal_add_assoc]) 1);
-qed "hypreal_of_nat_one";
-
-Goalw [hypreal_of_nat_def]
-     "hypreal_of_nat n1 + hypreal_of_nat n2 = hypreal_of_nat (n1 + n2)";
-by (full_simp_tac (simpset() addsimps hypreal_add_ac) 1);
-by (simp_tac (simpset() addsimps [hypreal_of_posnat_add,
-                                  hypreal_add_assoc RS sym]) 1);
-qed "hypreal_of_nat_add";
-
-Goal "hypreal_of_nat 2 = 1hr + 1hr";
-by (simp_tac (simpset() addsimps [hypreal_of_nat_one 
-    RS sym,hypreal_of_nat_add]) 1);
-qed "hypreal_of_nat_two";
-
-Goalw [hypreal_of_nat_def] 
-      "(n < m) = (hypreal_of_nat n < hypreal_of_nat m)";
-by (auto_tac (claset() addIs [hypreal_add_less_mono1],simpset()));
-qed "hypreal_of_nat_less_iff";
-Addsimps [hypreal_of_nat_less_iff RS sym];
-
-(*------------------------------------------------------------*)
-(* naturals embedded in hyperreals                            *)
-(* is a hyperreal c.f. NS extension                           *)
-(*------------------------------------------------------------*)
-
-Goalw [hypreal_of_nat_def,real_of_nat_def] 
-      "hypreal_of_nat  m = Abs_hypreal(hyprel^^{%n. real_of_nat m})";
-by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_def,
-    hypreal_of_real_of_posnat,hypreal_minus,hypreal_one_def,hypreal_add]));
-qed "hypreal_of_nat_iff";
-
-Goal "inj hypreal_of_nat";
-by (rtac injI 1);
-by (auto_tac (claset() addSDs [FreeUltrafilterNat_P],
-        simpset() addsimps [split_if_mem1, hypreal_of_nat_iff,
-        real_add_right_cancel,inj_real_of_nat RS injD]));
-qed "inj_hypreal_of_nat";
-
-Goalw [hypreal_of_nat_def,hypreal_of_real_def,hypreal_of_posnat_def,
-       real_of_posnat_def,hypreal_one_def,real_of_nat_def] 
-       "hypreal_of_nat n = hypreal_of_real (real_of_nat n)";
-by (simp_tac (simpset() addsimps [hypreal_add,hypreal_minus]) 1);
-qed "hypreal_of_nat_real_of_nat";
-
-Goal "hypreal_of_posnat (Suc n) = hypreal_of_posnat n + 1hr";
-by (stac (hypreal_of_posnat_add_one RS sym) 1);
-by (Simp_tac 1);
-qed "hypreal_of_posnat_Suc";
-
-Goalw [hypreal_of_nat_def] 
-      "hypreal_of_nat (Suc n) = hypreal_of_nat n + 1hr";
-by (simp_tac (simpset() addsimps [hypreal_of_posnat_Suc] @ hypreal_add_ac) 1);
-qed "hypreal_of_nat_Suc";
-
-Goal "0 < r ==> 0 < r/(1hr+1hr)";
-by (dtac (hypreal_zero_less_two RS hypreal_inverse_gt_zero 
-          RS hypreal_mult_less_mono1) 1);
-by (auto_tac (claset(), simpset() addsimps [hypreal_divide_def]));
-qed "hypreal_half_gt_zero";
-
-(* this proof is so much simpler than one for reals!! *)
-Goal "[| 0 < r; r < x |] ==> inverse x < inverse (r::hypreal)";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_inverse,
-    hypreal_less,hypreal_zero_def]));
-by (ultra_tac (claset() addIs [real_inverse_less_swap],simpset()) 1);
-qed "hypreal_inverse_less_swap";
-
-Goal "[| 0 < r; 0 < x|] ==> (r < x) = (inverse x < inverse (r::hypreal))";
-by (auto_tac (claset() addIs [hypreal_inverse_less_swap],simpset()));
-by (res_inst_tac [("t","r")] (hypreal_inverse_inverse RS subst) 1);
-by (res_inst_tac [("t","x")] (hypreal_inverse_inverse RS subst) 1);
-by (rtac hypreal_inverse_less_swap 1);
-by (auto_tac (claset(), simpset() addsimps [hypreal_inverse_gt_zero]));
-qed "hypreal_inverse_less_iff";
-
-Goal "[| 0 < z; x < y |] ==> x * inverse z < y * inverse (z::hypreal)";
-by (blast_tac (claset() addSIs [hypreal_mult_less_mono1,
-                                hypreal_inverse_gt_zero]) 1);
-qed "hypreal_mult_inverse_less_mono1";
-
-Goal "[| 0 < z; x < y |] ==> inverse z * x < inverse (z::hypreal) * y";
-by (blast_tac (claset() addSIs [hypreal_mult_less_mono2,
-                                hypreal_inverse_gt_zero]) 1);
-qed "hypreal_mult_inverse_less_mono2";
-
-Goal "[| (0::hypreal) < z; x*z < y*z |] ==> x < y";
-by (forw_inst_tac [("x","x*z")] hypreal_mult_inverse_less_mono1 1);
-by (dtac (hypreal_not_refl2 RS not_sym) 2);
-by (auto_tac (claset() addSDs [hypreal_mult_inverse],
-              simpset() addsimps hypreal_mult_ac));
-qed "hypreal_less_mult_right_cancel";
-
-Goal "[| (0::hypreal) < z; z*x < z*y |] ==> x < y";
-by (auto_tac (claset() addIs [hypreal_less_mult_right_cancel],
-    simpset() addsimps [hypreal_mult_commute]));
-qed "hypreal_less_mult_left_cancel";
-
-Goal "[| 0 < r; (0::hypreal) < ra; r < x; ra < y |] ==> r*ra < x*y";
-by (forw_inst_tac [("R2.0","r")] hypreal_less_trans 1);
-by (dres_inst_tac [("z","ra"),("x","r")] hypreal_mult_less_mono1 2);
-by (dres_inst_tac [("z","x"),("x","ra")] hypreal_mult_less_mono2 3);
-by (auto_tac (claset() addIs [hypreal_less_trans],simpset()));
-qed "hypreal_mult_less_gt_zero"; 
-
-Goal "[| 0 < r; (0::hypreal) < ra; r <= x; ra <= y |] ==> r*ra <= x*y";
-by (REPEAT(dtac hypreal_le_imp_less_or_eq 1));
-by (rtac hypreal_less_or_eq_imp_le 1);
-by (auto_tac (claset() addIs [hypreal_mult_less_mono1,
-    hypreal_mult_less_mono2,hypreal_mult_less_gt_zero],
-    simpset()));
-qed "hypreal_mult_le_ge_zero"; 
-
-(*----------------------------------------------------------------------------
-     Some convenient biconditionals for products of signs 
- ----------------------------------------------------------------------------*)
-
-Goal "((0::hypreal) < x*y) = (0 < x & 0 < y | x < 0 & y < 0)";
- by (auto_tac (claset(), 
-               simpset() addsimps [order_le_less, linorder_not_less, 
-                              hypreal_mult_order, hypreal_mult_less_zero1]));
-by (ALLGOALS (rtac ccontr)); 
-by (auto_tac (claset(), 
-              simpset() addsimps [order_le_less, linorder_not_less]));
-by (ALLGOALS (etac rev_mp)); 
-by (ALLGOALS (dtac hypreal_mult_less_zero THEN' assume_tac));
-by (auto_tac (claset() addDs [order_less_not_sym], 
-              simpset() addsimps [hypreal_mult_commute]));  
-qed "hypreal_zero_less_mult_iff";
-
-Goal "((0::hypreal) <= x*y) = (0 <= x & 0 <= y | x <= 0 & y <= 0)";
-by (auto_tac (claset() addDs [hypreal_mult_zero_disj],
-              simpset() addsimps [order_le_less, linorder_not_less,
-                                  hypreal_zero_less_mult_iff]));
-qed "hypreal_zero_le_mult_iff";
-
-Goal "(x*y < (0::hypreal)) = (0 < x & y < 0 | x < 0 & 0 < y)";
-by (auto_tac (claset(), 
-              simpset() addsimps [hypreal_zero_le_mult_iff, 
-                                  linorder_not_le RS sym]));
-by (auto_tac (claset() addDs [order_less_not_sym],  
-              simpset() addsimps [linorder_not_le]));
-qed "hypreal_mult_less_zero_iff";
-
-Goal "(x*y <= (0::hypreal)) = (0 <= x & y <= 0 | x <= 0 & 0 <= y)";
-by (auto_tac (claset() addDs [order_less_not_sym], 
-              simpset() addsimps [hypreal_zero_less_mult_iff, 
-                                  linorder_not_less RS sym]));
-qed "hypreal_mult_le_zero_iff";
-
--- a/src/HOL/Real/Hyperreal/HyperOrd.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,15 +0,0 @@
-(*  Title:	 Real/Hyperreal/HyperOrd.thy
-    Author:      Jacques D. Fleuriot
-    Copyright:   2000 University of Edinburgh
-    Description: Type "hypreal" is a linear order and also 
-                 satisfies plus_ac0: + is an AC-operator with zero
-*)
-
-HyperOrd = HyperDef +
-
-instance hypreal :: order (hypreal_le_refl,hypreal_le_trans,hypreal_le_anti_sym,hypreal_less_le)
-instance hypreal :: linorder (hypreal_le_linear)
-
-instance hypreal :: plus_ac0(hypreal_add_commute,hypreal_add_assoc,hypreal_add_zero_left)
-
-end
--- a/src/HOL/Real/Hyperreal/HyperPow.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,581 +0,0 @@
-(*  Title       : HyperPow.ML
-    Author      : Jacques D. Fleuriot  
-    Copyright   : 1998  University of Cambridge
-    Description : Natural Powers of hyperreals theory
-
-*)
-
-Goal "(0::hypreal) ^ (Suc n) = 0";
-by (Auto_tac);
-qed "hrealpow_zero";
-Addsimps [hrealpow_zero];
-
-Goal "r ~= (0::hypreal) --> r ^ n ~= 0";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [hypreal_mult_not_0E],
-    simpset() addsimps [hypreal_zero_not_eq_one RS not_sym]));
-qed_spec_mp "hrealpow_not_zero";
-
-Goal "r ~= (0::hypreal) --> inverse(r ^ n) = (inverse r) ^ n";
-by (induct_tac "n" 1);
-by (Auto_tac);
-by (forw_inst_tac [("n","n")] hrealpow_not_zero 1);
-by (auto_tac (claset(), simpset() addsimps [hypreal_inverse_distrib]));
-qed_spec_mp "hrealpow_inverse";
-
-Goal "abs (r::hypreal) ^ n = abs (r ^ n)";
-by (induct_tac "n" 1);
-by (auto_tac (claset(), simpset() addsimps [hrabs_mult,hrabs_one]));
-qed "hrealpow_hrabs";
-
-Goal "(r::hypreal) ^ (n + m) = (r ^ n) * (r ^ m)";
-by (induct_tac "n" 1);
-by (auto_tac (claset(), simpset() addsimps hypreal_mult_ac));
-qed "hrealpow_add";
-
-Goal "(r::hypreal) ^ 1 = r";
-by (Simp_tac 1);
-qed "hrealpow_one";
-Addsimps [hrealpow_one];
-
-Goal "(r::hypreal) ^ 2 = r * r";
-by (Simp_tac 1);
-qed "hrealpow_two";
-
-Goal "(0::hypreal) < r --> 0 <= r ^ n";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addDs [hypreal_less_imp_le] 
-                       addIs [hypreal_le_mult_order],
-             simpset() addsimps [hypreal_zero_less_one RS hypreal_less_imp_le]));
-qed_spec_mp "hrealpow_ge_zero";
-
-Goal "(0::hypreal) < r --> 0 < r ^ n";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [hypreal_mult_order],
-              simpset() addsimps [hypreal_zero_less_one]));
-qed_spec_mp "hrealpow_gt_zero";
-
-Goal "(0::hypreal) <= r --> 0 <= r ^ n";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [hypreal_le_mult_order],
-          simpset() addsimps [hypreal_zero_less_one RS hypreal_less_imp_le]));
-qed_spec_mp "hrealpow_ge_zero2";
-
-Goal "(0::hypreal) < x & x <= y --> x ^ n <= y ^ n";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addSIs [hypreal_mult_le_mono],
-              simpset() addsimps [hypreal_le_refl]));
-by (asm_simp_tac (simpset() addsimps [hrealpow_ge_zero]) 1);
-qed_spec_mp "hrealpow_le";
-
-Goal "(0::hypreal) < x & x < y & 0 < n --> x ^ n < y ^ n";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [hypreal_mult_less_mono,gr0I] 
-                       addDs [hrealpow_gt_zero],
-              simpset()));
-qed "hrealpow_less";
-
-Goal "1hr ^ n = 1hr";
-by (induct_tac "n" 1);
-by (Auto_tac);
-qed "hrealpow_eq_one";
-Addsimps [hrealpow_eq_one];
-
-Goal "abs(-(1hr ^ n)) = 1hr";
-by (simp_tac (simpset() addsimps [hrabs_one]) 1);
-qed "hrabs_minus_hrealpow_one";
-Addsimps [hrabs_minus_hrealpow_one];
-
-Goal "abs((-1hr) ^ n) = 1hr";
-by (induct_tac "n" 1);
-by (auto_tac (claset(),
-      simpset() addsimps [hrabs_mult, hrabs_one]));
-qed "hrabs_hrealpow_minus_one";
-Addsimps [hrabs_hrealpow_minus_one];
-
-Goal "((r::hypreal) * s) ^ n = (r ^ n) * (s ^ n)";
-by (induct_tac "n" 1);
-by (auto_tac (claset(), simpset() addsimps hypreal_mult_ac));
-qed "hrealpow_mult";
-
-Goal "(0::hypreal) <= r ^ 2";
-by (simp_tac (simpset() addsimps [hrealpow_two]) 1);
-qed "hrealpow_two_le";
-Addsimps [hrealpow_two_le];
-
-Goal "(0::hypreal) <= u ^ 2 + v ^ 2";
-by (auto_tac (claset() addIs [hypreal_le_add_order], simpset()));
-qed "hrealpow_two_le_add_order";
-Addsimps [hrealpow_two_le_add_order];
-
-Goal "(0::hypreal) <= u ^ 2 + v ^ 2 + w ^ 2";
-by (auto_tac (claset() addSIs [hypreal_le_add_order], simpset()));
-qed "hrealpow_two_le_add_order2";
-Addsimps [hrealpow_two_le_add_order2];
-
-Goal "(x ^ 2 + y ^ 2 + z ^ 2 = (0::hypreal)) = (x = 0 & y = 0 & z = 0)";
-by (simp_tac (HOL_ss addsimps [hypreal_three_squares_add_zero_iff, 
-                               hrealpow_two]) 1);
-qed "hrealpow_three_squares_add_zero_iff";
-Addsimps [hrealpow_three_squares_add_zero_iff];
-
-Goal "abs(x ^ 2) = (x::hypreal) ^ 2";
-by (simp_tac (simpset() addsimps [hrabs_eqI1]) 1);
-qed "hrabs_hrealpow_two";
-Addsimps [hrabs_hrealpow_two];
-
-Goal "abs(x) ^ 2 = (x::hypreal) ^ 2";
-by (simp_tac (simpset() addsimps [hrealpow_hrabs, hrabs_eqI1] 
-                        delsimps [hpowr_Suc]) 1);
-qed "hrealpow_two_hrabs";
-Addsimps [hrealpow_two_hrabs];
-
-Goal "!!r. 1hr < r ==> 1hr < r ^ 2";
-by (auto_tac (claset(), simpset() addsimps [hrealpow_two]));
-by (cut_facts_tac [hypreal_zero_less_one] 1);
-by (forw_inst_tac [("R1.0","0")] hypreal_less_trans 1);
-by (assume_tac 1);
-by (dres_inst_tac [("z","r"),("x","1hr")] hypreal_mult_less_mono1 1);
-by (auto_tac (claset() addIs [hypreal_less_trans], simpset()));
-qed "hrealpow_two_gt_one";
-
-Goal "!!r. 1hr <= r ==> 1hr <= r ^ 2";
-by (etac (hypreal_le_imp_less_or_eq RS disjE) 1);
-by (etac (hrealpow_two_gt_one RS hypreal_less_imp_le) 1);
-by (asm_simp_tac (simpset() addsimps [hypreal_le_refl]) 1);
-qed "hrealpow_two_ge_one";
-
-Goal "!!r. (0::hypreal) < r ==> 0 < r ^ Suc n";
-by (forw_inst_tac [("n","n")] hrealpow_ge_zero 1);
-by (forw_inst_tac [("n1","n")]
-    ((hypreal_not_refl2 RS not_sym) RS hrealpow_not_zero RS not_sym) 1);
-by (auto_tac (claset() addSDs [hypreal_le_imp_less_or_eq]
-                       addIs [hypreal_mult_order],
-              simpset()));
-qed "hrealpow_Suc_gt_zero";
-
-Goal "!!r. (0::hypreal) <= r ==> 0 <= r ^ Suc n";
-by (etac (hypreal_le_imp_less_or_eq RS disjE) 1);
-by (etac (hrealpow_ge_zero) 1);
-by (asm_simp_tac (simpset() addsimps [hypreal_le_refl]) 1);
-qed "hrealpow_Suc_ge_zero";
-
-Goal "1hr <= (1hr +1hr) ^ n";
-by (res_inst_tac [("j","1hr ^ n")] hypreal_le_trans 1);
-by (rtac hrealpow_le 2);
-by (auto_tac (claset() addIs [hypreal_less_imp_le],
-    simpset() addsimps [hypreal_zero_less_one,
-    hypreal_one_less_two,hypreal_le_refl]));
-qed "two_hrealpow_ge_one";
-
-Goal "hypreal_of_nat n < (1hr + 1hr) ^ n";
-by (induct_tac "n" 1);
-by (auto_tac (claset(),
-        simpset() addsimps [hypreal_of_nat_Suc,hypreal_of_nat_zero,
-          hypreal_zero_less_one,hypreal_add_mult_distrib,hypreal_of_nat_one]));
-by (blast_tac (claset() addSIs [hypreal_add_less_le_mono,
-    two_hrealpow_ge_one]) 1);
-qed "two_hrealpow_gt";
-Addsimps [two_hrealpow_gt,two_hrealpow_ge_one];
-
-Goal "(-1hr) ^ (2*n) = 1hr";
-by (induct_tac "n" 1);
-by (Auto_tac);
-qed "hrealpow_minus_one";
-
-Goal "(-1hr) ^ (n + n) = 1hr";
-by (induct_tac "n" 1);
-by (Auto_tac);
-qed "hrealpow_minus_one2";
-Addsimps [hrealpow_minus_one2];
-
-Goal "(-(x::hypreal)) ^ 2 = x ^ 2";
-by (Auto_tac);
-qed "hrealpow_minus_two";
-Addsimps [hrealpow_minus_two];
-
-Goal "(0::hypreal) < r & r < 1hr --> r ^ Suc n < r ^ n";
-by (induct_tac "n" 1);
-by (auto_tac (claset(),
-              simpset() addsimps [hypreal_mult_less_mono2]));
-qed_spec_mp "hrealpow_Suc_less";
-
-Goal "(0::hypreal) <= r & r < 1hr --> r ^ Suc n <= r ^ n";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [hypreal_less_imp_le]
-                       addSDs [hypreal_le_imp_less_or_eq,hrealpow_Suc_less],
-      simpset() addsimps [hypreal_le_refl,hypreal_mult_less_mono2]));
-qed_spec_mp "hrealpow_Suc_le";
-
-(* a few more theorems needed here *)
-Goal "1hr <= r --> r ^ n <= r ^ Suc n";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addSIs [hypreal_mult_le_le_mono1], simpset()));
-by (rtac ccontr 1 THEN dtac not_hypreal_leE 1);
-by (dtac hypreal_le_less_trans 1 THEN assume_tac 1);
-by (etac (hypreal_zero_less_one RS hypreal_less_asym) 1);
-qed "hrealpow_less_Suc";
-
-Goal "Abs_hypreal(hyprel^^{%n. X n}) ^ m = Abs_hypreal(hyprel^^{%n. (X n) ^ m})";
-by (nat_ind_tac "m" 1);
-by (auto_tac (claset(),
-              simpset() addsimps [hypreal_one_def,hypreal_mult]));
-qed "hrealpow";
-
-Goal "(x + (y::hypreal)) ^ 2 = \
-\     x ^ 2 + y ^ 2 + (hypreal_of_nat 2)*x*y";
-by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib2,
-    hypreal_add_mult_distrib,hypreal_of_nat_two] 
-    @ hypreal_add_ac @ hypreal_mult_ac) 1);
-qed "hrealpow_sum_square_expand";
-
-(*---------------------------------------------------------------
-   we'll prove the following theorem by going down to the
-   level of the ultrafilter and relying on the analogous
-   property for the real rather than prove it directly 
-   using induction: proof is much simpler this way!
- ---------------------------------------------------------------*)
-Goalw [hypreal_zero_def] 
-  "[|(0::hypreal) <= x;0 <= y;x ^ Suc n <= y ^ Suc n |] ==> x <= y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-              simpset() addsimps [hrealpow,hypreal_le,hypreal_mult]));
-by (ultra_tac (claset() addIs [realpow_increasing], simpset()) 1);
-qed "hrealpow_increasing";
-
-goalw HyperPow.thy [hypreal_zero_def] 
-  "!!x. [|(0::hypreal) <= x;0 <= y;x ^ Suc n = y ^ Suc n |] ==> x = y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(), simpset() addsimps [hrealpow,hypreal_mult,hypreal_le]));
-by (ultra_tac (claset() addIs [realpow_Suc_cancel_eq],
-    simpset()) 1);
-qed "hrealpow_Suc_cancel_eq";
-
-Goal "x : HFinite --> x ^ n : HFinite";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [HFinite_mult], simpset()));
-qed_spec_mp "hrealpow_HFinite";
-
-(*---------------------------------------------------------------
-                  Hypernaturals Powers
- --------------------------------------------------------------*)
-Goalw [congruent_def]
-     "congruent hyprel \
-\    (%X Y. hyprel^^{%n. ((X::nat=>real) n ^ (Y::nat=>nat) n)})";
-by (safe_tac (claset() addSIs [ext]));
-by (ALLGOALS(Fuf_tac));
-qed "hyperpow_congruent";
-
-Goalw [hyperpow_def]
-  "Abs_hypreal(hyprel^^{%n. X n}) pow Abs_hypnat(hypnatrel^^{%n. Y n}) = \
-\  Abs_hypreal(hyprel^^{%n. X n ^ Y n})";
-by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
-by (auto_tac (claset() addSIs [lemma_hyprel_refl,bexI],
-    simpset() addsimps [hyprel_in_hypreal RS 
-    Abs_hypreal_inverse,equiv_hyprel,hyperpow_congruent]));
-by (Fuf_tac 1);
-qed "hyperpow";
-
-Goalw [hypreal_zero_def,hypnat_one_def]
-      "(0::hypreal) pow (n + 1hn) = 0";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [hyperpow,hypnat_add]));
-qed "hyperpow_zero";
-Addsimps [hyperpow_zero];
-
-Goalw [hypreal_zero_def]
-      "r ~= (0::hypreal) --> r pow n ~= 0";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset(), simpset() addsimps [hyperpow]));
-by (dtac FreeUltrafilterNat_Compl_mem 1);
-by (fuf_empty_tac (claset() addIs [realpow_not_zero RS notE],
-    simpset()) 1);
-qed_spec_mp "hyperpow_not_zero";
-
-Goalw [hypreal_zero_def] 
-      "r ~= (0::hypreal) --> inverse(r pow n) = (inverse r) pow n";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem],
-    simpset() addsimps [hypreal_inverse,hyperpow]));
-by (rtac FreeUltrafilterNat_subset 1);
-by (auto_tac (claset() addDs [realpow_not_zero] 
-                       addIs [realpow_inverse], 
-              simpset()));
-qed "hyperpow_inverse";
-
-Goal "abs r pow n = abs (r pow n)";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-              simpset() addsimps [hypreal_hrabs, hyperpow,realpow_abs]));
-qed "hyperpow_hrabs";
-
-Goal "r pow (n + m) = (r pow n) * (r pow m)";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-          simpset() addsimps [hyperpow,hypnat_add, hypreal_mult,realpow_add]));
-qed "hyperpow_add";
-
-Goalw [hypnat_one_def] "r pow 1hn = r";
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset(), simpset() addsimps [hyperpow]));
-qed "hyperpow_one";
-Addsimps [hyperpow_one];
-
-Goalw [hypnat_one_def] 
-      "r pow (1hn + 1hn) = r * r";
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-           simpset() addsimps [hyperpow,hypnat_add, hypreal_mult,realpow_two]));
-qed "hyperpow_two";
-
-Goalw [hypreal_zero_def]
-      "(0::hypreal) < r --> 0 < r pow n";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSEs [FreeUltrafilterNat_subset, realpow_gt_zero],
-       simpset() addsimps [hyperpow,hypreal_less, hypreal_le]));
-qed_spec_mp "hyperpow_gt_zero";
-
-Goal "(0::hypreal) < r --> 0 <= r pow n";
-by (blast_tac (claset() addSIs [hyperpow_gt_zero, hypreal_less_imp_le]) 1);
-qed_spec_mp "hyperpow_ge_zero";
-
-Goalw [hypreal_zero_def]
-      "(0::hypreal) <= r --> 0 <= r pow n";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSEs [FreeUltrafilterNat_subset, realpow_ge_zero2],
-              simpset() addsimps [hyperpow,hypreal_le]));
-qed "hyperpow_ge_zero2";
-
-Goalw [hypreal_zero_def]
-      "(0::hypreal) < x & x <= y --> x pow n <= y pow n";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addIs [realpow_le,
-    (FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset)],
-    simpset() addsimps [hyperpow,hypreal_le,hypreal_less]));
-qed_spec_mp "hyperpow_le";
-
-Goalw [hypreal_one_def] "1hr pow n = 1hr";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [hyperpow]));
-qed "hyperpow_eq_one";
-Addsimps [hyperpow_eq_one];
-
-Goalw [hypreal_one_def]
-      "abs(-(1hr pow n)) = 1hr";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), 
-      simpset() addsimps [abs_one, hyperpow,hypreal_hrabs]));
-qed "hrabs_minus_hyperpow_one";
-Addsimps [hrabs_minus_hyperpow_one];
-
-Goalw [hypreal_one_def]
-      "abs((-1hr) pow n) = 1hr";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),
-       simpset() addsimps [hyperpow,hypreal_minus,hypreal_hrabs]));
-qed "hrabs_hyperpow_minus_one";
-Addsimps [hrabs_hyperpow_minus_one];
-
-Goal "(r * s) pow n = (r pow n) * (s pow n)";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","s")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-       simpset() addsimps [hyperpow, hypreal_mult,realpow_mult]));
-qed "hyperpow_mult";
-
-Goal "(0::hypreal) <= r pow (1hn + 1hn)";
-by (simp_tac (simpset() addsimps [hyperpow_two]) 1);
-qed "hyperpow_two_le";
-Addsimps [hyperpow_two_le];
-
-Goal "abs(x pow (1hn + 1hn)) = x pow (1hn + 1hn)";
-by (simp_tac (simpset() addsimps [hrabs_eqI1]) 1);
-qed "hrabs_hyperpow_two";
-Addsimps [hrabs_hyperpow_two];
-
-Goal "abs(x) pow (1hn + 1hn)  = x pow (1hn + 1hn)";
-by (simp_tac (simpset() addsimps [hyperpow_hrabs,hrabs_eqI1]) 1);
-qed "hyperpow_two_hrabs";
-Addsimps [hyperpow_two_hrabs];
-
-Goal "!!r. 1hr < r ==> 1hr < r pow (1hn + 1hn)";
-by (auto_tac (claset(), simpset() addsimps [hyperpow_two]));
-by (cut_facts_tac [hypreal_zero_less_one] 1);
-by (forw_inst_tac [("R1.0","0")] hypreal_less_trans 1);
-by (assume_tac 1);
-by (dres_inst_tac [("z","r"),("x","1hr")] 
-    hypreal_mult_less_mono1 1);
-by (auto_tac (claset() addIs [hypreal_less_trans], simpset()));
-qed "hyperpow_two_gt_one";
-
-Goal "!!r. 1hr <= r ==> 1hr <= r pow (1hn + 1hn)";
-by (auto_tac (claset() addSDs [hypreal_le_imp_less_or_eq] 
-     addIs [hyperpow_two_gt_one,hypreal_less_imp_le],
-     simpset() addsimps [hypreal_le_refl]));
-qed "hyperpow_two_ge_one";
-
-Goalw [hypnat_one_def,hypreal_zero_def]
-      "!!r. (0::hypreal) < r ==> 0 < r pow (n + 1hn)";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSEs [FreeUltrafilterNat_subset]
-                       addDs [realpow_Suc_gt_zero], 
-              simpset() addsimps [hyperpow, hypreal_less,hypnat_add]));
-qed "hyperpow_Suc_gt_zero";
-
-Goal "!!r. (0::hypreal) <= r ==> 0 <= r pow (n + 1hn)";
-by (auto_tac (claset() addSDs [hypreal_le_imp_less_or_eq] 
-    addIs [hyperpow_ge_zero,hypreal_less_imp_le], 
-    simpset() addsimps [hypreal_le_refl]));
-qed "hyperpow_Suc_ge_zero";
-
-Goal "1hr <= (1hr +1hr) pow n";
-by (res_inst_tac [("j","1hr pow n")] hypreal_le_trans 1);
-by (rtac hyperpow_le 2);
-by (auto_tac (claset() addIs [hypreal_less_imp_le],
-    simpset() addsimps [hypreal_zero_less_one,
-    hypreal_one_less_two,hypreal_le_refl]));
-qed "two_hyperpow_ge_one";
-Addsimps [two_hyperpow_ge_one];
-
-Addsimps [simplify (simpset()) realpow_minus_one];
-Goalw [hypreal_one_def]
-      "(-1hr) pow ((1hn + 1hn)*n) = 1hr";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),
-              simpset() addsimps [hyperpow, hypnat_add,hypreal_minus]));
-qed "hyperpow_minus_one2";
-Addsimps [hyperpow_minus_one2];
-
-Goalw [hypreal_zero_def,
-      hypreal_one_def,hypnat_one_def]
-     "(0::hypreal) < r & r < 1hr --> r pow (n + 1hn) < r pow n";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSDs [conjI RS realpow_Suc_less] addEs
-    [FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset ],
-    simpset() addsimps [hyperpow,hypreal_less,hypnat_add]));
-qed_spec_mp "hyperpow_Suc_less";
-
-Goalw [hypreal_zero_def,
-      hypreal_one_def,hypnat_one_def]
-     "0 <= r & r < 1hr --> r pow (n + 1hn) <= r pow n";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSDs [conjI RS realpow_Suc_le] addEs
-    [FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset ],
-    simpset() addsimps [hyperpow,hypreal_le,hypnat_add,
-    hypreal_less]));
-qed_spec_mp "hyperpow_Suc_le";
-
-Goalw [hypreal_zero_def,
-      hypreal_one_def,hypnat_one_def]
-     "(0::hypreal) <= r & r < 1hr & n < N --> r pow N <= r pow n";
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-        simpset() addsimps [hyperpow, hypreal_le,hypreal_less,hypnat_less]));
-by (etac (FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset) 1);
-by (etac FreeUltrafilterNat_Int 1);
-by (auto_tac (claset() addSDs [conjI RS realpow_less_le],
-    simpset()));
-qed_spec_mp "hyperpow_less_le";
-
-Goal "!!r. [| (0::hypreal) <= r; r < 1hr |] ==> \
-\          ALL N n. n < N --> r pow N <= r pow n";
-by (blast_tac (claset() addSIs [hyperpow_less_le]) 1);
-qed "hyperpow_less_le2";
-
-Goal "!!r. [| 0 <= r; r < 1hr; \
-\             N : HNatInfinite \
-\          |] ==> ALL n:SHNat. r pow N <= r pow n";
-by (auto_tac (claset() addSIs [hyperpow_less_le],
-              simpset() addsimps [HNatInfinite_iff]));
-qed "hyperpow_SHNat_le";
-
-Goalw [hypreal_of_real_def,hypnat_of_nat_def] 
-      "(hypreal_of_real r) pow (hypnat_of_nat n) = hypreal_of_real (r ^ n)";
-by (auto_tac (claset(), simpset() addsimps [hyperpow]));
-qed "hyperpow_realpow";
-
-Goalw [SReal_def]
-     "(hypreal_of_real r) pow (hypnat_of_nat n) : SReal";
-by (auto_tac (claset(), simpset() addsimps [hyperpow_realpow]));
-qed "hyperpow_SReal";
-Addsimps [hyperpow_SReal];
-
-Goal "!!N. N : HNatInfinite ==> (0::hypreal) pow N = 0";
-by (dtac HNatInfinite_is_Suc 1);
-by (Auto_tac);
-qed "hyperpow_zero_HNatInfinite";
-Addsimps [hyperpow_zero_HNatInfinite];
-
-Goal "!!r. [| (0::hypreal) <= r; r < 1hr; n <= N |] ==> r pow N <= r pow n";
-by (dres_inst_tac [("y","N")] hypnat_le_imp_less_or_eq 1);
-by (auto_tac (claset() addIs [hyperpow_less_le],
-    simpset() addsimps [hypreal_le_refl]));
-qed "hyperpow_le_le";
-
-Goal "!!r. [| (0::hypreal) < r; r < 1hr |] ==> r pow (n + 1hn) <= r";
-by (dres_inst_tac [("n","1hn")] (hypreal_less_imp_le RS 
-    hyperpow_le_le) 1);
-by (Auto_tac);
-qed "hyperpow_Suc_le_self";
-
-Goal "!!r. [| (0::hypreal) <= r; r < 1hr |] ==> r pow (n + 1hn) <= r";
-by (dres_inst_tac [("n","1hn")] hyperpow_le_le 1);
-by (Auto_tac);
-qed "hyperpow_Suc_le_self2";
-
-Goalw [Infinitesimal_def]
-     "!!x. [| x : Infinitesimal; 0 < N |] \
-\          ==> (abs x) pow N <= abs x";
-by (auto_tac (claset() addSIs [hyperpow_Suc_le_self2],
-    simpset() addsimps [hyperpow_hrabs RS sym,
-    hypnat_gt_zero_iff2,hrabs_ge_zero,SReal_one,
-    hypreal_zero_less_one]));
-qed "lemma_Infinitesimal_hyperpow";
-
-Goal "!!x. [| x : Infinitesimal; 0 < N |] \
-\           ==> x pow N : Infinitesimal";
-by (rtac hrabs_le_Infinitesimal 1);
-by (dtac Infinitesimal_hrabs 1);
-by (auto_tac (claset() addSIs [lemma_Infinitesimal_hyperpow],
-    simpset() addsimps [hyperpow_hrabs RS sym]));
-qed "Infinitesimal_hyperpow";
-
-goalw HyperPow.thy [hypnat_of_nat_def] 
-     "(x ^ n : Infinitesimal) = \
-\     (x pow (hypnat_of_nat n) : Infinitesimal)";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(), simpset() addsimps [hrealpow, hyperpow]));
-qed "hrealpow_hyperpow_Infinitesimal_iff";
-
-goal HyperPow.thy 
-     "!!x. [| x : Infinitesimal; 0 < n |] \
-\           ==> x ^ n : Infinitesimal";
-by (auto_tac (claset() addSIs [Infinitesimal_hyperpow],
-    simpset() addsimps [hrealpow_hyperpow_Infinitesimal_iff,
-                        hypnat_of_nat_less_iff,hypnat_of_nat_zero] 
-              delsimps [hypnat_of_nat_less_iff RS sym]));
-qed "Infinitesimal_hrealpow";
-
-
-
-
-
-
-
--- a/src/HOL/Real/Hyperreal/HyperPow.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,26 +0,0 @@
-(*  Title       : HyperPow.thy
-    Author      : Jacques D. Fleuriot  
-    Copyright   : 1998  University of Cambridge
-    Description : Powers theory for hyperreals
-*)
-
-HyperPow = HRealAbs + HyperNat + RealPow +  
-
-instance hypreal :: {power}
-
-consts hpowr :: "[hypreal,nat] => hypreal"  
-primrec
-     hpowr_0   "r ^ 0       = 1hr"
-     hpowr_Suc "r ^ (Suc n) = (r::hypreal) * (r ^ n)"
-
-consts
-  "pow"  :: [hypreal,hypnat] => hypreal     (infixr 80)
-
-defs
-
-  (* hypernatural powers of hyperreals *)
-  hyperpow_def
-  "(R::hypreal) pow (N::hypnat) 
-      == Abs_hypreal(UN X:Rep_hypreal(R). UN Y: Rep_hypnat(N).
-             hyprel^^{%n::nat. (X n) ^ (Y n)})"
-end
--- a/src/HOL/Real/Hyperreal/Hyperreal.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-
-theory Hyperreal = HSeries:
-
-end
--- a/src/HOL/Real/Hyperreal/Lim.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2353 +0,0 @@
-(*  Title       : Lim.ML
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Theory of limits, continuity and 
-                  differentiation of real=>real functions
-*)
-
-
-fun ARITH_PROVE str = prove_goal thy str 
-                      (fn prems => [cut_facts_tac prems 1,arith_tac 1]);
-
-
-(*---------------------------------------------------------------
-   Theory of limits, continuity and differentiation of 
-   real=>real functions 
- ----------------------------------------------------------------*)
-
-Goalw [LIM_def] "(%x. k) -- x --> k";
-by Auto_tac;
-qed "LIM_const";
-Addsimps [LIM_const];
-
-(***-----------------------------------------------------------***)
-(***  Some Purely Standard Proofs - Can be used for comparison ***)
-(***-----------------------------------------------------------***)
- 
-(*--------------- 
-    LIM_add    
- ---------------*)
-Goalw [LIM_def] 
-     "[| f -- x --> l; g -- x --> m |] \
-\     ==> (%x. f(x) + g(x)) -- x --> (l + m)";
-by (Step_tac 1);
-by (REPEAT(dres_inst_tac [("x","r/#2")] spec 1));
-by (dtac (rename_numerals (real_zero_less_two RS real_inverse_gt_zero 
-          RSN (2,real_mult_less_mono2))) 1);
-by (Asm_full_simp_tac 1);
-by (Clarify_tac 1);
-by (res_inst_tac [("R1.0","s"),("R2.0","sa")] 
-    real_linear_less2 1);
-by (res_inst_tac [("x","s")] exI 1);
-by (res_inst_tac [("x","sa")] exI 2);
-by (res_inst_tac [("x","sa")] exI 3);
-by (Step_tac 1);
-by (REPEAT(dres_inst_tac [("x","xa")] spec 1) 
-    THEN step_tac (claset() addSEs [real_less_trans]) 1);
-by (REPEAT(dres_inst_tac [("x","xa")] spec 2) 
-    THEN step_tac (claset() addSEs [real_less_trans]) 2);
-by (REPEAT(dres_inst_tac [("x","xa")] spec 3) 
-    THEN step_tac (claset() addSEs [real_less_trans]) 3);
-by (ALLGOALS(rtac (abs_sum_triangle_ineq RS real_le_less_trans)));
-by (ALLGOALS(rtac (real_sum_of_halves RS subst)));
-by (auto_tac (claset() addIs [real_add_less_mono],simpset()));
-qed "LIM_add";
-
-Goalw [LIM_def] "f -- a --> L ==> (%x. -f(x)) -- a --> -L";
-by (full_simp_tac (simpset() addsimps [real_minus_add_distrib RS sym] 
-                    delsimps [real_minus_add_distrib, real_minus_minus]) 1);
-qed "LIM_minus";
-
-(*----------------------------------------------
-     LIM_add_minus
- ----------------------------------------------*)
-Goal "[| f -- x --> l; g -- x --> m |] \
-\     ==> (%x. f(x) + -g(x)) -- x --> (l + -m)";
-by (blast_tac (claset() addDs [LIM_add,LIM_minus]) 1);
-qed "LIM_add_minus";
-
-(*----------------------------------------------
-     LIM_zero
- ----------------------------------------------*)
-Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> #0";
-by (res_inst_tac [("z1","l")] (rename_numerals (real_add_minus RS subst)) 1);
-by (rtac LIM_add_minus 1 THEN Auto_tac);
-qed "LIM_zero";
-
-(*--------------------------
-   Limit not zero
- --------------------------*)
-Goalw [LIM_def] "k ~= #0 ==> ~ ((%x. k) -- x --> #0)";
-by (res_inst_tac [("R1.0","k"),("R2.0","#0")] real_linear_less2 1);
-by (auto_tac (claset(), simpset() addsimps [real_abs_def]));
-by (res_inst_tac [("x","-k")] exI 1);
-by (res_inst_tac [("x","k")] exI 2);
-by Auto_tac;
-by (ALLGOALS(dres_inst_tac [("y","s")] real_dense));
-by Safe_tac;
-by (ALLGOALS(res_inst_tac [("x","r + x")] exI));
-by Auto_tac;  
-qed "LIM_not_zero";
-
-(* [| k ~= #0; (%x. k) -- x --> #0 |] ==> R *)
-bind_thm("LIM_not_zeroE", LIM_not_zero RS notE);
-
-Goal "(%x. k) -- x --> L ==> k = L";
-by (rtac ccontr 1);
-by (dtac LIM_zero 1);
-by (rtac LIM_not_zeroE 1 THEN assume_tac 2);
-by (arith_tac 1);
-qed "LIM_const_eq";
-
-(*------------------------
-     Limit is Unique
- ------------------------*)
-Goal "[| f -- x --> L; f -- x --> M |] ==> L = M";
-by (dtac LIM_minus 1);
-by (dtac LIM_add 1 THEN assume_tac 1);
-by (auto_tac (claset() addSDs [LIM_const_eq RS sym],  simpset()));
-qed "LIM_unique";
-
-(*-------------
-    LIM_mult_zero
- -------------*)
-Goalw [LIM_def] "[| f -- x --> #0; g -- x --> #0 |] \
-\         ==> (%x. f(x)*g(x)) -- x --> #0";
-by (Step_tac 1);
-by (dres_inst_tac [("x","#1")] spec 1);
-by (dres_inst_tac [("x","r")] spec 1);
-by (cut_facts_tac [real_zero_less_one] 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [abs_mult]) 1);
-by (Clarify_tac 1);
-by (res_inst_tac [("R1.0","s"),("R2.0","sa")] 
-    real_linear_less2 1);
-by (res_inst_tac [("x","s")] exI 1);
-by (res_inst_tac [("x","sa")] exI 2);
-by (res_inst_tac [("x","sa")] exI 3);
-by (Step_tac 1);
-by (REPEAT(dres_inst_tac [("x","xa")] spec 1) 
-    THEN step_tac (claset() addSEs [real_less_trans]) 1);
-by (REPEAT(dres_inst_tac [("x","xa")] spec 2) 
-    THEN step_tac (claset() addSEs [real_less_trans]) 2);
-by (REPEAT(dres_inst_tac [("x","xa")] spec 3) 
-    THEN step_tac (claset() addSEs [real_less_trans]) 3);
-by (ALLGOALS(res_inst_tac [("t","r")] (real_mult_1 RS subst)));
-by (ALLGOALS(rtac abs_mult_less2));
-by Auto_tac;
-qed "LIM_mult_zero";
-
-Goalw [LIM_def] "(%x. x) -- a --> a";
-by Auto_tac;
-qed "LIM_self";
-
-(*--------------------------------------------------------------
-   Limits are equal for functions equal except at limit point
- --------------------------------------------------------------*)
-Goalw [LIM_def] 
-      "[| ALL x. x ~= a --> (f x = g x) |] \
-\           ==> (f -- a --> l) = (g -- a --> l)";
-by (auto_tac (claset(), simpset() addsimps [real_add_minus_iff]));
-qed "LIM_equal";
-
-Goal "[| (%x. f(x) + -g(x)) -- a --> #0; \
-\        g -- a --> l |] \
-\      ==> f -- a --> l";
-by (dtac LIM_add 1 THEN assume_tac 1);
-by (auto_tac (claset(), simpset() addsimps [real_add_assoc]));
-qed "LIM_trans";
-
-(***-------------------------------------------------------------***)
-(***           End of Purely Standard Proofs                     ***)
-(***-------------------------------------------------------------***)
-(*--------------------------------------------------------------
-       Standard and NS definitions of Limit
- --------------------------------------------------------------*)
-Goalw [LIM_def,NSLIM_def,inf_close_def] 
-      "f -- x --> L ==> f -- x --NS> L";
-by (asm_full_simp_tac
-    (simpset() addsimps [Infinitesimal_FreeUltrafilterNat_iff]) 1);
-by (Step_tac 1);
-by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-      simpset() addsimps [real_add_minus_iff, starfun, hypreal_minus, 
-                          hypreal_of_real_def, hypreal_add]));
-by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2 THEN Step_tac 1);
-by (dres_inst_tac [("x","u")] spec 1 THEN Clarify_tac 1);
-by (dres_inst_tac [("x","s")] spec 1 THEN Clarify_tac 1);
-by (subgoal_tac "ALL n::nat. (xa n) ~= x & \
-\                    abs ((xa n) + - x) < s --> abs (f (xa n) + - L) < u" 1);
-by (Blast_tac 2);
-by (dtac FreeUltrafilterNat_all 1);
-by (Ultra_tac 1);
-qed "LIM_NSLIM";
- 
-(*---------------------------------------------------------------------
-    Limit: NS definition ==> standard definition
- ---------------------------------------------------------------------*)
-
-Goal "ALL s. #0 < s --> (EX xa.  xa ~= x & \
-\        abs (xa + - x) < s  & r <= abs (f xa + -L)) \
-\     ==> ALL n::nat. EX xa.  xa ~= x & \
-\             abs(xa + -x) < inverse(real_of_posnat n) & r <= abs(f xa + -L)";
-by (Step_tac 1);
-by (cut_inst_tac [("n1","n")]
-    (real_of_posnat_gt_zero RS real_inverse_gt_zero) 1);
-by Auto_tac;
-val lemma_LIM = result();
-
-Goal "ALL s. #0 < s --> (EX xa.  xa ~= x & \
-\        abs (xa + - x) < s  & r <= abs (f xa + -L)) \
-\     ==> EX X. ALL n::nat. X n ~= x & \
-\               abs(X n + -x) < inverse(real_of_posnat n) & r <= abs(f (X n) + -L)";
-by (dtac lemma_LIM 1);
-by (dtac choice 1);
-by (Blast_tac 1);
-val lemma_skolemize_LIM2 = result();
-
-Goal "ALL n. X n ~= x & \
-\         abs (X n + - x) < inverse (real_of_posnat  n) & \
-\         r <= abs (f (X n) + - L) ==> \
-\         ALL n. abs (X n + - x) < inverse (real_of_posnat  n)";
-by (Auto_tac );
-val lemma_simp = result();
- 
-(*-------------------
-    NSLIM => LIM
- -------------------*)
-
-Goalw [LIM_def,NSLIM_def,inf_close_def] 
-      "f -- x --NS> L ==> f -- x --> L";
-by (asm_full_simp_tac
-    (simpset() addsimps [Infinitesimal_FreeUltrafilterNat_iff]) 1);
-by (EVERY1[Step_tac, rtac ccontr, Asm_full_simp_tac]);
-by (fold_tac [real_le_def]);
-by (dtac lemma_skolemize_LIM2 1);
-by (Step_tac 1);
-by (dres_inst_tac [("x","Abs_hypreal(hyprel^^{X})")] spec 1);
-by (asm_full_simp_tac
-    (simpset() addsimps [starfun, hypreal_minus, 
-                         hypreal_of_real_def,hypreal_add]) 1);
-by (Step_tac 1);
-by (dtac (lemma_simp RS real_seq_to_hypreal_Infinitesimal) 1);
-by (asm_full_simp_tac
-    (simpset() addsimps 
-       [Infinitesimal_FreeUltrafilterNat_iff,hypreal_of_real_def,
-        hypreal_minus, hypreal_add]) 1);
-by (Blast_tac 1); 
-by (rotate_tac 2 1);
-by (dres_inst_tac [("x","r")] spec 1);
-by (Clarify_tac 1);
-by (dtac FreeUltrafilterNat_all 1);
-by (Ultra_tac 1);
-qed "NSLIM_LIM";
-
-
-(**** Key result ****)
-Goal "(f -- x --> L) = (f -- x --NS> L)";
-by (blast_tac (claset() addIs [LIM_NSLIM,NSLIM_LIM]) 1);
-qed "LIM_NSLIM_iff";
-
-(*-------------------------------------------------------------------*)
-(*   Proving properties of limits using nonstandard definition and   *)
-(*   hence, the properties hold for standard limits as well          *)
-(*-------------------------------------------------------------------*)
-(*------------------------------------------------
-      NSLIM_mult and hence (trivially) LIM_mult
- ------------------------------------------------*)
-
-Goalw [NSLIM_def]
-     "[| f -- x --NS> l; g -- x --NS> m |] \
-\     ==> (%x. f(x) * g(x)) -- x --NS> (l * m)";
-by (auto_tac (claset() addSIs [inf_close_mult_HFinite],  simpset()));
-qed "NSLIM_mult";
-
-Goal "[| f -- x --> l; g -- x --> m |] \
-\     ==> (%x. f(x) * g(x)) -- x --> (l * m)";
-by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_mult]) 1);
-qed "LIM_mult2";
-
-(*----------------------------------------------
-      NSLIM_add and hence (trivially) LIM_add
-      Note the much shorter proof
- ----------------------------------------------*)
-Goalw [NSLIM_def]
-     "[| f -- x --NS> l; g -- x --NS> m |] \
-\     ==> (%x. f(x) + g(x)) -- x --NS> (l + m)";
-by (auto_tac (claset() addSIs [inf_close_add], simpset()));
-qed "NSLIM_add";
-
-Goal "[| f -- x --> l; g -- x --> m |] \
-\     ==> (%x. f(x) + g(x)) -- x --> (l + m)";
-by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_add]) 1);
-qed "LIM_add2";
-
-(*----------------------------------------------
-     NSLIM_const
- ----------------------------------------------*)
-Goalw [NSLIM_def] "(%x. k) -- x --NS> k";
-by Auto_tac;
-qed "NSLIM_const";
-
-Addsimps [NSLIM_const];
-
-Goal "(%x. k) -- x --> k";
-by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);
-qed "LIM_const2";
-
-(*----------------------------------------------
-     NSLIM_minus
- ----------------------------------------------*)
-Goalw [NSLIM_def] 
-      "f -- a --NS> L ==> (%x. -f(x)) -- a --NS> -L";
-by Auto_tac;  
-qed "NSLIM_minus";
-
-Goal "f -- a --> L ==> (%x. -f(x)) -- a --> -L";
-by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_minus]) 1);
-qed "LIM_minus2";
-
-(*----------------------------------------------
-     NSLIM_add_minus
- ----------------------------------------------*)
-Goal "[| f -- x --NS> l; g -- x --NS> m |] \
-\     ==> (%x. f(x) + -g(x)) -- x --NS> (l + -m)";
-by (blast_tac (claset() addDs [NSLIM_add,NSLIM_minus]) 1);
-qed "NSLIM_add_minus";
-
-Goal "[| f -- x --> l; g -- x --> m |] \
-\     ==> (%x. f(x) + -g(x)) -- x --> (l + -m)";
-by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
-    NSLIM_add_minus]) 1);
-qed "LIM_add_minus2";
-
-(*-----------------------------
-    NSLIM_inverse
- -----------------------------*)
-Goalw [NSLIM_def] 
-     "[| f -- a --NS> L;  L ~= #0 |] \
-\     ==> (%x. inverse(f(x))) -- a --NS> (inverse L)";
-by (Clarify_tac 1);
-by (dtac spec 1);
-by (auto_tac (claset(), 
-              simpset() addsimps [hypreal_of_real_inf_close_inverse]));  
-qed "NSLIM_inverse";
-
-Goal "[| f -- a --> L; \
-\        L ~= #0 |] ==> (%x. inverse(f(x))) -- a --> (inverse L)";
-by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_inverse]) 1);
-qed "LIM_inverse";
-
-(*------------------------------
-    NSLIM_zero
- ------------------------------*)
-Goal "f -- a --NS> l ==> (%x. f(x) + -l) -- a --NS> #0";
-by (res_inst_tac [("z1","l")] (rename_numerals (real_add_minus RS subst)) 1);
-by (rtac NSLIM_add_minus 1 THEN Auto_tac);
-qed "NSLIM_zero";
-
-Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> #0";
-by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_zero]) 1);
-qed "LIM_zero2";
-
-Goal "(%x. f(x) - l) -- x --NS> #0 ==> f -- x --NS> l";
-by (dres_inst_tac [("g","%x. l"),("m","l")] NSLIM_add 1);
-by (auto_tac (claset(),simpset() addsimps [real_diff_def, real_add_assoc]));
-qed "NSLIM_zero_cancel";
-
-Goal "(%x. f(x) - l) -- x --> #0 ==> f -- x --> l";
-by (dres_inst_tac [("g","%x. l"),("m","l")] LIM_add 1);
-by (auto_tac (claset(),simpset() addsimps [real_diff_def, real_add_assoc]));
-qed "LIM_zero_cancel";
-
-(*--------------------------
-   NSLIM_not_zero
- --------------------------*)
-Goalw [NSLIM_def] "k ~= #0 ==> ~ ((%x. k) -- x --NS> #0)";
-by Auto_tac;
-by (res_inst_tac [("x","hypreal_of_real x + ehr")] exI 1);
-by (auto_tac (claset() addIs [Infinitesimal_add_inf_close_self 
-    RS inf_close_sym],simpset()));
-by (dres_inst_tac [("x1","-hypreal_of_real x")]
-    (hypreal_add_left_cancel RS iffD2) 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym,
-                                           hypreal_epsilon_not_zero]) 1);
-qed "NSLIM_not_zero";
-
-(* [| k ~= #0; (%x. k) -- x --NS> #0 |] ==> R *)
-bind_thm("NSLIM_not_zeroE", NSLIM_not_zero RS notE);
-
-Goal "k ~= #0 ==> ~ ((%x. k) -- x --> #0)";
-by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_not_zero]) 1);
-qed "LIM_not_zero2";
-
-(*-------------------------------------
-   NSLIM of constant function
- -------------------------------------*)
-Goal "(%x. k) -- x --NS> L ==> k = L";
-by (rtac ccontr 1);
-by (dtac NSLIM_zero 1);
-by (rtac NSLIM_not_zeroE 1 THEN assume_tac 2);
-by (arith_tac 1);
-qed "NSLIM_const_eq";
-
-Goal "(%x. k) -- x --> L ==> k = L";
-by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
-    NSLIM_const_eq]) 1);
-qed "LIM_const_eq2";
-
-(*------------------------
-     NS Limit is Unique
- ------------------------*)
-(* can actually be proved more easily by unfolding def! *)
-Goal "[| f -- x --NS> L; f -- x --NS> M |] ==> L = M";
-by (dtac NSLIM_minus 1);
-by (dtac NSLIM_add 1 THEN assume_tac 1);
-by (auto_tac (claset() addSDs [NSLIM_const_eq RS sym], simpset()));
-qed "NSLIM_unique";
-
-Goal "[| f -- x --> L; f -- x --> M |] ==> L = M";
-by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_unique]) 1);
-qed "LIM_unique2";
-
-(*--------------------
-    NSLIM_mult_zero
- --------------------*)
-Goal "[| f -- x --NS> #0; g -- x --NS> #0 |] \
-\         ==> (%x. f(x)*g(x)) -- x --NS> #0";
-by (dtac NSLIM_mult 1 THEN Auto_tac);
-qed "NSLIM_mult_zero";
-
-(* we can use the corresponding thm LIM_mult2 *)
-(* for standard definition of limit           *)
-
-Goal "[| f -- x --> #0; g -- x --> #0 |] \
-\     ==> (%x. f(x)*g(x)) -- x --> #0";
-by (dtac LIM_mult2 1 THEN Auto_tac);
-qed "LIM_mult_zero2";
-
-(*----------------------------
-    NSLIM_self
- ----------------------------*)
-Goalw [NSLIM_def] "(%x. x) -- a --NS> a";
-by (auto_tac (claset() addIs [starfun_Idfun_inf_close],simpset()));
-qed "NSLIM_self";
-
-Goal "(%x. x) -- a --> a";
-by (simp_tac (simpset() addsimps [LIM_NSLIM_iff,NSLIM_self]) 1);
-qed "LIM_self2";
-
-(*-----------------------------------------------------------------------------
-   Derivatives and Continuity - NS and Standard properties
- -----------------------------------------------------------------------------*)
-(*---------------
-    Continuity 
- ---------------*)
-
-Goalw [isNSCont_def] 
-      "[| isNSCont f a; y @= hypreal_of_real a |] \
-\           ==> (*f* f) y @= hypreal_of_real (f a)";
-by (Blast_tac 1);
-qed "isNSContD";
-
-Goalw [isNSCont_def,NSLIM_def] 
-      "isNSCont f a ==> f -- a --NS> (f a) ";
-by (Blast_tac 1);
-qed "isNSCont_NSLIM";
-
-Goalw [isNSCont_def,NSLIM_def] 
-      "f -- a --NS> (f a) ==> isNSCont f a";
-by Auto_tac;
-by (res_inst_tac [("Q","y = hypreal_of_real a")] 
-    (excluded_middle RS disjE) 1);
-by Auto_tac;
-qed "NSLIM_isNSCont";
-
-(*-----------------------------------------------------
-    NS continuity can be defined using NS Limit in
-    similar fashion to standard def of continuity
- -----------------------------------------------------*)
-Goal "(isNSCont f a) = (f -- a --NS> (f a))";
-by (blast_tac (claset() addIs [isNSCont_NSLIM,NSLIM_isNSCont]) 1);
-qed "isNSCont_NSLIM_iff";
-
-(*----------------------------------------------
-  Hence, NS continuity can be given
-  in terms of standard limit
- ---------------------------------------------*)
-Goal "(isNSCont f a) = (f -- a --> (f a))";
-by (asm_full_simp_tac (simpset() addsimps 
-    [LIM_NSLIM_iff,isNSCont_NSLIM_iff]) 1);
-qed "isNSCont_LIM_iff";
-
-(*-----------------------------------------------
-  Moreover, it's trivial now that NS continuity 
-  is equivalent to standard continuity
- -----------------------------------------------*)
-Goalw [isCont_def] "(isNSCont f a) = (isCont f a)";
-by (rtac isNSCont_LIM_iff 1);
-qed "isNSCont_isCont_iff";
-
-(*----------------------------------------
-  Standard continuity ==> NS continuity 
- ----------------------------------------*)
-Goal "isCont f a ==> isNSCont f a";
-by (etac (isNSCont_isCont_iff RS iffD2) 1);
-qed "isCont_isNSCont";
-
-(*----------------------------------------
-  NS continuity ==> Standard continuity 
- ----------------------------------------*)
-Goal "isNSCont f a ==> isCont f a";
-by (etac (isNSCont_isCont_iff RS iffD1) 1);
-qed "isNSCont_isCont";
-
-(*--------------------------------------------------------------------------
-                 Alternative definition of continuity
- --------------------------------------------------------------------------*)
-(* Prove equivalence between NS limits - *)
-(* seems easier than using standard def  *)
-Goalw [NSLIM_def] "(f -- a --NS> L) = ((%h. f(a + h)) -- #0 --NS> L)";
-by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_zero]));
-by (dres_inst_tac [("x","hypreal_of_real a + x")] spec 1);
-by (dres_inst_tac [("x","-hypreal_of_real a + x")] spec 2);
-by (Step_tac 1);
-by (Asm_full_simp_tac 1);
-by (rtac ((mem_infmal_iff RS iffD2) RS 
-    (Infinitesimal_add_inf_close_self RS inf_close_sym)) 1);
-by (rtac (inf_close_minus_iff2 RS iffD1) 4);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 3);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 2);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 4);
-by (auto_tac (claset(),
-       simpset() addsimps [starfun, hypreal_of_real_def, hypreal_minus,
-              hypreal_add, real_add_assoc, inf_close_refl, hypreal_zero_def]));
-qed "NSLIM_h_iff";
-
-Goal "(f -- a --NS> f a) = ((%h. f(a + h)) -- #0 --NS> f a)";
-by (rtac NSLIM_h_iff 1);
-qed "NSLIM_isCont_iff";
-
-Goal "(f -- a --> f a) = ((%h. f(a + h)) -- #0 --> f(a))";
-by (simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_isCont_iff]) 1);
-qed "LIM_isCont_iff";
-
-Goalw [isCont_def] "(isCont f x) = ((%h. f(x + h)) -- #0 --> f(x))";
-by (simp_tac (simpset() addsimps [LIM_isCont_iff]) 1);
-qed "isCont_iff";
-
-(*--------------------------------------------------------------------------
-   Immediate application of nonstandard criterion for continuity can offer 
-   very simple proofs of some standard property of continuous functions
- --------------------------------------------------------------------------*)
-(*------------------------
-     sum continuous
- ------------------------*)
-Goal "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) + g(x)) a";
-by (auto_tac (claset() addIs [inf_close_add],
-              simpset() addsimps [isNSCont_isCont_iff RS sym, isNSCont_def]));
-qed "isCont_add";
-
-(*------------------------
-     mult continuous
- ------------------------*)
-Goal "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) * g(x)) a";
-by (auto_tac (claset() addSIs [starfun_mult_HFinite_inf_close],
-              simpset() delsimps [starfun_mult RS sym]
-			addsimps [isNSCont_isCont_iff RS sym, isNSCont_def]));
-qed "isCont_mult";
-
-(*-------------------------------------------
-     composition of continuous functions
-     Note very short straightforard proof!
- ------------------------------------------*)
-Goal "[| isCont f a; isCont g (f a) |] \
-\     ==> isCont (g o f) a";
-by (auto_tac (claset(),simpset() addsimps [isNSCont_isCont_iff RS sym,
-              isNSCont_def,starfun_o RS sym]));
-qed "isCont_o";
-
-Goal "[| isCont f a; isCont g (f a) |] \
-\     ==> isCont (%x. g (f x)) a";
-by (auto_tac (claset() addDs [isCont_o],simpset() addsimps [o_def]));
-qed "isCont_o2";
-
-Goalw [isNSCont_def] "isNSCont f a ==> isNSCont (%x. - f x) a";
-by Auto_tac; 
-qed "isNSCont_minus";
-
-Goal "isCont f a ==> isCont (%x. - f x) a";
-by (auto_tac (claset(),simpset() addsimps [isNSCont_isCont_iff RS sym,
-              isNSCont_minus]));
-qed "isCont_minus";
-
-Goalw [isCont_def]  
-      "[| isCont f x; f x ~= #0 |] ==> isCont (%x. inverse (f x)) x";
-by (blast_tac (claset() addIs [LIM_inverse]) 1);
-qed "isCont_inverse";
-
-Goal "[| isNSCont f x; f x ~= #0 |] ==> isNSCont (%x. inverse (f x)) x";
-by (auto_tac (claset() addIs [isCont_inverse],simpset() addsimps 
-    [isNSCont_isCont_iff]));
-qed "isNSCont_inverse";
-
-Goalw [real_diff_def] 
-      "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) - g(x)) a";
-by (auto_tac (claset() addIs [isCont_add,isCont_minus],simpset()));
-qed "isCont_diff";
-
-Goalw [isCont_def]  "isCont (%x. k) a";
-by (Simp_tac 1);
-qed "isCont_const";
-Addsimps [isCont_const];
-
-Goalw [isNSCont_def]  "isNSCont (%x. k) a";
-by (Simp_tac 1);
-qed "isNSCont_const";
-Addsimps [isNSCont_const];
-
-Goalw [isNSCont_def]  "isNSCont abs a";
-by (auto_tac (claset() addIs [inf_close_hrabs],
-              simpset() addsimps [hypreal_of_real_hrabs RS sym,
-                                  starfun_rabs_hrabs]));
-qed "isNSCont_rabs";
-Addsimps [isNSCont_rabs];
-
-Goal "isCont abs a";
-by (auto_tac (claset(), simpset() addsimps [isNSCont_isCont_iff RS sym]));
-qed "isCont_rabs";
-Addsimps [isCont_rabs];
-
-(****************************************************************
-(%* Leave as commented until I add topology theory or remove? *%)
-(%*------------------------------------------------------------
-  Elementary topology proof for a characterisation of 
-  continuity now: a function f is continuous if and only 
-  if the inverse image, {x. f(x) : A}, of any open set A 
-  is always an open set
- ------------------------------------------------------------*%)
-Goal "[| isNSopen A; ALL x. isNSCont f x |] \
-\              ==> isNSopen {x. f x : A}";
-by (auto_tac (claset(),simpset() addsimps [isNSopen_iff1]));
-by (dtac (mem_monad_inf_close RS inf_close_sym) 1);
-by (dres_inst_tac [("x","a")] spec 1);
-by (dtac isNSContD 1 THEN assume_tac 1);
-by (dtac bspec 1 THEN assume_tac 1);
-by (dres_inst_tac [("x","( *f* f) x")] inf_close_mem_monad2 1);
-by (blast_tac (claset() addIs [starfun_mem_starset]) 1);
-qed "isNSCont_isNSopen";
-
-Goalw [isNSCont_def]
-          "ALL A. isNSopen A --> isNSopen {x. f x : A} \
-\              ==> isNSCont f x";
-by (auto_tac (claset() addSIs [(mem_infmal_iff RS iffD1) RS 
-     (inf_close_minus_iff RS iffD2)],simpset() addsimps 
-      [Infinitesimal_def,SReal_iff]));
-by (dres_inst_tac [("x","{z. abs(z + -f(x)) < ya}")] spec 1);
-by (etac (isNSopen_open_interval RSN (2,impE)) 1);
-by (auto_tac (claset(),simpset() addsimps [isNSopen_def,isNSnbhd_def]));
-by (dres_inst_tac [("x","x")] spec 1);
-by (auto_tac (claset() addDs [inf_close_sym RS inf_close_mem_monad],
-    simpset() addsimps [hypreal_of_real_zero RS sym,STAR_starfun_rabs_add_minus]));
-qed "isNSopen_isNSCont";
-
-Goal "(ALL x. isNSCont f x) = \
-\     (ALL A. isNSopen A --> isNSopen {x. f(x) : A})";
-by (blast_tac (claset() addIs [isNSCont_isNSopen,
-    isNSopen_isNSCont]) 1);
-qed "isNSCont_isNSopen_iff";
-
-(%*------- Standard version of same theorem --------*%)
-Goal "(ALL x. isCont f x) = \
-\         (ALL A. isopen A --> isopen {x. f(x) : A})";
-by (auto_tac (claset() addSIs [isNSCont_isNSopen_iff],
-              simpset() addsimps [isNSopen_isopen_iff RS sym,
-              isNSCont_isCont_iff RS sym]));
-qed "isCont_isopen_iff";
-*******************************************************************)
-
-(*-----------------------------------------------------------------
-                        Uniform continuity
- ------------------------------------------------------------------*)
-Goalw [isNSUCont_def] 
-      "[| isNSUCont f; x @= y|] ==> (*f* f) x @= (*f* f) y";
-by (Blast_tac 1);
-qed "isNSUContD";
-
-Goalw [isUCont_def,isCont_def,LIM_def]
-     "isUCont f ==> EX x. isCont f x";
-by (Force_tac 1);
-qed "isUCont_isCont";
-
-Goalw [isNSUCont_def,isUCont_def,inf_close_def] 
-     "isUCont f ==> isNSUCont f";
-by (asm_full_simp_tac (simpset() addsimps 
-    [Infinitesimal_FreeUltrafilterNat_iff]) 1);
-by (Step_tac 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [starfun,
-    hypreal_minus, hypreal_add]));
-by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2 THEN Step_tac 1);
-by (dres_inst_tac [("x","u")] spec 1 THEN Clarify_tac 1);
-by (dres_inst_tac [("x","s")] spec 1 THEN Clarify_tac 1);
-by (subgoal_tac "ALL n::nat. abs ((xa n) + - (xb n)) < s --> abs (f (xa n) + - f (xb n)) < u" 1);
-by (Blast_tac 2);
-by (thin_tac "ALL x y. abs (x + - y) < s --> abs (f x + - f y) < u" 1);
-by (dtac FreeUltrafilterNat_all 1);
-by (Ultra_tac 1);
-qed "isUCont_isNSUCont";
-
-Goal "ALL s. #0 < s --> (EX z y. abs (z + - y) < s & r <= abs (f z + -f y)) \
-\     ==> ALL n::nat. EX z y.  \
-\              abs(z + -y) < inverse(real_of_posnat n) & \
-\              r <= abs(f z + -f y)";
-by (Step_tac 1);
-by (cut_inst_tac [("n1","n")] (real_of_posnat_gt_zero RS real_inverse_gt_zero) 1);
-by Auto_tac;
-val lemma_LIMu = result();
-
-Goal "ALL s. #0 < s --> (EX z y. abs (z + - y) < s  & r <= abs (f z + -f y)) \
-\     ==> EX X Y. ALL n::nat. \
-\              abs(X n + -(Y n)) < inverse(real_of_posnat n) & \
-\              r <= abs(f (X n) + -f (Y n))";
-by (dtac lemma_LIMu 1);
-by (dtac choice 1);
-by (Step_tac 1);
-by (dtac choice 1);
-by (Blast_tac 1);
-val lemma_skolemize_LIM2u = result();
-
-Goal "ALL n. abs (X n + -Y n) < inverse (real_of_posnat  n) & \
-\         r <= abs (f (X n) + - f(Y n)) ==> \
-\         ALL n. abs (X n + - Y n) < inverse (real_of_posnat  n)";
-by (Auto_tac );
-val lemma_simpu = result();
-
-Goalw [isNSUCont_def,isUCont_def,inf_close_def] 
-     "isNSUCont f ==> isUCont f";
-by (asm_full_simp_tac (simpset() addsimps 
-                       [Infinitesimal_FreeUltrafilterNat_iff]) 1);
-by (EVERY1[Step_tac, rtac ccontr, Asm_full_simp_tac]);
-by (fold_tac [real_le_def]);
-by (dtac lemma_skolemize_LIM2u 1);
-by (Step_tac 1);
-by (dres_inst_tac [("x","Abs_hypreal(hyprel^^{X})")] spec 1);
-by (dres_inst_tac [("x","Abs_hypreal(hyprel^^{Y})")] spec 1);
-by (asm_full_simp_tac
-    (simpset() addsimps [starfun, hypreal_minus,hypreal_add]) 1);
-by Auto_tac;
-by (dtac (lemma_simpu RS real_seq_to_hypreal_Infinitesimal2) 1);
-by (asm_full_simp_tac (simpset() addsimps 
-     [Infinitesimal_FreeUltrafilterNat_iff, hypreal_minus,hypreal_add]) 1);
-by (Blast_tac 1);
-by (rotate_tac 2 1);
-by (dres_inst_tac [("x","r")] spec 1);
-by (Clarify_tac 1);
-by (dtac FreeUltrafilterNat_all 1);
-by (Ultra_tac 1);
-qed "isNSUCont_isUCont";
-
-(*------------------------------------------------------------------
-                         Derivatives
- ------------------------------------------------------------------*)
-Goalw [deriv_def] 
-      "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- #0 --> D)";
-by (Blast_tac 1);        
-qed "DERIV_iff";
-
-Goalw [deriv_def] 
-      "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- #0 --NS> D)";
-by (simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);
-qed "DERIV_NS_iff";
-
-Goalw [deriv_def] 
-      "DERIV f x :> D \
-\      ==> (%h. (f(x + h) + - f(x))/h) -- #0 --> D";
-by (Blast_tac 1);        
-qed "DERIVD";
-
-Goalw [deriv_def] "DERIV f x :> D ==> \
-\          (%h. (f(x + h) + - f(x))/h) -- #0 --NS> D";
-by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);
-qed "NS_DERIVD";
-
-(* Uniqueness *)
-Goalw [deriv_def] 
-      "[| DERIV f x :> D; DERIV f x :> E |] ==> D = E";
-by (blast_tac (claset() addIs [LIM_unique]) 1);
-qed "DERIV_unique";
-
-Goalw [nsderiv_def] 
-     "[| NSDERIV f x :> D; NSDERIV f x :> E |] ==> D = E";
-by (cut_facts_tac [Infinitesimal_epsilon, hypreal_epsilon_not_zero] 1);
-by (auto_tac (claset() addSDs [inst "x" "ehr" bspec] 
-                       addSIs [inj_hypreal_of_real RS injD] 
-                       addDs [inf_close_trans3],
-              simpset()));
-qed "NSDeriv_unique";
-
-(*------------------------------------------------------------------------
-                          Differentiable
- ------------------------------------------------------------------------*)
-
-Goalw [differentiable_def] 
-      "f differentiable x ==> EX D. DERIV f x :> D";
-by (assume_tac 1);
-qed "differentiableD";
-
-Goalw [differentiable_def] 
-      "DERIV f x :> D ==> f differentiable x";
-by (Blast_tac 1);
-qed "differentiableI";
-
-Goalw [NSdifferentiable_def] 
-      "f NSdifferentiable x ==> EX D. NSDERIV f x :> D";
-by (assume_tac 1);
-qed "NSdifferentiableD";
-
-Goalw [NSdifferentiable_def] 
-      "NSDERIV f x :> D ==> f NSdifferentiable x";
-by (Blast_tac 1);
-qed "NSdifferentiableI";
-
-(*--------------------------------------------------------
-      Alternative definition for differentiability
- -------------------------------------------------------*)
-
-Goalw [LIM_def] 
- "((%h. (f(a + h) + - f(a))/h) -- #0 --> D) = \
-\ ((%x. (f(x) + -f(a)) / (x + -a)) -- a --> D)";
-by (Step_tac 1);
-by (ALLGOALS(dtac spec));
-by (Step_tac 1);
-by (Blast_tac 1 THEN Blast_tac 2);
-by (ALLGOALS(res_inst_tac [("x","s")] exI));
-by (Step_tac 1);
-by (dres_inst_tac [("x","x + -a")] spec 1);
-by (dres_inst_tac [("x","x + a")] spec 2);
-by (auto_tac (claset(), simpset() addsimps real_add_ac));
-qed "DERIV_LIM_iff";
-
-Goalw [deriv_def] "(DERIV f x :> D) = \
-\         ((%z. (f(z) + -f(x)) / (z + -x)) -- x --> D)";
-by (simp_tac (simpset() addsimps [DERIV_LIM_iff]) 1);
-qed "DERIV_iff2";
-
-(*--------------------------------------------------------
-  Equivalence of NS and standard defs of differentiation
- -------------------------------------------------------*)
-(*-------------------------------------------
-   First NSDERIV in terms of NSLIM 
- -------------------------------------------*)
-
-(*--- first equivalence ---*)
-Goalw [nsderiv_def,NSLIM_def] 
-      "(NSDERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- #0 --NS> D)";
-by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_zero]));
-by (dres_inst_tac [("x","xa")] bspec 1);
-by (rtac ccontr 3);
-by (dres_inst_tac [("x","h")] spec 3);
-by (auto_tac (claset(),
-              simpset() addsimps [mem_infmal_iff, starfun_lambda_cancel]));
-qed "NSDERIV_NSLIM_iff";
-
-(*--- second equivalence ---*)
-Goal "(NSDERIV f x :> D) = \
-\         ((%z. (f(z) + -f(x)) / (z + -x)) -- x --NS> D)";
-by (full_simp_tac (simpset() addsimps 
-     [NSDERIV_NSLIM_iff, DERIV_LIM_iff, LIM_NSLIM_iff RS sym]) 1);
-qed "NSDERIV_NSLIM_iff2";
-
-(* while we're at it! *)
-Goalw [real_diff_def]
-     "(NSDERIV f x :> D) = \
-\     (ALL xa. \
-\       xa ~= hypreal_of_real x & xa @= hypreal_of_real x --> \
-\       (*f* (%z. (f z - f x) / (z - x))) xa @= hypreal_of_real D)";
-by (auto_tac (claset(), simpset() addsimps [NSDERIV_NSLIM_iff2, NSLIM_def]));
-qed "NSDERIV_iff2";
-
-Addsimps [inf_close_refl];
-
-
-(*FIXME: replace by simprocs for cancellation of common factors*)
-Goal "h ~= (0::hypreal) ==> x*h/h = x";
-by (asm_simp_tac 
-    (simpset() addsimps [hypreal_divide_def, hypreal_mult_inverse_left, 
-                         hypreal_mult_assoc]) 1);
-qed "hypreal_divide_mult_self_eq"; 
-Addsimps [hypreal_divide_mult_self_eq];
-
-(*FIXME: replace by simprocs for cancellation of common factors*)
-Goal "h ~= (0::hypreal) ==> (h*x)/h = x";
-by (asm_simp_tac 
-    (simpset() addsimps [inst "z" "h" hypreal_mult_commute]) 1);
-qed "hypreal_times_divide_self_eq"; 
-Addsimps [hypreal_times_divide_self_eq];
-
-(*FIXME: replace by simprocs for cancellation of common factors*)
-Goal "h ~= (0::hypreal) ==> h/h = 1hr";
-by (asm_simp_tac 
-    (simpset() addsimps [hypreal_divide_def, hypreal_mult_inverse_left]) 1);
-qed "hypreal_divide_self_eq"; 
-Addsimps [hypreal_divide_self_eq];
-
-
-Goal "(NSDERIV f x :> D) ==> \
-\    (ALL xa. \
-\       xa @= hypreal_of_real x --> \
-\       (*f* (%z. f z - f x)) xa @= hypreal_of_real D * (xa - hypreal_of_real x))";
-by (auto_tac (claset(), simpset() addsimps [NSDERIV_iff2]));
-by (case_tac "xa = hypreal_of_real x" 1);
-by (auto_tac (claset(), 
-              simpset() addsimps [hypreal_diff_def, hypreal_of_real_zero]));
-by (dres_inst_tac [("x","xa")] spec 1);
-by Auto_tac;
-by (dres_inst_tac [("c","xa - hypreal_of_real x"),("b","hypreal_of_real D")]
-     inf_close_mult1 1);
-by (ALLGOALS(dtac (hypreal_not_eq_minus_iff RS iffD1)));
-by (subgoal_tac "(*f* (%z. z - x)) xa ~= (0::hypreal)" 2);
-by (rotate_tac ~1 2);
-by (auto_tac (claset(),
-    simpset() addsimps [real_diff_def, hypreal_diff_def, 
-		(inf_close_minus_iff RS iffD1) RS (mem_infmal_iff RS iffD2),  
-			Infinitesimal_subset_HFinite RS subsetD]));
-qed "NSDERIVD5";
-
-Goal "(NSDERIV f x :> D) ==> \
-\     (ALL h: Infinitesimal. \
-\              ((*f* f)(hypreal_of_real x + h) - \
-\                hypreal_of_real (f x))@= (hypreal_of_real D) * h)";
-by (auto_tac (claset(),simpset() addsimps [nsderiv_def]));
-by (case_tac "h = (0::hypreal)" 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_diff_def]));
-by (dres_inst_tac [("x","h")] bspec 1);
-by (dres_inst_tac [("c","h")] inf_close_mult1 2);
-by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],
-              simpset() addsimps [hypreal_diff_def]));
-qed "NSDERIVD4";
-
-Goal "(NSDERIV f x :> D) ==> \
-\     (ALL h: Infinitesimal - {0}. \
-\              ((*f* f)(hypreal_of_real x + h) - \
-\                hypreal_of_real (f x))@= (hypreal_of_real D) * h)";
-by (auto_tac (claset(),simpset() addsimps [nsderiv_def]));
-by (rtac ccontr 1 THEN dres_inst_tac [("x","h")] bspec 1);
-by (dres_inst_tac [("c","h")] inf_close_mult1 2);
-by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],
-              simpset() addsimps [hypreal_mult_assoc, hypreal_diff_def]));
-qed "NSDERIVD3";
-
-(*--------------------------------------------------------------
-          Now equivalence between NSDERIV and DERIV
- -------------------------------------------------------------*)
-Goalw [deriv_def] "(NSDERIV f x :> D) = (DERIV f x :> D)";
-by (simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,LIM_NSLIM_iff]) 1);
-qed "NSDERIV_DERIV_iff";
-
-(*---------------------------------------------------
-         Differentiability implies continuity 
-         nice and simple "algebraic" proof
- --------------------------------------------------*)
-Goalw [nsderiv_def]
-      "NSDERIV f x :> D ==> isNSCont f x";
-by (auto_tac (claset(),simpset() addsimps 
-        [isNSCont_NSLIM_iff,NSLIM_def]));
-by (dtac (inf_close_minus_iff RS iffD1) 1);
-by (dtac (hypreal_not_eq_minus_iff RS iffD1) 1);
-by (dres_inst_tac [("x","-hypreal_of_real x + xa")] bspec 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypreal_add_assoc RS sym]) 2);
-by (auto_tac (claset(),simpset() addsimps 
-    [mem_infmal_iff RS sym,hypreal_add_commute]));
-by (dres_inst_tac [("c","xa + -hypreal_of_real x")] inf_close_mult1 1);
-by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite
-    RS subsetD],simpset() addsimps [hypreal_mult_assoc]));
-by (dres_inst_tac [("x3","D")] (HFinite_hypreal_of_real RSN
-    (2,Infinitesimal_HFinite_mult) RS (mem_infmal_iff RS iffD1)) 1);
-by (blast_tac (claset() addIs [inf_close_trans,
-    hypreal_mult_commute RS subst,
-    (inf_close_minus_iff RS iffD2)]) 1);
-qed "NSDERIV_isNSCont";
-
-(* Now Sandard proof *)
-Goal "DERIV f x :> D ==> isCont f x";
-by (asm_full_simp_tac (simpset() addsimps 
-    [NSDERIV_DERIV_iff RS sym, isNSCont_isCont_iff RS sym,
-     NSDERIV_isNSCont]) 1);
-qed "DERIV_isCont";
-
-(*----------------------------------------------------------------------------
-      Differentiation rules for combinations of functions
-      follow from clear, straightforard, algebraic 
-      manipulations
- ----------------------------------------------------------------------------*)
-(*-------------------------
-    Constant function
- ------------------------*)
-
-(* use simple constant nslimit theorem *)
-Goal "(NSDERIV (%x. k) x :> #0)";
-by (simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff]) 1);
-qed "NSDERIV_const";
-Addsimps [NSDERIV_const];
-
-Goal "(DERIV (%x. k) x :> #0)";
-by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1);
-qed "DERIV_const";
-Addsimps [DERIV_const];
-
-(*-----------------------------------------------------
-    Sum of functions- proved easily
- ----------------------------------------------------*)
-
-
-Goal "[| NSDERIV f x :> Da;  NSDERIV g x :> Db |] \
-\     ==> NSDERIV (%x. f x + g x) x :> Da + Db";
-by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,
-           NSLIM_def]) 1 THEN REPEAT(Step_tac 1));
-by (auto_tac (claset(),
-       simpset() addsimps [hypreal_add_divide_distrib]));
-by (thin_tac "xa @= hypreal_of_real #0" 1 THEN dtac inf_close_add 1);
-by (auto_tac (claset(),simpset() addsimps hypreal_add_ac));
-qed "NSDERIV_add";
-
-(* Standard theorem *)
-Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \
-\     ==> DERIV (%x. f x + g x) x :> Da + Db";
-by (asm_full_simp_tac (simpset() addsimps [NSDERIV_add,
-    NSDERIV_DERIV_iff RS sym]) 1);
-qed "DERIV_add";
-
-(*-----------------------------------------------------
-  Product of functions - Proof is trivial but tedious
-  and long due to rearrangement of terms  
- ----------------------------------------------------*)
-(* lemma - terms manipulation tedious as usual*)
-
-Goal "((a::hypreal)*b) + -(c*d) = (b*(a + -c)) + \
-\                           (c*(b + -d))";
-by (full_simp_tac (simpset() addsimps [hypreal_add_mult_distrib2,
-    hypreal_minus_mult_eq2 RS sym,hypreal_mult_commute]) 1);
-val lemma_nsderiv1 = result();
-
-Goal "[| (x + y) / z = hypreal_of_real D + yb; z ~= 0; \
-\        z : Infinitesimal; yb : Infinitesimal |] \
-\     ==> x + y @= 0";
-by (forw_inst_tac [("c1","z")] (hypreal_mult_right_cancel RS iffD2) 1 
-    THEN assume_tac 1);
-by (thin_tac "(x + y) / z = hypreal_of_real  D + yb" 1);
-by (auto_tac (claset() addSIs [Infinitesimal_HFinite_mult2, HFinite_add],
-              simpset() addsimps [hypreal_mult_assoc, mem_infmal_iff RS sym]));
-by (etac (Infinitesimal_subset_HFinite RS subsetD) 1);
-val lemma_nsderiv2 = result();
-
-
-Goal "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \
-\     ==> NSDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))";
-by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff, NSLIM_def]) 1 
-    THEN REPEAT(Step_tac 1));
-by (auto_tac (claset(),
-       simpset() addsimps [starfun_lambda_cancel, hypreal_of_real_zero,
-              lemma_nsderiv1]));
-by (simp_tac (simpset() addsimps [hypreal_add_divide_distrib]) 1); 
-by (REPEAT(dtac (bex_Infinitesimal_iff2 RS iffD2) 1));
-by (auto_tac (claset(),
-        simpset() delsimps [hypreal_times_divide1_eq]
-		  addsimps [hypreal_times_divide1_eq RS sym]));
-by (dres_inst_tac [("D","Db")] lemma_nsderiv2 1);
-by (dtac (inf_close_minus_iff RS iffD2 RS (bex_Infinitesimal_iff2 RS iffD2)) 4);
-by (auto_tac (claset() addSIs [inf_close_add_mono1],
-      simpset() addsimps [hypreal_add_mult_distrib, hypreal_add_mult_distrib2, 
-			  hypreal_mult_commute, hypreal_add_assoc]));
-by (res_inst_tac [("w1","hypreal_of_real Db * hypreal_of_real (f x)")]
-    (hypreal_add_commute RS subst) 1);
-by (auto_tac (claset() addSIs [Infinitesimal_add_inf_close_self2 RS inf_close_sym,
-			       Infinitesimal_add, Infinitesimal_mult,
-			       Infinitesimal_hypreal_of_real_mult,
-			       Infinitesimal_hypreal_of_real_mult2],
-	      simpset() addsimps [hypreal_add_assoc RS sym]));
-qed "NSDERIV_mult";
-
-Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \
-\     ==> DERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))";
-by (asm_full_simp_tac (simpset() addsimps [NSDERIV_mult,
-                                           NSDERIV_DERIV_iff RS sym]) 1);
-qed "DERIV_mult";
-
-(*----------------------------
-   Multiplying by a constant
- ---------------------------*)
-Goal "NSDERIV f x :> D \
-\     ==> NSDERIV (%x. c * f x) x :> c*D";
-by (asm_full_simp_tac 
-    (simpset() addsimps [real_times_divide1_eq RS sym, NSDERIV_NSLIM_iff,
-                         real_minus_mult_eq2, real_add_mult_distrib2 RS sym] 
-             delsimps [real_times_divide1_eq, real_minus_mult_eq2 RS sym]) 1);
-by (etac (NSLIM_const RS NSLIM_mult) 1);
-qed "NSDERIV_cmult";
-
-(* let's do the standard proof though theorem *)
-(* LIM_mult2 follows from a NS proof          *)
-
-Goalw [deriv_def] 
-      "DERIV f x :> D \
-\      ==> DERIV (%x. c * f x) x :> c*D";
-by (asm_full_simp_tac 
-    (simpset() addsimps [real_times_divide1_eq RS sym, NSDERIV_NSLIM_iff,
-                         real_minus_mult_eq2, real_add_mult_distrib2 RS sym] 
-             delsimps [real_times_divide1_eq, real_minus_mult_eq2 RS sym]) 1);
-by (etac (LIM_const RS LIM_mult2) 1);
-qed "DERIV_cmult";
-
-(*--------------------------------
-   Negation of function
- -------------------------------*)
-Goal "NSDERIV f x :> D ==> NSDERIV (%x. -(f x)) x :> -D";
-by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff]) 1);
-by (res_inst_tac [("t","f x")] (real_minus_minus RS subst) 1);
-by (asm_simp_tac (simpset() addsimps [real_minus_add_distrib RS sym,
-                                      real_minus_mult_eq1 RS sym] 
-                   delsimps [real_minus_add_distrib, real_minus_minus]) 1);
-by (etac NSLIM_minus 1);
-qed "NSDERIV_minus";
-
-Goal "DERIV f x :> D \
-\     ==> DERIV (%x. -(f x)) x :> -D";
-by (asm_full_simp_tac (simpset() addsimps 
-    [NSDERIV_minus,NSDERIV_DERIV_iff RS sym]) 1);
-qed "DERIV_minus";
-
-(*-------------------------------
-   Subtraction
- ------------------------------*)
-Goal "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \
-\     ==> NSDERIV (%x. f x + -g x) x :> Da + -Db";
-by (blast_tac (claset() addDs [NSDERIV_add,NSDERIV_minus]) 1);
-qed "NSDERIV_add_minus";
-
-Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \
-\     ==> DERIV (%x. f x + -g x) x :> Da + -Db";
-by (blast_tac (claset() addDs [DERIV_add,DERIV_minus]) 1);
-qed "DERIV_add_minus";
-
-Goalw [real_diff_def]
-     "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \
-\     ==> NSDERIV (%x. f x - g x) x :> Da - Db";
-by (blast_tac (claset() addIs [NSDERIV_add_minus]) 1);
-qed "NSDERIV_diff";
-
-Goalw [real_diff_def]
-     "[| DERIV f x :> Da; DERIV g x :> Db |] \
-\      ==> DERIV (%x. f x - g x) x :> Da - Db";
-by (blast_tac (claset() addIs [DERIV_add_minus]) 1);
-qed "DERIV_diff";
-
-(*---------------------------------------------------------------
-                     (NS) Increment
- ---------------------------------------------------------------*)
-Goalw [increment_def] 
-      "f NSdifferentiable x ==> \
-\     increment f x h = (*f* f) (hypreal_of_real(x) + h) + \
-\     -hypreal_of_real (f x)";
-by (Blast_tac 1);
-qed "incrementI";
-
-Goal "NSDERIV f x :> D ==> \
-\    increment f x h = (*f* f) (hypreal_of_real(x) + h) + \
-\    -hypreal_of_real (f x)";
-by (etac (NSdifferentiableI RS incrementI) 1);
-qed "incrementI2";
-
-(* The Increment theorem -- Keisler p. 65 *)
-Goal "[| NSDERIV f x :> D; h: Infinitesimal; h ~= 0 |] \
-\     ==> EX e: Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h";
-by (forw_inst_tac [("h","h")] incrementI2 1 THEN rewtac nsderiv_def);
-by (dtac bspec 1 THEN Auto_tac);
-by (dtac (bex_Infinitesimal_iff2 RS iffD2) 1 THEN Step_tac 1);
-by (forw_inst_tac [("b1","hypreal_of_real(D) + y")] 
-    (hypreal_mult_right_cancel RS iffD2) 1);
-by (thin_tac "((*f* f) (hypreal_of_real(x) + h) + \
-\   - hypreal_of_real (f x)) / h = hypreal_of_real(D) + y" 2);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_times_divide1_eq RS sym]
-             delsimps [hypreal_times_divide1_eq]) 1);
-by (auto_tac (claset(),
-              simpset() addsimps [hypreal_add_mult_distrib]));
-qed "increment_thm";
-
-Goal "[| NSDERIV f x :> D; h @= 0; h ~= 0 |] \
-\      ==> EX e: Infinitesimal. increment f x h = \
-\             hypreal_of_real(D)*h + e*h";
-by (blast_tac (claset() addSDs [mem_infmal_iff RS iffD2] 
-    addSIs [increment_thm]) 1);
-qed "increment_thm2";
-
-Goal "[| NSDERIV f x :> D; h @= 0; h ~= 0 |] \
-\     ==> increment f x h @= 0";
-by (dtac increment_thm2 1 THEN auto_tac (claset() addSIs 
-    [Infinitesimal_HFinite_mult2,HFinite_add],simpset() addsimps 
-    [hypreal_add_mult_distrib RS sym,mem_infmal_iff RS sym]));
-by (etac (Infinitesimal_subset_HFinite RS subsetD) 1);
-qed "increment_inf_close_zero";
-
-(*---------------------------------------------------------------
-   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 *)
-Goalw [nsderiv_def] 
-      "[| NSDERIV g x :> D; \
-\              (*f* g) (hypreal_of_real(x) + xa) = hypreal_of_real(g x);\
-\              xa : Infinitesimal;\
-\              xa ~= 0 \
-\           |] ==> D = #0";
-by (dtac bspec 1);
-by Auto_tac;
-by (asm_full_simp_tac (simpset() addsimps [hypreal_of_real_zero RS sym]) 1);
-qed "NSDERIV_zero";
-
-(* can be proved differently using NSLIM_isCont_iff *)
-Goalw [nsderiv_def] 
-      "[| NSDERIV f x :> D; \
-\              h: Infinitesimal; h ~= 0 \
-\           |] ==> (*f* f) (hypreal_of_real(x) + h) + -hypreal_of_real(f x) @= 0";    
-by (asm_full_simp_tac (simpset() addsimps 
-    [mem_infmal_iff RS sym]) 1);
-by (rtac Infinitesimal_ratio 1);
-by (rtac inf_close_hypreal_of_real_HFinite 3);
-by Auto_tac;
-qed "NSDERIV_inf_close";
-
-(*--------------------------------------------------------------- 
-   from one version of differentiability 
- 
-                f(x) - f(a)
-              --------------- @= Db
-                  x - a
- ---------------------------------------------------------------*)
-Goal "[| NSDERIV f (g x) :> Da; \
-\        (*f* g) (hypreal_of_real(x) + xa) ~= hypreal_of_real (g x); \
-\        (*f* g) (hypreal_of_real(x) + xa) @= hypreal_of_real (g x) \
-\     |] ==> ((*f* f) ((*f* g) (hypreal_of_real(x) + xa)) \
-\                  + - hypreal_of_real (f (g x))) \
-\             / ((*f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real (g x)) \
-\            @= hypreal_of_real(Da)";
-by (auto_tac (claset(),
-       simpset() addsimps [NSDERIV_NSLIM_iff2, NSLIM_def]));
-qed "NSDERIVD1";
-
-(*-------------------------------------------------------------- 
-   from other version of differentiability 
-
-                f(x + h) - f(x)
-               ----------------- @= Db
-                       h
- --------------------------------------------------------------*)
-Goal "[| NSDERIV g x :> Db; xa: Infinitesimal; xa ~= 0 |] \
-\     ==> ((*f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real(g x)) / xa \
-\         @= hypreal_of_real(Db)";
-by (auto_tac (claset(),
-    simpset() addsimps [NSDERIV_NSLIM_iff, NSLIM_def, 
-		hypreal_of_real_zero, mem_infmal_iff, starfun_lambda_cancel]));
-qed "NSDERIVD2";
-
-(*------------------------------------------------------
-  This proof uses both definitions of differentiability.
- ------------------------------------------------------*)
-Goal "[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |] \
-\     ==> NSDERIV (f o g) x :> Da * Db";
-by (asm_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,
-    NSLIM_def,hypreal_of_real_zero,mem_infmal_iff RS sym]) 1 THEN Step_tac 1);
-by (forw_inst_tac [("f","g")] NSDERIV_inf_close 1);
-by (auto_tac (claset(),
-              simpset() addsimps [starfun_lambda_cancel2, starfun_o RS sym]));
-by (case_tac "(*f* g) (hypreal_of_real(x) + xa) = hypreal_of_real (g x)" 1);
-by (dres_inst_tac [("g","g")] NSDERIV_zero 1);
-by (auto_tac (claset(),
-    simpset() addsimps [hypreal_divide_def, hypreal_of_real_zero]));
-by (res_inst_tac [("z1","(*f* g) (hypreal_of_real(x) + xa) + -hypreal_of_real (g x)"),
-    ("y1","inverse xa")] (lemma_chain RS ssubst) 1);
-by (etac (hypreal_not_eq_minus_iff RS iffD1) 1);
-by (rtac inf_close_mult_hypreal_of_real 1);
-by (fold_tac [hypreal_divide_def]);
-by (blast_tac (claset() addIs [NSDERIVD1,
-    inf_close_minus_iff RS iffD2]) 1);
-by (blast_tac (claset() addIs [NSDERIVD2]) 1);
-qed "NSDERIV_chain";
-
-(* standard version *)
-Goal "[| DERIV f (g x) :> Da; \
-\                 DERIV g x :> Db \
-\              |] ==> DERIV (f o g) x :> Da * Db";
-by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym,
-    NSDERIV_chain]) 1);
-qed "DERIV_chain";
-
-Goal "[| DERIV f (g x) :> Da; DERIV g x :> Db |] \
-\     ==> DERIV (%x. f (g x)) x :> Da * Db";
-by (auto_tac (claset() addDs [DERIV_chain], simpset() addsimps [o_def]));
-qed "DERIV_chain2";
-
-(*------------------------------------------------------------------
-           Differentiation of natural number powers
- ------------------------------------------------------------------*)
-Goal "NSDERIV (%x. x) x :> #1";
-by (auto_tac (claset(),
-     simpset() addsimps [NSDERIV_NSLIM_iff,
-          NSLIM_def ,starfun_Id, hypreal_of_real_zero,
-           hypreal_of_real_one]));
-qed "NSDERIV_Id";
-Addsimps [NSDERIV_Id];
-
-(*derivative of the identity function*)
-Goal "DERIV (%x. x) x :> #1";
-by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1);
-qed "DERIV_Id";
-Addsimps [DERIV_Id];
-
-bind_thm ("isCont_Id", DERIV_Id RS DERIV_isCont);
-
-(*derivative of linear multiplication*)
-Goal "DERIV (op * c) x :> c";
-by (cut_inst_tac [("c","c"),("x","x")] (DERIV_Id RS DERIV_cmult) 1);
-by (Asm_full_simp_tac 1);
-qed "DERIV_cmult_Id";
-Addsimps [DERIV_cmult_Id];
-
-Goal "NSDERIV (op * c) x :> c";
-by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff]) 1);
-qed "NSDERIV_cmult_Id";
-Addsimps [NSDERIV_cmult_Id];
-
-Goal "DERIV (%x. x ^ n) x :> real_of_nat n * (x ^ (n - 1))";
-by (induct_tac "n" 1);
-by (dtac (DERIV_Id RS DERIV_mult) 2);
-by (auto_tac (claset(),simpset() addsimps 
-    [real_add_mult_distrib]));
-by (case_tac "0 < n" 1);
-by (dres_inst_tac [("x","x")] realpow_minus_mult 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [real_mult_assoc,real_add_commute]));
-qed "DERIV_pow";
-
-(* NS version *)
-Goal "NSDERIV (%x. x ^ n) x :> real_of_nat n * (x ^ (n - 1))";
-by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,DERIV_pow]) 1);
-qed "NSDERIV_pow";
-
-(*---------------------------------------------------------------
-                    Power of -1 
- ---------------------------------------------------------------*)
-
-(*Can't get rid of x ~= #0 because it isn't continuous at zero*)
-Goalw [nsderiv_def]
-     "x ~= #0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ 2))";
-by (rtac ballI 1 THEN Asm_full_simp_tac 1 THEN Step_tac 1);
-by (forward_tac [Infinitesimal_add_not_zero] 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 2); 
-by (auto_tac (claset(),
-     simpset() addsimps [starfun_inverse_inverse, realpow_two] 
-               delsimps [hypreal_minus_mult_eq1 RS sym,
-                         hypreal_minus_mult_eq2 RS sym]));
-by (asm_full_simp_tac
-     (simpset() addsimps [hypreal_inverse_add,
-          hypreal_inverse_distrib RS sym, hypreal_minus_inverse RS sym] 
-          @ hypreal_add_ac @ hypreal_mult_ac 
-       delsimps [hypreal_minus_mult_eq1 RS sym,
-                 hypreal_minus_mult_eq2 RS sym] ) 1);
-by (asm_simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym,
-                                      hypreal_add_mult_distrib2] 
-         delsimps [hypreal_minus_mult_eq1 RS sym, 
-                   hypreal_minus_mult_eq2 RS sym]) 1);
-by (res_inst_tac [("y"," inverse(- hypreal_of_real x * hypreal_of_real x)")] 
-                 inf_close_trans 1);
-by (rtac inverse_add_Infinitesimal_inf_close2 1);
-by (auto_tac (claset() addSDs [hypreal_of_real_HFinite_diff_Infinitesimal], 
-         simpset() addsimps [hypreal_minus_inverse RS sym,
-                             HFinite_minus_iff RS sym, 
-                             Infinitesimal_minus_iff RS sym]));
-by (rtac Infinitesimal_HFinite_mult2 1); 
-by Auto_tac;  
-qed "NSDERIV_inverse";
-
-
-Goal "x ~= #0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ 2))";
-by (asm_simp_tac (simpset() addsimps [NSDERIV_inverse,
-         NSDERIV_DERIV_iff RS sym] delsimps [realpow_Suc]) 1);
-qed "DERIV_inverse";
-
-(*--------------------------------------------------------------
-        Derivative of inverse 
- -------------------------------------------------------------*)
-Goal "[| DERIV f x :> d; f(x) ~= #0 |] \
-\     ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))";
-by (rtac (real_mult_commute RS subst) 1);
-by (asm_simp_tac (simpset() addsimps [real_minus_mult_eq1,
-    realpow_inverse] delsimps [realpow_Suc, 
-    real_minus_mult_eq1 RS sym]) 1);
-by (fold_goals_tac [o_def]);
-by (blast_tac (claset() addSIs [DERIV_chain,DERIV_inverse]) 1);
-qed "DERIV_inverse_fun";
-
-Goal "[| NSDERIV f x :> d; f(x) ~= #0 |] \
-\     ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))";
-by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,
-            DERIV_inverse_fun] delsimps [realpow_Suc]) 1);
-qed "NSDERIV_inverse_fun";
-
-(*--------------------------------------------------------------
-        Derivative of quotient 
- -------------------------------------------------------------*)
-Goal "[| DERIV f x :> d; DERIV g x :> e; g(x) ~= #0 |] \
-\      ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) + -(e*f(x))) / (g(x) ^ 2)";
-by (dres_inst_tac [("f","g")] DERIV_inverse_fun 1);
-by (dtac DERIV_mult 2);
-by (REPEAT(assume_tac 1));
-by (asm_full_simp_tac
-    (simpset() addsimps [real_divide_def, real_add_mult_distrib2,
-                         realpow_inverse,real_minus_mult_eq1] @ real_mult_ac 
-       delsimps [realpow_Suc, real_minus_mult_eq1 RS sym,
-                 real_minus_mult_eq2 RS sym]) 1);
-qed "DERIV_quotient";
-
-Goal "[| NSDERIV f x :> d; DERIV g x :> e; g(x) ~= #0 |] \
-\      ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) \
-\                           + -(e*f(x))) / (g(x) ^ 2)";
-by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,
-            DERIV_quotient] delsimps [realpow_Suc]) 1);
-qed "NSDERIV_quotient";
- 
-(* ------------------------------------------------------------------------ *)
-(* Caratheodory formulation of derivative at a point: standard proof        *)
-(* ------------------------------------------------------------------------ *)
-
-Goal "(DERIV f x :> l) = \
-\     (EX g. (ALL z. f z - f x = g z * (z - x)) & isCont g x & g x = l)";
-by (Step_tac 1);
-by (res_inst_tac 
-    [("x","%z. if  z = x then l else (f(z) - f(x)) / (z - x)")] exI 1);
-by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,
-    ARITH_PROVE "z ~= x ==> z - x ~= (#0::real)"]));
-by (auto_tac (claset(),simpset() addsimps [isCont_iff,DERIV_iff]));
-by (ALLGOALS(rtac (LIM_equal RS iffD1)));
-by (auto_tac (claset(),simpset() addsimps [real_diff_def,real_mult_assoc]));
-qed "CARAT_DERIV";
-
-Goal "NSDERIV f x :> l ==> \
-\     EX g. (ALL z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l";
-by (auto_tac (claset(),simpset() addsimps [NSDERIV_DERIV_iff,
-    isNSCont_isCont_iff,CARAT_DERIV]));
-qed "CARAT_NSDERIV";
-
-(* How about a NS proof? *)
-Goal "(ALL z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l \
-\     ==> NSDERIV f x :> l";
-by (auto_tac (claset(), 
-              simpset() delsimprocs real_cancel_factor
-                        addsimps [NSDERIV_iff2]));
-by (auto_tac (claset(),
-              simpset() addsimps [hypreal_mult_assoc]));
-by (asm_full_simp_tac (simpset() addsimps [hypreal_eq_minus_iff3 RS sym,
-                                           hypreal_diff_def]) 1);
-by (asm_full_simp_tac (simpset() addsimps [isNSCont_def]) 1);
-qed "CARAT_DERIVD";
- 
-
-
-(*--------------------------------------------------------------------------*)
-(* Lemmas about nested intervals and proof by bisection (cf.Harrison)       *)
-(* All considerably tidied by lcp                                           *)
-(*--------------------------------------------------------------------------*)
-
-Goal "(ALL n. (f::nat=>real) n <= f (Suc n)) --> f m <= f(m + no)";
-by (induct_tac "no" 1);
-by (auto_tac (claset() addIs [real_le_trans],simpset()));
-qed_spec_mp "lemma_f_mono_add";
-
-Goal "[| ALL n. f(n) <= f(Suc n); \
-\        ALL n. g(Suc n) <= g(n); \
-\        ALL n. f(n) <= g(n) |] \
-\     ==> Bseq f";
-by (res_inst_tac [("k","f 0"),("K","g 0")] BseqI2 1 THEN rtac allI 1);
-by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [real_le_trans], simpset()));
-by (res_inst_tac [("j","g(Suc na)")] real_le_trans 1);
-by (induct_tac "na" 2);
-by (auto_tac (claset() addIs [real_le_trans], simpset()));
-qed "f_inc_g_dec_Beq_f";
-
-Goal "[| ALL n. f(n) <= f(Suc n); \
-\        ALL n. g(Suc n) <= g(n); \
-\        ALL n. f(n) <= g(n) |] \
-\     ==> Bseq g";
-by (stac (Bseq_minus_iff RS sym) 1);
-by (res_inst_tac [("g","%x. -(f x)")] f_inc_g_dec_Beq_f 1); 
-by Auto_tac;  
-qed "f_inc_g_dec_Beq_g";
-
-Goal "[| ALL n. f n <= f (Suc n);  convergent f |] ==> f n <= lim f";
-by (rtac real_leI 1);
-by (auto_tac (claset(), 
-      simpset() addsimps [convergent_LIMSEQ_iff, LIMSEQ_iff, monoseq_Suc]));
-by (dtac real_less_sum_gt_zero 1);
-by (dres_inst_tac [("x","f n + - lim f")] spec 1);
-by Safe_tac;
-by (dres_inst_tac [("P","%na. no<=na --> ?Q na"),("x","no + n")] spec 2);
-by Auto_tac;
-by (subgoal_tac "lim f <= f(no + n)" 1);
-by (induct_tac "no" 2);
-by (auto_tac (claset() addIs [real_le_trans],
-              simpset() addsimps [real_diff_def, real_abs_def]));
-by (dres_inst_tac [("i","f(no + n)"),("no1","no")] 
-    (lemma_f_mono_add RSN (2,real_less_le_trans)) 1);
-by (auto_tac (claset(), simpset() addsimps [add_commute]));
-qed "f_inc_imp_le_lim";
-
-Goal "convergent g ==> lim (%x. - g x) = - (lim g)";
-by (rtac (LIMSEQ_minus RS limI) 1); 
-by (asm_full_simp_tac (simpset() addsimps [convergent_LIMSEQ_iff]) 1); 
-qed "lim_uminus";
-
-Goal "[| ALL n. g(Suc n) <= g(n);  convergent g |] ==> lim g <= g n";
-by (subgoal_tac "- (g n) <= - (lim g)" 1);
-by (cut_inst_tac [("f", "%x. - (g x)")] f_inc_imp_le_lim 2);
-by (auto_tac (claset(), 
-              simpset() addsimps [lim_uminus, convergent_minus_iff RS sym]));  
-qed "g_dec_imp_lim_le";
-
-Goal "[| ALL n. f(n) <= f(Suc n); \
-\        ALL n. g(Suc n) <= g(n); \
-\        ALL n. f(n) <= g(n) |] \
-\     ==> EX l m. l <= m &  ((ALL n. f(n) <= l) & f ----> l) & \
-\                           ((ALL n. m <= g(n)) & g ----> m)";
-by (subgoal_tac "monoseq f & monoseq g" 1);
-by (force_tac (claset(), simpset() addsimps [LIMSEQ_iff,monoseq_Suc]) 2);
-by (subgoal_tac "Bseq f & Bseq g" 1);
-by (blast_tac (claset() addIs [f_inc_g_dec_Beq_f, f_inc_g_dec_Beq_g]) 2); 
-by (auto_tac (claset() addSDs [Bseq_monoseq_convergent],
-              simpset() addsimps [convergent_LIMSEQ_iff]));
-by (res_inst_tac [("x","lim f")] exI 1);
-by (res_inst_tac [("x","lim g")] exI 1);
-by (auto_tac (claset() addIs [LIMSEQ_le], simpset()));
-by (auto_tac (claset(), 
-              simpset() addsimps [f_inc_imp_le_lim, g_dec_imp_lim_le, 
-                                  convergent_LIMSEQ_iff]));  
-qed "lemma_nest";
-
-Goal "[| ALL n. f(n) <= f(Suc n); \
-\        ALL n. g(Suc n) <= g(n); \
-\        ALL n. f(n) <= g(n); \
-\        (%n. f(n) - g(n)) ----> 0 |] \
-\     ==> EX l. ((ALL n. f(n) <= l) & f ----> l) & \
-\               ((ALL n. l <= g(n)) & g ----> l)";
-by (dtac lemma_nest 1 THEN Auto_tac);
-by (subgoal_tac "l = m" 1);
-by (dres_inst_tac [("X","f")] LIMSEQ_diff 2);
-by (auto_tac (claset() addIs [LIMSEQ_unique], simpset()));
-qed "lemma_nest_unique";
-
-
-Goal "a <= b ==> \
-\  ALL n. fst (Bolzano_bisect P a b n) <= snd (Bolzano_bisect P a b n)";
-by (rtac allI 1);
-by (induct_tac "n" 1);
-by (auto_tac (claset(), simpset() addsimps [Let_def, split_def]));  
-qed "Bolzano_bisect_le";
-
-Goal "a <= b ==> \
-\  ALL n. fst(Bolzano_bisect P a b n) <= fst (Bolzano_bisect P a b (Suc n))";
-by (rtac allI 1);
-by (induct_tac "n" 1);
-by (auto_tac (claset(), 
-              simpset() addsimps [Bolzano_bisect_le, Let_def, split_def]));  
-qed "Bolzano_bisect_fst_le_Suc";
-
-Goal "a <= b ==> \
-\  ALL n. snd(Bolzano_bisect P a b (Suc n)) <= snd (Bolzano_bisect P a b n)";
-by (rtac allI 1);
-by (induct_tac "n" 1);
-by (auto_tac (claset(), 
-              simpset() addsimps [Bolzano_bisect_le, Let_def, split_def]));  
-qed "Bolzano_bisect_Suc_le_snd";
-
-Goal "((x::real) = y / (#2 * z)) = (#2 * x = y/z)";
-by Auto_tac;  
-by (dres_inst_tac [("f","%u. (#1/#2)*u")] arg_cong 1); 
-by Auto_tac;  
-qed "eq_divide_2_times_iff";
-
-Goal "a <= b ==> \
-\     snd(Bolzano_bisect P a b n) - fst(Bolzano_bisect P a b n) = \
-\     (b-a) / (#2 ^ n)";
-by (induct_tac "n" 1);
-by (auto_tac (claset(), 
-      simpset() addsimps [eq_divide_2_times_iff, real_add_divide_distrib, 
-                          Let_def, split_def]));
-by (auto_tac (claset(), 
-              simpset() addsimps (real_add_ac@[Bolzano_bisect_le, real_diff_def])));
-qed "Bolzano_bisect_diff";
-
-val Bolzano_nest_unique =
-    [Bolzano_bisect_fst_le_Suc, Bolzano_bisect_Suc_le_snd, Bolzano_bisect_le] 
-    MRS lemma_nest_unique;
-
-(*P_prem is a looping simprule, so it works better if it isn't an assumption*)
-val P_prem::notP_prem::rest =
-Goal "[| !!a b c. [| P(a,b); P(b,c); a <= b; b <= c|] ==> P(a,c); \
-\        ~ P(a,b);  a <= b |] ==> \
-\     ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))";
-by (cut_facts_tac rest 1);
-by (induct_tac "n" 1);
-by (auto_tac (claset(), 
-              simpset() delsimps [surjective_pairing RS sym]
-			addsimps [notP_prem, Let_def, split_def]));  
-by (swap_res_tac [P_prem] 1);
-by (assume_tac 1); 
-by (auto_tac (claset(), simpset() addsimps [Bolzano_bisect_le]));  
-qed "not_P_Bolzano_bisect";
-
-(*Now we re-package P_prem as a formula*)
-Goal "[| ALL a b c. P(a,b) & P(b,c) & a <= b & b <= c --> P(a,c); \
-\        ~ P(a,b);  a <= b |] ==> \
-\     ALL n. ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))";
-by (blast_tac (claset() addSEs [not_P_Bolzano_bisect RSN (2,rev_notE)]) 1); 
-qed "not_P_Bolzano_bisect'";
-
-
-Goal "[| ALL a b c. P(a,b) & P(b,c) & a <= b & b <= c --> P(a,c); \
-\        ALL x. EX d::real. 0 < d & \
-\               (ALL a b. a <= x & x <= b & (b - a) < d --> P(a,b)); \
-\        a <= b |]  \
-\     ==> P(a,b)";
-by (rtac (inst "P1" "P" Bolzano_nest_unique RS exE) 1);
-by (REPEAT (assume_tac 1));
-by (rtac LIMSEQ_minus_cancel 1);
-by (asm_simp_tac (simpset() addsimps [Bolzano_bisect_diff,
-                                      LIMSEQ_divide_realpow_zero]) 1); 
-by (rtac ccontr 1);
-by (dtac not_P_Bolzano_bisect' 1); 
-by (REPEAT (assume_tac 1));
-by (rename_tac "l" 1);
-by (dres_inst_tac [("x","l")] spec 1 THEN Clarify_tac 1);
-by (rewtac LIMSEQ_def);
-by (dres_inst_tac [("P", "%r. #0<r --> ?Q r"), ("x","d/#2")] spec 1);
-by (dres_inst_tac [("P", "%r. #0<r --> ?Q r"), ("x","d/#2")] spec 1);
-by (dtac real_less_half_sum 1);
-by Safe_tac;
-(*linear arithmetic bug if we just use Asm_simp_tac*)
-by (ALLGOALS Asm_full_simp_tac);
-by (dres_inst_tac [("x","fst(Bolzano_bisect P a b (no + noa))")] spec 1);
-by (dres_inst_tac [("x","snd(Bolzano_bisect P a b (no + noa))")] spec 1);
-by Safe_tac;
-by (ALLGOALS Asm_simp_tac);
-by (res_inst_tac [("j","abs(fst(Bolzano_bisect P a b(no + noa)) - l) + \
-\                       abs(snd(Bolzano_bisect P a b(no + noa)) - l)")] 
-    real_le_less_trans 1);
-by (asm_simp_tac (simpset() addsimps [real_abs_def]) 1);  
-by (rtac (real_sum_of_halves RS subst) 1);
-by (rtac real_add_less_mono 1);
-by (ALLGOALS 
-    (asm_full_simp_tac (simpset() addsimps [symmetric real_diff_def])));
-qed "lemma_BOLZANO";
-
-
-Goal "((ALL a b c. (a <= b & b <= c & P(a,b) & P(b,c)) --> P(a,c)) & \
-\      (ALL x. EX d::real. 0 < d & \
-\               (ALL a b. a <= x & x <= b & (b - a) < d --> P(a,b)))) \
-\     --> (ALL a b. a <= b --> P(a,b))";
-by (Clarify_tac 1);
-by (blast_tac (claset() addIs [lemma_BOLZANO]) 1); 
-qed "lemma_BOLZANO2";
-
-
-(*----------------------------------------------------------------------------*)
-(* Intermediate Value Theorem (prove contrapositive by bisection)             *)
-(*----------------------------------------------------------------------------*)
-
-Goal "[| f(a) <= y & y <= f(b); \
-\        a <= b; \
-\        (ALL x. a <= x & x <= b --> isCont f x) |] \
-\     ==> EX x. a <= x & x <= b & f(x) = y";
-by (rtac contrapos_pp 1);
-by (assume_tac 1);
-by (cut_inst_tac
-    [("P","%(u,v). a <= u & u <= v & v <= b --> ~(f(u) <= y & y <= f(v))")] 
-    lemma_BOLZANO2 1);
-by (Step_tac 1);
-by (ALLGOALS(Asm_full_simp_tac));
-by (Blast_tac 2);
-by (asm_full_simp_tac (simpset() addsimps [isCont_iff,LIM_def]) 1);
-by (rtac ccontr 1);
-by (subgoal_tac "a <= x & x <= b" 1);
-by (Asm_full_simp_tac 2);
-by (dres_inst_tac [("P", "%d. #0<d --> ?P d"),("x","#1")] spec 2);
-by (Step_tac 2);
-by (Asm_full_simp_tac 2);
-by (Asm_full_simp_tac 2);
-by (REPEAT(blast_tac (claset() addIs [real_le_trans]) 2));
-by (REPEAT(dres_inst_tac [("x","x")] spec 1));
-by (Asm_full_simp_tac 1);
-by (dres_inst_tac [("P", "%r. ?P r --> (EX s. #0<s & ?Q r s)"),
-                   ("x","abs(y - f x)")] spec 1);
-by (Step_tac 1);
-by (asm_full_simp_tac (simpset() addsimps []) 1);
-by (dres_inst_tac [("x","s")] spec 1);
-by (Clarify_tac 1);
-by (cut_inst_tac [("R1.0","f x"),("R2.0","y")] real_linear 1);
-by Safe_tac;
-by (dres_inst_tac [("x","ba - x")] spec 1);
-by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [abs_iff])));
-by (dres_inst_tac [("x","aa - x")] spec 1);
-by (case_tac "x <= aa" 1);
-by (ALLGOALS Asm_full_simp_tac);
-by (dres_inst_tac [("z","x"),("w","aa")] real_le_anti_sym 1);
-by (assume_tac 1 THEN Asm_full_simp_tac 1);
-qed "IVT";
-
-
-Goal "[| f(b) <= y & y <= f(a); \
-\        a <= b; \
-\        (ALL x. a <= x & x <= b --> isCont f x) \
-\     |] ==> EX x. a <= x & x <= b & f(x) = y";
-by (subgoal_tac "- f a <= -y & -y <= - f b" 1);
-by (thin_tac "f b <= y & y <= f a" 1);
-by (dres_inst_tac [("f","%x. - f x")] IVT 1);
-by (auto_tac (claset() addIs [isCont_minus],simpset()));
-qed "IVT2";
-
-
-(*HOL style here: object-level formulations*)
-Goal "(f(a) <= y & y <= f(b) & a <= b & \
-\     (ALL x. a <= x & x <= b --> isCont f x)) \
-\     --> (EX x. a <= x & x <= b & f(x) = y)";
-by (blast_tac (claset() addIs [IVT]) 1);
-qed "IVT_objl";
-
-Goal "(f(b) <= y & y <= f(a) & a <= b & \
-\     (ALL x. a <= x & x <= b --> isCont f x)) \
-\     --> (EX x. a <= x & x <= b & f(x) = y)";
-by (blast_tac (claset() addIs [IVT2]) 1);
-qed "IVT2_objl";
-
-(*---------------------------------------------------------------------------*)
-(* By bisection, function continuous on closed interval is bounded above     *)
-(*---------------------------------------------------------------------------*)
-
-Goal "abs (real_of_nat x) = real_of_nat x";
-by (auto_tac (claset() addIs [abs_eqI1],simpset()));
-qed "abs_real_of_nat_cancel";
-Addsimps [abs_real_of_nat_cancel];
-
-Goal "~ abs(x) + 1r < x";
-by (rtac real_leD 1);
-by (auto_tac (claset() addIs [abs_ge_self RS real_le_trans],simpset()));
-qed "abs_add_one_not_less_self";
-Addsimps [abs_add_one_not_less_self];
-
-
-Goal "[| a <= b; ALL x. a <= x & x <= b --> isCont f x |]\
-\     ==> EX M. ALL x. a <= x & x <= b --> f(x) <= M";
-by (cut_inst_tac [("P","%(u,v). a <= u & u <= v & v <= b --> \
-\                         (EX M. ALL x. u <= x & x <= v --> f x <= M)")] 
-    lemma_BOLZANO2 1);
-by (Step_tac 1);
-by (ALLGOALS(Asm_full_simp_tac));
-by (cut_inst_tac [("x","M"),("y","Ma")] 
-    (CLAIM "x <= y | y <= (x::real)") 1);
-by (Step_tac 1);
-by (res_inst_tac [("x","Ma")] exI 1);
-by (Step_tac 1);
-by (cut_inst_tac [("x","xb"),("y","xa")] 
-    (CLAIM "x <= y | y <= (x::real)") 1);
-by (Step_tac 1);
-by (rtac real_le_trans 1 THEN assume_tac 2);
-by (Asm_full_simp_tac 1);
-by (Asm_full_simp_tac 1);
-by (res_inst_tac [("x","M")] exI 1);
-by (Step_tac 1);
-by (cut_inst_tac [("x","xb"),("y","xa")] 
-    (CLAIM "x <= y | y <= (x::real)") 1);
-by (Step_tac 1);
-by (Asm_full_simp_tac 1);
-by (rtac real_le_trans 1 THEN assume_tac 2);
-by (Asm_full_simp_tac 1);
-by (case_tac "a <= x & x <= b" 1);
-by (res_inst_tac [("x","#1")] exI 2);
-by (auto_tac (claset(),
-              simpset() addsimps [LIM_def,isCont_iff]));
-by (dres_inst_tac [("x","x")] spec 1 THEN Auto_tac);
-by (thin_tac "ALL M. EX x. a <= x & x <= b & ~ f x <= M" 1);
-by (dres_inst_tac [("x","#1")] spec 1);
-by Auto_tac;  
-by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
-by (res_inst_tac [("x","abs(f x) + #1")] exI 1 THEN Step_tac 1);
-by (dres_inst_tac [("x","xa - x")] spec 1 THEN Auto_tac);
-by (res_inst_tac [("j","abs(f xa)")] real_le_trans 3);
-by (res_inst_tac [("j","abs(f x) + abs(f(xa) - f(x))")] real_le_trans 4);
-by (auto_tac (claset() addIs [abs_triangle_ineq RSN (2, real_le_trans)],
-              simpset() addsimps [real_diff_def,abs_ge_self]));
-by (auto_tac (claset(),
-              simpset() addsimps [real_abs_def] addsplits [split_if_asm]));
-qed "isCont_bounded";
-
-(*----------------------------------------------------------------------------*)
-(* Refine the above to existence of least upper bound                         *)
-(*----------------------------------------------------------------------------*)
-
-Goal "((EX x. x : S) & (EX y. isUb UNIV S (y::real))) --> \
-\     (EX t. isLub UNIV S t)";
-by (blast_tac (claset() addIs [reals_complete]) 1);
-qed "lemma_reals_complete";
-
-Goal "[| a <= b; ALL x. a <= x & x <= b --> isCont f x |] \
-\        ==> EX M. (ALL x. a <= x & x <= b --> f(x) <= M) & \
-\                  (ALL N. N < M --> (EX x. a <= x & x <= b & N < f(x)))";
-by (cut_inst_tac [("S","Collect (%y. EX x. a <= x & x <= b & y = f x)")]
-    lemma_reals_complete 1);
-by Auto_tac;
-by (dtac isCont_bounded 1 THEN assume_tac 1);
-by (auto_tac (claset(),simpset() addsimps [isUb_def,leastP_def,
-    isLub_def,setge_def,setle_def]));
-by (rtac exI 1 THEN Auto_tac);
-by (REPEAT(dtac spec 1) THEN Auto_tac);
-by (dres_inst_tac [("x","x")] spec 1);
-by (auto_tac (claset() addSIs [real_leI],simpset()));
-qed "isCont_has_Ub";
-
-(*----------------------------------------------------------------------------*)
-(* Now show that it attains its upper bound                                   *)
-(*----------------------------------------------------------------------------*)
-
-Goal "[| a <= b; ALL x. a <= x & x <= b --> isCont f x |] \
-\        ==> EX M. (ALL x. a <= x & x <= b --> f(x) <= M) & \
-\                  (EX x. a <= x & x <= b & f(x) = M)";
-by (ftac isCont_has_Ub 1 THEN assume_tac 1);
-by (Clarify_tac 1);
-by (res_inst_tac [("x","M")] exI 1);
-by (Asm_full_simp_tac 1); 
-by (rtac ccontr 1);
-by (subgoal_tac "ALL x. a <= x & x <= b --> f x < M" 1 THEN Step_tac 1);
-by (rtac ccontr 2 THEN dtac real_leI 2);
-by (dres_inst_tac [("z","M")] real_le_anti_sym 2);
-by (REPEAT(Blast_tac 2));
-by (subgoal_tac "ALL x. a <= x & x <= b --> isCont (%x. inverse(M - f x)) x" 1);
-by (Step_tac 1);
-by (EVERY[rtac isCont_inverse 2, rtac isCont_diff 2, rtac notI 4]);
-by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [real_diff_eq_eq])));
-by (Blast_tac 2);
-by (subgoal_tac 
-    "EX k. ALL x. a <= x & x <= b --> (%x. inverse(M - (f x))) x <= k" 1);
-by (rtac isCont_bounded 2);
-by (Step_tac 1);
-by (subgoal_tac "ALL x. a <= x & x <= b --> #0 < inverse(M - f(x))" 1);
-by (Asm_full_simp_tac 1); 
-by (Step_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [real_less_diff_eq]) 2);
-by (subgoal_tac 
-    "ALL x. a <= x & x <= b --> (%x. inverse(M - (f x))) x < (k + #1)" 1);
-by (Step_tac 1);
-by (res_inst_tac [("j","k")] real_le_less_trans 2);
-by (asm_full_simp_tac (simpset() addsimps [real_zero_less_one]) 3);
-by (Asm_full_simp_tac 2); 
-by (subgoal_tac "ALL x. a <= x & x <= b --> \
-\                inverse(k + #1) < inverse((%x. inverse(M - (f x))) x)" 1);
-by (Step_tac 1);
-by (rtac real_inverse_less_swap 2);
-by (ALLGOALS Asm_full_simp_tac);
-by (dres_inst_tac [("P", "%N. N<M --> ?Q N"),
-                   ("x","M - inverse(k + #1)")] spec 1);
-by (Step_tac 1 THEN dtac real_leI 1);
-by (dtac (real_le_diff_eq RS iffD1) 1);
-by (REPEAT(dres_inst_tac [("x","a")] spec 1));
-by (Asm_full_simp_tac 1);
-by (asm_full_simp_tac 
-    (simpset() addsimps [real_inverse_eq_divide, pos_real_divide_le_eq]) 1); 
-by (cut_inst_tac [("x","k"),("y","M-f a")] real_0_less_mult_iff 1);
-by (Asm_full_simp_tac 1); 
-(*last one*)
-by (REPEAT(dres_inst_tac [("x","x")] spec 1));
-by (Asm_full_simp_tac 1); 
-qed "isCont_eq_Ub";
-
-
-(*----------------------------------------------------------------------------*)
-(* Same theorem for lower bound                                               *)
-(*----------------------------------------------------------------------------*)
-
-Goal "[| a <= b; ALL x. a <= x & x <= b --> isCont f x |] \
-\        ==> EX M. (ALL x. a <= x & x <= b --> M <= f(x)) & \
-\                  (EX x. a <= x & x <= b & f(x) = M)";
-by (subgoal_tac "ALL x. a <= x & x <= b --> isCont (%x. -(f x)) x" 1);
-by (blast_tac (claset() addIs [isCont_minus]) 2);
-by (dres_inst_tac [("f","(%x. -(f x))")] isCont_eq_Ub 1);
-by (Step_tac 1);
-by Auto_tac;
-qed "isCont_eq_Lb";
-
-
-(* ------------------------------------------------------------------------- *)
-(* Another version.                                                          *)
-(* ------------------------------------------------------------------------- *)
-
-Goal "[|a <= b; ALL x. a <= x & x <= b --> isCont f x |] \
-\     ==> EX L M. (ALL x. a <= x & x <= b --> L <= f(x) & f(x) <= M) & \
-\         (ALL y. L <= y & y <= M --> (EX x. a <= x & x <= b & (f(x) = y)))";
-by (ftac isCont_eq_Lb 1);
-by (ftac isCont_eq_Ub 2);
-by (REPEAT(assume_tac 1));
-by (Step_tac 1);
-by (res_inst_tac [("x","f x")] exI 1);
-by (res_inst_tac [("x","f xa")] exI 1);
-by (Asm_full_simp_tac 1);
-by (Step_tac 1);
-by (cut_inst_tac [("x","x"),("y","xa")] (CLAIM "x <= y | y <= (x::real)") 1);
-by (Step_tac 1);
-by (cut_inst_tac [("f","f"),("a","x"),("b","xa"),("y","y")] IVT_objl 1);
-by (cut_inst_tac [("f","f"),("a","xa"),("b","x"),("y","y")] IVT2_objl 2);
-by (Step_tac 1);
-by (res_inst_tac [("x","xb")] exI 2);
-by (res_inst_tac [("x","xb")] exI 4);
-by (ALLGOALS(Asm_full_simp_tac));
-qed "isCont_Lb_Ub";
-
-(*----------------------------------------------------------------------------*)
-(* If f'(x) > 0 then x is locally strictly increasing at the right            *)
-(*----------------------------------------------------------------------------*)
-
-Goalw [deriv_def,LIM_def] 
-    "[| DERIV f x :> l;  #0 < l |] ==> \
-\      EX d. #0 < d & (ALL h. #0 < h & h < d --> f(x) < f(x + h))";
-by (dtac spec 1 THEN Auto_tac);
-by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
-by (subgoal_tac "#0 < l*h" 1);
-by (asm_full_simp_tac (simpset() addsimps [real_0_less_mult_iff]) 2); 
-by (dres_inst_tac [("x","h")] spec 1);
-by (asm_full_simp_tac
-    (simpset() addsimps [real_abs_def, real_inverse_eq_divide, 
-                 pos_real_le_divide_eq, pos_real_less_divide_eq]
-              addsplits [split_if_asm]) 1); 
-qed "DERIV_left_inc";
-
-Goalw [deriv_def,LIM_def] 
-    "[| DERIV f x :> l;  l < #0 |] ==> \
-\      EX d. #0 < d & (ALL h. #0 < h & h < d --> f(x) < f(x - h))";
-by (dres_inst_tac [("x","-l")] spec 1 THEN Auto_tac);
-by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
-by (subgoal_tac "l*h < #0" 1);
-by (asm_full_simp_tac (simpset() addsimps [real_mult_less_0_iff]) 2); 
-by (dres_inst_tac [("x","-h")] spec 1);
-by (asm_full_simp_tac
-    (simpset() addsimps [real_abs_def, real_inverse_eq_divide, 
-                         pos_real_less_divide_eq,
-                         symmetric real_diff_def]
-               addsplits [split_if_asm]) 1);
-by (subgoal_tac "#0 < (f (x - h) - f x)/h" 1);
-by (arith_tac 2);
-by (asm_full_simp_tac
-    (simpset() addsimps [pos_real_less_divide_eq]) 1); 
-qed "DERIV_left_dec";
-
-
-Goal "[| DERIV f x :> l; \
-\        EX d. #0 < d & (ALL y. abs(x - y) < d --> f(y) <= f(x)) |] \
-\     ==> l = #0";
-by (res_inst_tac [("R1.0","l"),("R2.0","#0")] real_linear_less2 1);
-by (Step_tac 1);
-by (dtac DERIV_left_dec 1);
-by (dtac DERIV_left_inc 3);
-by (Step_tac 1);
-by (dres_inst_tac [("d1.0","d"),("d2.0","da")] real_lbound_gt_zero 1);
-by (dres_inst_tac [("d1.0","d"),("d2.0","da")] real_lbound_gt_zero 3);
-by (Step_tac 1);
-by (dres_inst_tac [("x","x - e")] spec 1);
-by (dres_inst_tac [("x","x + e")] spec 2);
-by (auto_tac (claset(), simpset() addsimps [real_abs_def]));
-qed "DERIV_local_max";
-
-(*----------------------------------------------------------------------------*)
-(* Similar theorem for a local minimum                                        *)
-(*----------------------------------------------------------------------------*)
-
-Goal "[| DERIV f x :> l; \
-\        EX d::real. #0 < d & (ALL y. abs(x - y) < d --> f(x) <= f(y)) |] \
-\     ==> l = #0";
-by (dtac (DERIV_minus RS DERIV_local_max) 1); 
-by Auto_tac;  
-qed "DERIV_local_min";
-
-(*----------------------------------------------------------------------------*)
-(* In particular if a function is locally flat                                *)
-(*----------------------------------------------------------------------------*)
-
-Goal "[| DERIV f x :> l; \
-\        EX d. #0 < d & (ALL y. abs(x - y) < d --> f(x) = f(y)) |] \
-\     ==> l = #0";
-by (auto_tac (claset() addSDs [DERIV_local_max],simpset()));
-qed "DERIV_local_const";
-
-(*----------------------------------------------------------------------------*)
-(* Lemma about introducing open ball in open interval                         *)
-(*----------------------------------------------------------------------------*)
-
-Goal "[| a < x;  x < b |] ==> \
-\       EX d::real. #0 < d &  (ALL y. abs(x - y) < d --> a < y & y < b)";
-by (simp_tac (simpset() addsimps [abs_interval_iff]) 1);
-by (cut_inst_tac [("x","x - a"),("y","b - x")] 
-    (CLAIM "x <= y | y <= (x::real)") 1);
-by (Step_tac 1);
-by (res_inst_tac [("x","x - a")] exI 1);
-by (res_inst_tac [("x","b - x")] exI 2);
-by Auto_tac;
-by (auto_tac (claset(),simpset() addsimps [real_less_diff_eq]));
-qed "lemma_interval_lt";
-
-Goal "[| a < x;  x < b |] ==> \
-\       EX d::real. #0 < d &  (ALL y. abs(x - y) < d --> a <= y & y <= b)";
-by (dtac lemma_interval_lt 1);
-by Auto_tac;
-by (auto_tac (claset() addSIs [exI] ,simpset()));
-qed "lemma_interval";
-
-(*-----------------------------------------------------------------------
-            Rolle's Theorem
-   If f is defined and continuous on the finite closed interval [a,b]
-   and differentiable a least on the open interval (a,b), and f(a) = f(b),
-   then x0 : (a,b) such that f'(x0) = #0
- ----------------------------------------------------------------------*)
-
-Goal "[| a < b; f(a) = f(b); \
-\        ALL x. a <= x & x <= b --> isCont f x; \
-\        ALL x. a < x & x < b --> f differentiable x \
-\     |] ==> EX z. a < z & z < b & DERIV f z :> #0";
-by (ftac (real_less_imp_le RS isCont_eq_Ub) 1);
-by (EVERY1[assume_tac,Step_tac]);
-by (ftac (real_less_imp_le RS isCont_eq_Lb) 1);
-by (EVERY1[assume_tac,Step_tac]);
-by (case_tac "a < x & x < b" 1 THEN etac conjE 1);
-by (Asm_full_simp_tac 2);
-by (forw_inst_tac [("a","a"),("x","x")] lemma_interval 1);
-by (EVERY1[assume_tac,etac exE]);
-by (res_inst_tac [("x","x")] exI 1 THEN Asm_full_simp_tac 1);
-by (subgoal_tac "(EX l. DERIV f x :> l) & \
-\        (EX d. #0 < d & (ALL y. abs(x - y) < d --> f(y) <= f(x)))" 1);
-by (Clarify_tac 1 THEN rtac conjI 2);
-by (blast_tac (claset() addIs [differentiableD]) 2);
-by (Blast_tac 2);
-by (ftac DERIV_local_max 1);
-by (EVERY1[Blast_tac,Blast_tac]);
-by (case_tac "a < xa & xa < b" 1 THEN etac conjE 1);
-by (Asm_full_simp_tac 2);
-by (forw_inst_tac [("a","a"),("x","xa")] lemma_interval 1);
-by (EVERY1[assume_tac,etac exE]);
-by (res_inst_tac [("x","xa")] exI 1 THEN Asm_full_simp_tac 1);
-by (subgoal_tac "(EX l. DERIV f xa :> l) & \
-\        (EX d. #0 < d & (ALL y. abs(xa - y) < d --> f(xa) <= f(y)))" 1);
-by (Clarify_tac 1 THEN rtac conjI 2);
-by (blast_tac (claset() addIs [differentiableD]) 2);
-by (Blast_tac 2);
-by (ftac DERIV_local_min 1);
-by (EVERY1[Blast_tac,Blast_tac]);
-by (subgoal_tac "ALL x. a <= x & x <= b --> f(x) = f(b)" 1);
-by (Clarify_tac 2);
-by (rtac real_le_anti_sym 2);
-by (subgoal_tac "f b = f x" 2);
-by (Asm_full_simp_tac 2);
-by (res_inst_tac [("x1","a"),("y1","x")] (real_le_imp_less_or_eq RS disjE) 2);
-by (assume_tac 2);
-by (dres_inst_tac [("z","x"),("w","b")] real_le_anti_sym 2);
-by (subgoal_tac "f b = f xa" 5);
-by (Asm_full_simp_tac 5);
-by (res_inst_tac [("x1","a"),("y1","xa")] (real_le_imp_less_or_eq RS disjE) 5);
-by (assume_tac 5);
-by (dres_inst_tac [("z","xa"),("w","b")] real_le_anti_sym 5);
-by (REPEAT(Asm_full_simp_tac 2));
-by (dtac real_dense 1 THEN etac exE 1);
-by (res_inst_tac [("x","r")] exI 1 THEN Asm_full_simp_tac 1);
-by (etac conjE 1);
-by (forw_inst_tac [("a","a"),("x","r")] lemma_interval 1);
-by (EVERY1[assume_tac, etac exE]);
-by (subgoal_tac "(EX l. DERIV f r :> l) & \
-\        (EX d. #0 < d & (ALL y. abs(r - y) < d --> f(r) = f(y)))" 1);
-by (Clarify_tac 1 THEN rtac conjI 2);
-by (blast_tac (claset() addIs [differentiableD]) 2);
-by (EVERY1[ftac DERIV_local_const, Blast_tac, Blast_tac]);
-by (res_inst_tac [("x","d")] exI 1);
-by (EVERY1[rtac conjI, Blast_tac, rtac allI, rtac impI]);
-by (res_inst_tac [("s","f b")] trans 1);
-by (blast_tac (claset() addSDs [real_less_imp_le]) 1);
-by (rtac sym 1 THEN Blast_tac 1);
-qed "Rolle";
-
-(*----------------------------------------------------------------------------*)
-(* Mean value theorem                                                         *)
-(*----------------------------------------------------------------------------*)
-
-Goal "f a - (f b - f a)/(b - a) * a = \
-\     f b - (f b - f a)/(b - a) * (b::real)";
-by (case_tac "a = b" 1);
-by (asm_full_simp_tac (simpset() addsimps [REAL_DIVIDE_ZERO]) 1); 
-by (res_inst_tac [("c1","b - a")] (real_mult_left_cancel RS iffD1) 1);
-by (arith_tac 1);
-by (auto_tac (claset(),
-              simpset() addsimps [real_diff_mult_distrib2]));
-by (auto_tac (claset(),
-           simpset() addsimps [real_diff_mult_distrib]));
-qed "lemma_MVT";
-
-Goal "[| a < b; \
-\        ALL x. a <= x & x <= b --> isCont f x; \
-\        ALL x. a < x & x < b --> f differentiable x |] \
-\     ==>  EX l z. a < z & z < b & DERIV f z :> l & \
-\                  (f(b) - f(a) = (b - a) * l)";
-by (dres_inst_tac [("f","%x. f(x) - (((f(b) - f(a)) / (b - a)) * x)")]
-    Rolle 1);
-by (rtac lemma_MVT 1);
-by (Step_tac 1);
-by (rtac isCont_diff 1 THEN Blast_tac 1);
-by (rtac (isCont_const RS isCont_mult) 1);
-by (rtac isCont_Id 1); 
-by (dres_inst_tac [("P", "%x. ?Pre x --> f differentiable x"), 
-                   ("x","x")] spec 1);
-by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 1);
-by (Step_tac 1);
-by (res_inst_tac [("x","xa - ((f(b) - f(a)) / (b - a))")] exI 1);
-by (rtac DERIV_diff 1 THEN assume_tac 1);
-(*derivative of a linear function is the constant...*)
-by (subgoal_tac "(%x. (f b - f a) * x / (b - a)) = \
-\                op * ((f b - f a) / (b - a))" 1);
-by (rtac ext 2 THEN Simp_tac 2);
-by (Asm_full_simp_tac 1); 
-(*final case*)
-by (res_inst_tac [("x","((f(b) - f(a)) / (b - a))")] exI 1);
-by (res_inst_tac [("x","z")] exI 1);
-by (Step_tac 1);
-by (Asm_full_simp_tac 2); 
-by (subgoal_tac "DERIV (%x. ((f(b) - f(a)) / (b - a)) * x) z :> \
-\                           ((f(b) - f(a)) / (b - a))" 1);
-by (rtac DERIV_cmult_Id 2); 
-by (dtac DERIV_add 1 THEN assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [real_add_assoc, real_diff_def]) 1);
-qed "MVT";
-
-(*----------------------------------------------------------------------------*)
-(* Theorem that function is constant if its derivative is 0 over an interval. *)
-(*----------------------------------------------------------------------------*)
-
-Goal "[| a < b; \
-\        ALL x. a <= x & x <= b --> isCont f x; \
-\        ALL x. a < x & x < b --> DERIV f x :> #0 |] \
-\       ==> (f b = f a)";
-by (dtac MVT 1 THEN assume_tac 1);
-by (blast_tac (claset() addIs [differentiableI]) 1);
-by (auto_tac (claset() addSDs [DERIV_unique],simpset() 
-    addsimps [real_diff_eq_eq]));
-qed "DERIV_isconst_end";
-
-Goal "[| a < b; \
-\        ALL x. a <= x & x <= b --> isCont f x; \
-\        ALL x. a < x & x < b --> DERIV f x :> #0 |] \
-\       ==> ALL x. a <= x & x <= b --> f x = f a";
-by (Step_tac 1);
-by (dres_inst_tac [("x","a")] real_le_imp_less_or_eq 1);
-by (Step_tac 1);
-by (dres_inst_tac [("b","x")] DERIV_isconst_end 1);
-by Auto_tac;
-qed "DERIV_isconst1";
-
-Goal "[| a < b; \
-\        ALL x. a <= x & x <= b --> isCont f x; \
-\        ALL x. a < x & x < b --> DERIV f x :> #0; \
-\        a <= x; x <= b |] \
-\       ==> f x = f a";
-by (blast_tac (claset() addDs [DERIV_isconst1]) 1);
-qed "DERIV_isconst2";
-
-Goal "ALL x. DERIV f x :> #0 ==> f(x) = f(y)";
-by (res_inst_tac [("R1.0","x"),("R2.0","y")] real_linear_less2 1);
-by (rtac sym 1);
-by (auto_tac (claset() addIs [DERIV_isCont,DERIV_isconst_end],simpset()));
-qed "DERIV_isconst_all";
-
-Goal "[|a ~= b; ALL x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b - a) * k";
-by (res_inst_tac [("R1.0","a"),("R2.0","b")] real_linear_less2 1);
-by Auto_tac;
-by (ALLGOALS(dres_inst_tac [("f","f")] MVT));
-by (auto_tac (claset() addDs [DERIV_isCont,DERIV_unique],simpset() addsimps 
-    [differentiable_def]));
-by (auto_tac (claset() addDs [DERIV_unique],
-       simpset() addsimps [real_add_mult_distrib, real_diff_def,
-                           real_minus_mult_eq1 RS sym]));
-qed "DERIV_const_ratio_const";
-
-Goal "[|a ~= b; ALL x. DERIV f x :> k |] ==> (f(b) - f(a))/(b - a) = k";
-by (res_inst_tac [("c1","b - a")] (real_mult_right_cancel RS iffD1) 1);
-by (auto_tac (claset() addSDs [DERIV_const_ratio_const], 
-              simpset() addsimps [real_mult_assoc]));
-qed "DERIV_const_ratio_const2";
-
-Goal "((a + b) /#2 - a) = (b - a)/(#2::real)";
-by Auto_tac;  
-qed "real_average_minus_first";
-Addsimps [real_average_minus_first];
-
-Goal "((b + a)/#2 - a) = (b - a)/(#2::real)";
-by Auto_tac;  
-qed "real_average_minus_second";
-Addsimps [real_average_minus_second];
-
-
-(* Gallileo's "trick": average velocity = av. of end velocities *)
-Goal "[|a ~= (b::real); ALL x. DERIV v x :> k|] \
-\     ==> v((a + b)/#2) = (v a + v b)/#2";
-by (res_inst_tac [("R1.0","a"),("R2.0","b")] real_linear_less2 1);
-by Auto_tac;
-by (ftac DERIV_const_ratio_const2 1 THEN assume_tac 1);
-by (ftac DERIV_const_ratio_const2 2 THEN assume_tac 2);
-by (dtac real_less_half_sum 1);
-by (dtac real_gt_half_sum 2); 
-by (ftac (real_not_refl2 RS DERIV_const_ratio_const2) 1 THEN assume_tac 1);
-by (dtac ((real_not_refl2 RS not_sym) RS DERIV_const_ratio_const2) 2
-    THEN assume_tac 2);
-by (ALLGOALS (dres_inst_tac [("f","%u. (b-a)*u")] arg_cong)); 
-by (auto_tac (claset(), simpset() addsimps [real_inverse_eq_divide])); 
-by (asm_full_simp_tac (simpset() addsimps [real_add_commute, eq_commute]) 1);  
-qed "DERIV_const_average";
-
-
-(* ------------------------------------------------------------------------ *)
-(* Dull lemma that an continuous injection on an interval must have a strict*)
-(* maximum at an end point, not in the middle.                              *)
-(* ------------------------------------------------------------------------ *)
-
-Goal "[|#0 < d; ALL z. abs(z - x) <= d --> g(f z) = z; \
-\       ALL z. abs(z - x) <= d --> isCont f z |]  \
-\     ==> ~(ALL z. abs(z - x) <= d --> f(z) <= f(x))";
-by (rtac notI 1);
-by (rotate_tac 3 1);
-by (forw_inst_tac [("x","x - d")] spec 1);
-by (forw_inst_tac [("x","x + d")] spec 1);
-by (Step_tac 1);
-by (cut_inst_tac [("x","f(x - d)"),("y","f(x + d)")] 
-    (ARITH_PROVE "x <= y | y <= (x::real)") 4);
-by (etac disjE 4);
-by (REPEAT(arith_tac 1));
-by (cut_inst_tac [("f","f"),("a","x - d"),("b","x"),("y","f(x + d)")]
-    IVT_objl 1);
-by (Step_tac 1);
-by (arith_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);
-by (dres_inst_tac [("f","g")] arg_cong 1);
-by (rotate_tac 2 1);
-by (forw_inst_tac [("x","xa")] spec 1);
-by (dres_inst_tac [("x","x + d")] spec 1);
-by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);
-(* 2nd case: similar *)
-by (cut_inst_tac [("f","f"),("a","x"),("b","x + d"),("y","f(x - d)")]
-    IVT2_objl 1);
-by (Step_tac 1);
-by (arith_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);
-by (dres_inst_tac [("f","g")] arg_cong 1);
-by (rotate_tac 2 1);
-by (forw_inst_tac [("x","xa")] spec 1);
-by (dres_inst_tac [("x","x - d")] spec 1);
-by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);
-qed "lemma_isCont_inj";
-
-(* ------------------------------------------------------------------------ *)
-(* Similar version for lower bound                                          *)
-(* ------------------------------------------------------------------------ *)
-
-Goal "[|#0 < d; ALL z. abs(z - x) <= d --> g(f z) = z; \
-\       ALL z. abs(z - x) <= d --> isCont f z |]  \
-\     ==> ~(ALL z. abs(z - x) <= d --> f(x) <= f(z))";
-by (auto_tac (claset() addSDs [(asm_full_simplify (simpset()) 
-    (read_instantiate [("f","%x. - f x"),("g","%y. g(-y)"),("x","x"),("d","d")]
-     lemma_isCont_inj))],simpset() addsimps [isCont_minus]));
-qed "lemma_isCont_inj2";
-
-(* ------------------------------------------------------------------------ *)
-(* Show there's an interval surrounding f(x) in f[[x - d, x + d]]           *)
-(* Also from John's theory                                                  *)
-(* ------------------------------------------------------------------------ *)
-
-Addsimps [zero_eq_numeral_0,one_eq_numeral_1];
-
-val lemma_le = ARITH_PROVE "#0 <= (d::real) ==> -d <= d";
-
-(* FIXME: awful proof - needs improvement *)
-Goal "[| #0 < d; ALL z. abs(z - x) <= d --> g(f z) = z; \
-\        ALL z. abs(z - x) <= d --> isCont f z |] \
-\      ==> EX e. #0 < e & \
-\                 (ALL y. \
-\                     abs(y - f(x)) <= e --> \
-\                     (EX z. abs(z - x) <= d & (f z = y)))";
-by (ftac real_less_imp_le 1);
-by (dtac (lemma_le RS (asm_full_simplify (simpset()) (read_instantiate 
-    [("f","f"),("a","x - d"),("b","x + d")] isCont_Lb_Ub))) 1);
-by (Step_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);
-by (subgoal_tac "L <= f x & f x <= M" 1);
-by (dres_inst_tac [("P", "%v. ?P v --> ?Q v & ?R v"), ("x","x")] spec 2);
-by (Asm_full_simp_tac 2);
-by (subgoal_tac "L < f x & f x < M" 1);
-by (Step_tac 1);
-by (dres_inst_tac [("x","L")] (ARITH_PROVE "x < y ==> #0 < y - (x::real)") 1);
-by (dres_inst_tac [("x","f x")] (ARITH_PROVE "x < y ==> #0 < y - (x::real)") 1);
-by (dres_inst_tac [("d1.0","f x - L"),("d2.0","M - f x")] 
-    (rename_numerals real_lbound_gt_zero) 1);
-by (Step_tac 1);
-by (res_inst_tac [("x","e")] exI 1);
-by (Step_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);
-by (dres_inst_tac [("P","%v. ?PP v --> (EX xa. ?Q v xa)"),("x","y")] spec 1);
-by (Step_tac 1 THEN REPEAT(arith_tac 1));
-by (res_inst_tac [("x","xa")] exI 1);
-by (arith_tac 1);
-by (ALLGOALS(etac (ARITH_PROVE "[|x <= y; x ~= y |] ==> x < (y::real)")));
-by (ALLGOALS(rotate_tac 3));
-by (dtac lemma_isCont_inj2 1);
-by (assume_tac 2);
-by (dtac lemma_isCont_inj 3);
-by (assume_tac 4);
-by (TRYALL(assume_tac));
-by (Step_tac 1);
-by (ALLGOALS(dres_inst_tac [("x","z")] spec));
-by (ALLGOALS(arith_tac));
-qed "isCont_inj_range";
-
-
-(* ------------------------------------------------------------------------ *)
-(* Continuity of inverse function                                           *)
-(* ------------------------------------------------------------------------ *)
-
-Goal "[| #0 < d; ALL z. abs(z - x) <= d --> g(f(z)) = z; \
-\        ALL z. abs(z - x) <= d --> isCont f z |] \
-\     ==> isCont g (f x)";
-by (simp_tac (simpset() addsimps [isCont_iff,LIM_def]) 1);
-by (Step_tac 1);
-by (dres_inst_tac [("d1.0","r")] (rename_numerals real_lbound_gt_zero) 1);
-by (assume_tac 1 THEN Step_tac 1);
-by (subgoal_tac "ALL z. abs(z - x) <= e --> (g(f z) = z)" 1);
-by (Force_tac 2);
-by (subgoal_tac "ALL z. abs(z - x) <= e --> isCont f z" 1);
-by (Force_tac 2);
-by (dres_inst_tac [("d","e")] isCont_inj_range 1);
-by (assume_tac 2 THEN assume_tac 1);
-by (Step_tac 1);
-by (res_inst_tac [("x","ea")] exI 1);
-by Auto_tac;
-by (rotate_tac 4 1);
-by (dres_inst_tac [("x","f(x) + xa")] spec 1); 
-by Auto_tac;
-by (dtac sym 1 THEN Auto_tac);
-by (arith_tac 1);
-qed "isCont_inverse";
-
--- a/src/HOL/Real/Hyperreal/Lim.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,79 +0,0 @@
-(*  Title       : Lim.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Theory of limits, continuity and 
-                  differentiation of real=>real functions
-*)
-
-Lim = SEQ + RealAbs + 
-
-(*-----------------------------------------------------------------------
-    Limits, continuity and differentiation: standard and NS definitions
- -----------------------------------------------------------------------*)
-
-constdefs
-  LIM :: [real=>real,real,real] => bool
-				("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60)
-  "f -- a --> L ==
-     ALL r. #0 < r --> 
-	     (EX s. #0 < s & (ALL x. (x ~= a & (abs(x + -a) < s)
-			  --> abs(f x + -L) < r)))"
-
-  NSLIM :: [real=>real,real,real] => bool
-			      ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60)
-  "f -- a --NS> L == (ALL x. (x ~= hypreal_of_real a & 
-		      x @= hypreal_of_real a -->
-		      (*f* f) x @= hypreal_of_real L))"   
-
-  isCont :: [real=>real,real] => bool
-  "isCont f a == (f -- a --> (f a))"        
-
-  (* NS definition dispenses with limit notions *)
-  isNSCont :: [real=>real,real] => bool
-  "isNSCont f a == (ALL y. y @= hypreal_of_real a --> 
-			   (*f* f) y @= hypreal_of_real (f a))"
-
-  (* differentiation: D is derivative of function f at x *)
-  deriv:: [real=>real,real,real] => bool
-			    ("(DERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60)
-  "DERIV f x :> D == ((%h. (f(x + h) + -f(x))/h) -- #0 --> D)"
-
-  nsderiv :: [real=>real,real,real] => bool
-			    ("(NSDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60)
-  "NSDERIV f x :> D == (ALL h: Infinitesimal - {0}. 
-			((*f* f)(hypreal_of_real x + h) + 
-			 - hypreal_of_real (f x))/h @= hypreal_of_real D)"
-
-  differentiable :: [real=>real,real] => bool   (infixl 60)
-  "f differentiable x == (EX D. DERIV f x :> D)"
-
-  NSdifferentiable :: [real=>real,real] => bool   (infixl 60)
-  "f NSdifferentiable x == (EX D. NSDERIV f x :> D)"
-
-  increment :: [real=>real,real,hypreal] => hypreal
-  "increment f x h == (@inc. f NSdifferentiable x & 
-		       inc = (*f* f)(hypreal_of_real x + h) + -hypreal_of_real (f x))"
-
-  isUCont :: (real=>real) => bool
-  "isUCont f ==  (ALL r. #0 < r --> 
-		      (EX s. #0 < s & (ALL x y. abs(x + -y) < s
-			    --> abs(f x + -f y) < r)))"
-
-  isNSUCont :: (real=>real) => bool
-  "isNSUCont f == (ALL x y. x @= y --> (*f* f) x @= (*f* f) y)"
-
-
-(*Used in the proof of the Bolzano theorem*)
-consts
-  Bolzano_bisect :: "[real*real=>bool, real, real, nat] => (real*real)"
-
-primrec
-  "Bolzano_bisect P a b 0 = (a,b)"
-  "Bolzano_bisect P a b (Suc n) =
-      (let (x,y) = Bolzano_bisect P a b n
-       in if P(x, (x+y)/#2) then ((x+y)/#2, y)
-                            else (x, (x+y)/#2) )"
-  
-
-end
-
--- a/src/HOL/Real/Hyperreal/NSA.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2291 +0,0 @@
-(*  Title       : NSA.ML
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Infinite numbers, Infinitesimals,
-                  infinitely close relation  etc.
-*) 
-
-fun CLAIM_SIMP str thms =
-               prove_goal (the_context()) str
-               (fn prems => [cut_facts_tac prems 1,
-                   auto_tac (claset(),simpset() addsimps thms)]);
-fun CLAIM str = CLAIM_SIMP str [];
-
-(*--------------------------------------------------------------------
-     Closure laws for members of (embedded) set standard real SReal
- --------------------------------------------------------------------*)
-
-Goalw [SReal_def] "[| x: SReal; y: SReal |] ==> x + y: SReal";
-by (Step_tac 1);
-by (res_inst_tac [("x","r + ra")] exI 1);
-by (Simp_tac 1);
-qed "SReal_add";
-
-Goalw [SReal_def] "[| x: SReal; y: SReal |] ==> x * y: SReal";
-by (Step_tac 1);
-by (res_inst_tac [("x","r * ra")] exI 1);
-by (simp_tac (simpset() addsimps [hypreal_of_real_mult]) 1);
-qed "SReal_mult";
-
-Goalw [SReal_def] "x: SReal ==> inverse x : SReal";
-by (blast_tac (claset() addIs [hypreal_of_real_inverse RS sym]) 1); 
-qed "SReal_inverse";
-
-Goalw [SReal_def] "x: SReal ==> -x : SReal";
-by (blast_tac (claset() addIs [hypreal_of_real_minus RS sym]) 1); 
-qed "SReal_minus";
-
-Goal "[| x + y : SReal; y: SReal |] ==> x: SReal";
-by (dres_inst_tac [("x","y")] SReal_minus 1);
-by (dtac SReal_add 1);
-by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc]));
-qed "SReal_add_cancel";
-
-Goalw [SReal_def] "x: SReal ==> abs x : SReal";
-by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_hrabs]));
-qed "SReal_hrabs";
-
-Goalw [SReal_def] "hypreal_of_real #1 : SReal";
-by (auto_tac (claset() addIs [hypreal_of_real_one RS sym], simpset()));  
-qed "SReal_hypreal_of_real_one";
-
-Goalw [SReal_def] "hypreal_of_real 0 : SReal";
-by (auto_tac (claset() addIs [hypreal_of_real_zero RS sym], simpset()));  
-qed "SReal_hypreal_of_real_zero";
-
-Goal "1hr : SReal";
-by (simp_tac (simpset() addsimps [SReal_hypreal_of_real_one,
-                                  hypreal_of_real_one RS sym]) 1);
-qed "SReal_one";
-
-Goal "0 : SReal";
-by (simp_tac (simpset() addsimps [rename_numerals SReal_hypreal_of_real_zero,
-                                  hypreal_of_real_zero RS sym]) 1);
-qed "SReal_zero";
-
-Goal "1hr + 1hr : SReal";
-by (rtac ([SReal_one,SReal_one] MRS SReal_add) 1);
-qed "SReal_two";
-
-Addsimps [SReal_zero,SReal_two];
-
-Goalw [hypreal_divide_def] "r : SReal ==> r/(1hr + 1hr): SReal";
-by (blast_tac (claset() addSIs [SReal_mult, SReal_inverse, SReal_two]) 1);
-qed "SReal_half";
-
-(* in general: move before previous thms!*)
-Goalw [SReal_def] "hypreal_of_real x: SReal";
-by (Blast_tac 1);
-qed "SReal_hypreal_of_real";
-
-Addsimps [SReal_hypreal_of_real];
-
-(* Infinitesimal ehr not in SReal *)
-
-Goalw [SReal_def] "ehr ~: SReal";
-by (auto_tac (claset(),
-              simpset() addsimps [hypreal_of_real_not_eq_epsilon RS not_sym]));
-qed "SReal_epsilon_not_mem";
-
-Goalw [SReal_def] "whr ~: SReal";
-by (auto_tac (claset(),
-              simpset() addsimps [hypreal_of_real_not_eq_omega RS not_sym]));
-qed "SReal_omega_not_mem";
-
-Goalw [SReal_def] "{x. hypreal_of_real x : SReal} = (UNIV::real set)";
-by Auto_tac;
-qed "SReal_UNIV_real";
-
-Goalw [SReal_def] "(x: SReal) = (EX y. x = hypreal_of_real  y)";
-by Auto_tac;
-qed "SReal_iff";
-
-Goalw [SReal_def] "hypreal_of_real ``(UNIV::real set) = SReal";
-by Auto_tac;
-qed "hypreal_of_real_image";
-
-Goalw [SReal_def] "inv hypreal_of_real ``SReal = (UNIV::real set)";
-by Auto_tac;
-by (rtac (inj_hypreal_of_real RS inv_f_f RS subst) 1);
-by (Blast_tac 1);
-qed "inv_hypreal_of_real_image";
-
-Goalw [SReal_def] 
-      "[| EX x. x: P; P <= SReal |] ==> EX Q. P = hypreal_of_real `` Q";
-by (Best_tac 1); 
-qed "SReal_hypreal_of_real_image";
-
-Goal "[| x: SReal; y: SReal; x < y |] ==> EX r: SReal. x < r & r < y";
-by (auto_tac (claset(),simpset() addsimps [SReal_iff]));
-by (dtac real_dense 1 THEN Step_tac 1);
-by (res_inst_tac [("x","hypreal_of_real r")] bexI 1);
-by Auto_tac;
-qed "SReal_dense";
-
-(*------------------------------------------------------------------
-                   Completeness of SReal
- ------------------------------------------------------------------*)
-Goal "P <= SReal ==> ((EX x:P. y < x) = \ 
-\     (EX X. hypreal_of_real X : P & y < hypreal_of_real X))";
-by (blast_tac (claset() addSDs [SReal_iff RS iffD1]) 1);
-by (flexflex_tac );
-qed "SReal_sup_lemma";
-
-Goal "[| P <= SReal; EX x. x: P; EX y : SReal. ALL x: P. x < y |] \
-\     ==> (EX X. X: {w. hypreal_of_real w : P}) & \
-\         (EX Y. ALL X: {w. hypreal_of_real w : P}. X < Y)";
-by (rtac conjI 1);
-by (fast_tac (claset() addSDs [SReal_iff RS iffD1]) 1);
-by (Auto_tac THEN forward_tac [subsetD] 1 THEN assume_tac 1);
-by (dtac (SReal_iff RS iffD1) 1);
-by (Auto_tac THEN res_inst_tac [("x","ya")] exI 1);
-by (auto_tac (claset() addDs [bspec],
-   simpset() addsimps [hypreal_of_real_less_iff RS sym]));
-qed "SReal_sup_lemma2";
-
-(*------------------------------------------------------
-    lifting of ub and property of lub
- -------------------------------------------------------*)
-Goalw [isUb_def,setle_def] 
-      "(isUb (SReal) (hypreal_of_real `` Q) (hypreal_of_real Y)) = \
-\      (isUb (UNIV :: real set) Q Y)";
-by (blast_tac (claset() addIs [hypreal_of_real_le_iff RS iffD1,
-                hypreal_of_real_le_iff RS iffD2,SReal_hypreal_of_real]) 1);
-qed "hypreal_of_real_isUb_iff";
-
-Goalw [isLub_def,leastP_def] 
-      "isLub SReal (hypreal_of_real `` Q) (hypreal_of_real Y) \
-\      ==> isLub (UNIV :: real set) Q Y";
-by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_isUb_iff,
-               setge_def]));
-by (blast_tac (claset() addIs [hypreal_of_real_isUb_iff RS iffD2,
-               hypreal_of_real_le_iff RS iffD2]) 1);
-qed "hypreal_of_real_isLub1";
-
-Goalw [isLub_def,leastP_def] 
-      "isLub (UNIV :: real set) Q Y \
-\      ==> isLub SReal (hypreal_of_real `` Q) (hypreal_of_real Y)";
-by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_isUb_iff,
-               setge_def]));
-by (forw_inst_tac [("x2","x")] (isUbD2a RS (SReal_iff RS iffD1) RS exE) 1);
-by (assume_tac 2);
-by (dres_inst_tac [("x","xa")] spec 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_isUb_iff,
-    hypreal_of_real_le_iff RS iffD1]));
-qed "hypreal_of_real_isLub2";
-
-Goal
-    "(isLub SReal (hypreal_of_real `` Q) (hypreal_of_real Y)) = \
-\    (isLub (UNIV :: real set) Q Y)";
-by (blast_tac (claset() addIs [hypreal_of_real_isLub1,hypreal_of_real_isLub2]) 1);
-qed "hypreal_of_real_isLub_iff";
-
-(* lemmas *)
-Goalw [isUb_def] 
-     "isUb SReal P Y \
-\     ==> EX Yo. isUb SReal P (hypreal_of_real Yo)";
-by (auto_tac (claset(),simpset() addsimps [SReal_iff]));
-qed "lemma_isUb_hypreal_of_real";
-
-Goalw [isLub_def,leastP_def,isUb_def] 
-      "isLub SReal P Y ==> EX Yo. isLub SReal P (hypreal_of_real Yo)";
-by (auto_tac (claset(),simpset() addsimps [SReal_iff]));
-qed "lemma_isLub_hypreal_of_real";
-
-Goalw [isLub_def,leastP_def,isUb_def] 
-      "EX Yo. isLub SReal P (hypreal_of_real Yo) \
-\      ==> EX Y. isLub SReal P Y";
-by Auto_tac;
-qed "lemma_isLub_hypreal_of_real2";
-
-Goal "[| P <= SReal; EX x. x: P; \
-\        EX Y. isUb SReal P Y |] \
-\     ==> EX t. isLub SReal P t";
-by (forward_tac [SReal_hypreal_of_real_image] 1);
-by (Auto_tac THEN dtac lemma_isUb_hypreal_of_real 1);
-by (auto_tac (claset() addSIs [reals_complete,
-    lemma_isLub_hypreal_of_real2], simpset() addsimps 
-    [hypreal_of_real_isLub_iff,hypreal_of_real_isUb_iff]));
-qed "SReal_complete";
-
-(*--------------------------------------------------------------------
-        Set of finite elements is a subring of the extended reals
- --------------------------------------------------------------------*)
-Goalw [HFinite_def] "[|x : HFinite;y : HFinite|] ==> (x+y) : HFinite";
-by (blast_tac (claset() addSIs [SReal_add,hrabs_add_less]) 1);
-qed "HFinite_add";
-
-Goalw [HFinite_def] "[|x : HFinite; y : HFinite|] ==> (x*y) : HFinite";
-by (blast_tac (claset() addSIs [SReal_mult,hrabs_mult_less]) 1);
-qed "HFinite_mult";
-
-Goalw [HFinite_def] "(x:HFinite)=(-x:HFinite)";
-by (Simp_tac 1);
-qed "HFinite_minus_iff";
-
-Goal "[| x: HFinite; y: HFinite|] ==> (x + -y): HFinite";
-by (blast_tac (claset() addDs [HFinite_minus_iff RS iffD1] 
-                        addIs [HFinite_add]) 1);
-qed "HFinite_add_minus";
-
-Goalw [SReal_def,HFinite_def] "SReal <= HFinite";
-by Auto_tac;
-by (res_inst_tac [("x","1hr + abs(hypreal_of_real r)")] exI 1);
-by (auto_tac (claset(),
-      simpset() addsimps [hypreal_of_real_hrabs, hypreal_zero_less_one]));
-by (res_inst_tac [("x","#1 + abs r")] exI 1);
-by (simp_tac (simpset() addsimps [hypreal_of_real_one,  
-                                  hypreal_of_real_zero]) 1);
-qed "SReal_subset_HFinite";
-
-Goal "hypreal_of_real x : HFinite";
-by (auto_tac (claset() addIs [(SReal_subset_HFinite RS subsetD)],
-              simpset()));
-qed "HFinite_hypreal_of_real";
-
-Addsimps [HFinite_hypreal_of_real];
-
-Goalw [HFinite_def] "x : HFinite ==> EX t: SReal. abs x < t";
-by Auto_tac;
-qed "HFiniteD";
-
-Goalw [HFinite_def] "x : HFinite ==> abs x : HFinite";
-by (auto_tac (claset(),simpset() addsimps [hrabs_idempotent]));
-qed "HFinite_hrabs";
-
-Goalw [HFinite_def,Bex_def] "x ~: HFinite ==> abs x ~: HFinite";
-by (auto_tac (claset() addDs [spec],simpset() addsimps [hrabs_idempotent]));
-qed "not_HFinite_hrabs";
-
-Goal "0 : HFinite";
-by (rtac (SReal_zero RS (SReal_subset_HFinite RS subsetD)) 1);
-qed "HFinite_zero";
-Addsimps [HFinite_zero];
-
-Goal "1hr : HFinite";
-by (rtac (SReal_one RS (SReal_subset_HFinite RS subsetD)) 1);
-qed "HFinite_one";
-Addsimps [HFinite_one];
-
-Goal "[|x : HFinite; y <= x; 0 <= y |] ==> y: HFinite";
-by (case_tac "x <= 0" 1);
-by (dres_inst_tac [("j","x")] hypreal_le_trans 1);
-by (dtac hypreal_le_anti_sym 2);
-by (auto_tac (claset() addSDs [not_hypreal_leE],simpset()));
-by (auto_tac (claset() addSIs [bexI]
-    addIs [hypreal_le_less_trans],simpset() addsimps 
-    [hrabs_eqI1,hrabs_eqI2,hrabs_minus_eqI1,HFinite_def]));
-qed "HFinite_bounded"; 
-
-(*------------------------------------------------------------------
-       Set of infinitesimals is a subring of the hyperreals   
- ------------------------------------------------------------------*)
-Goalw [Infinitesimal_def]
-      "x : Infinitesimal ==> ALL r: SReal. 0 < r --> abs x < r";
-by Auto_tac;
-qed "InfinitesimalD";
-
-Goalw [Infinitesimal_def] "0 : Infinitesimal";
-by (simp_tac (simpset() addsimps [hrabs_zero]) 1);
-qed "Infinitesimal_zero";
-AddIffs [Infinitesimal_zero];
-
-Goalw [Infinitesimal_def] 
-      "[| x : Infinitesimal; y : Infinitesimal |] ==> (x + y) : Infinitesimal";
-by Auto_tac;
-by (rtac (hypreal_sum_of_halves RS subst) 1);
-by (dtac hypreal_half_gt_zero 1);
-by (dtac SReal_half 1);
-by (auto_tac (claset() addSDs [bspec] addIs [hrabs_add_less], simpset()));
-qed "Infinitesimal_add";
-
-Goalw [Infinitesimal_def] 
-     "(x:Infinitesimal) = (-x:Infinitesimal)";
-by (Full_simp_tac 1);
-qed "Infinitesimal_minus_iff";
-
-Goal "x ~:Infinitesimal = (-x ~:Infinitesimal)";
-by (full_simp_tac (simpset() addsimps [Infinitesimal_minus_iff RS sym]) 1);
-qed "not_Infinitesimal_minus_iff";
-
-Goal "[| x : Infinitesimal; y : Infinitesimal |] ==> (x + -y):Infinitesimal";
-by (asm_simp_tac
-    (simpset() addsimps [Infinitesimal_minus_iff RS sym, 
-                          Infinitesimal_add]) 1);
-qed "Infinitesimal_add_minus";
-
-Goalw [Infinitesimal_def] 
-      "[| x : Infinitesimal; y : Infinitesimal |] \
-\      ==> (x * y) : Infinitesimal";
-by Auto_tac;
-by (rtac (hypreal_mult_1_right RS subst) 1);
-by (rtac hrabs_mult_less 1);
-by (cut_facts_tac [SReal_one,hypreal_zero_less_one] 2);
-by (ALLGOALS(blast_tac (claset() addSDs [bspec])));
-qed "Infinitesimal_mult";
-
-Goal "[| x : Infinitesimal; y : HFinite |] ==> (x * y) : Infinitesimal";
-by (auto_tac (claset() addSDs [HFiniteD],
-              simpset() addsimps [Infinitesimal_def]));
-by (forward_tac [hrabs_less_gt_zero] 1);
-by (dtac (hypreal_inverse_gt_zero RSN (2,hypreal_mult_less_mono2)) 1 
-    THEN assume_tac 1);
-by (dtac SReal_inverse 1);
-by (dtac SReal_mult 1 THEN assume_tac 1);
-by (thin_tac "inverse t : SReal" 1);
-by (auto_tac (claset() addSDs [bspec], simpset() addsimps [hrabs_mult]));
-by (dtac hrabs_mult_less 1 THEN assume_tac 1);
-by (dtac (hypreal_not_refl2 RS not_sym) 1);
-by (auto_tac (claset() addSDs [hypreal_mult_inverse],
-              simpset() addsimps [hrabs_mult] @ hypreal_mult_ac));
-qed "Infinitesimal_HFinite_mult";
-
-Goal "[| x : Infinitesimal; y : HFinite |] \
-\     ==> (y * x) : Infinitesimal";
-by (auto_tac (claset() addDs [Infinitesimal_HFinite_mult],
-              simpset() addsimps [hypreal_mult_commute]));
-qed "Infinitesimal_HFinite_mult2";
-
-(*** rather long proof ***)
-Goalw [HInfinite_def,Infinitesimal_def] 
-     "x: HInfinite ==> inverse x: Infinitesimal";
-by Auto_tac;
-by (eres_inst_tac [("x","inverse r")] ballE 1);
-by (rtac (hrabs_inverse RS ssubst) 1);
-by (forw_inst_tac [("x1","r"),("R3.0","abs x")] 
-    (hypreal_inverse_gt_zero RS hypreal_less_trans) 1);
-by (assume_tac 1);
-by (forw_inst_tac [("s","abs x"),("t","0")] 
-              (hypreal_not_refl2 RS not_sym) 1);
-by (dtac ((hypreal_inverse_inverse RS sym) RS subst) 1);
-by (dres_inst_tac [("x","abs x")] hypreal_inverse_gt_zero 1);
-by (rtac (hypreal_inverse_less_iff RS ssubst) 1);
-by (auto_tac (claset() addSDs [SReal_inverse],
-    simpset() addsimps [hypreal_not_refl2 RS not_sym, hypreal_less_not_refl]));
-qed "HInfinite_inverse_Infinitesimal";
-
-Goalw [HInfinite_def] 
-   "[|x: HInfinite;y: HInfinite|] ==> (x*y): HInfinite";
-by Auto_tac;
-by (cut_facts_tac [SReal_one] 1);
-by (eres_inst_tac [("x","1hr")] ballE 1);
-by (eres_inst_tac [("x","r")] ballE 1);
-by (REPEAT(fast_tac (HOL_cs addSIs [hrabs_mult_gt]) 1));
-qed "HInfinite_mult";
-
-Goalw [HInfinite_def] 
-      "[|x: HInfinite; 0 <= y; 0 <= x|] ==> (x + y): HInfinite";
-by (auto_tac (claset() addSIs  
-    [hypreal_add_zero_less_le_mono],
-    simpset() addsimps [hrabs_eqI1,
-    hypreal_add_commute,hypreal_le_add_order]));
-qed "HInfinite_add_ge_zero";
-
-Goal "[|x: HInfinite; 0 <= y; 0 <= x|] ==> (y + x): HInfinite";
-by (auto_tac (claset() addSIs [HInfinite_add_ge_zero],
-    simpset() addsimps [hypreal_add_commute]));
-qed "HInfinite_add_ge_zero2";
-
-Goal "[|x: HInfinite; 0 < y; 0 < x|] ==> (x + y): HInfinite";
-by (blast_tac (claset() addIs [HInfinite_add_ge_zero,
-    hypreal_less_imp_le]) 1);
-qed "HInfinite_add_gt_zero";
-
-Goalw [HInfinite_def] "(-x: HInfinite) = (x: HInfinite)";
-by Auto_tac;
-qed "HInfinite_minus_iff";
-
-Goal "[|x: HInfinite; y <= 0; x <= 0|] ==> (x + y): HInfinite";
-by (dtac (HInfinite_minus_iff RS iffD2) 1);
-by (rtac (HInfinite_minus_iff RS iffD1) 1);
-by (auto_tac (claset() addIs [HInfinite_add_ge_zero],
-              simpset() addsimps [hypreal_minus_zero_le_iff]));
-qed "HInfinite_add_le_zero";
-
-Goal "[|x: HInfinite; y < 0; x < 0|] ==> (x + y): HInfinite";
-by (blast_tac (claset() addIs [HInfinite_add_le_zero,
-                               hypreal_less_imp_le]) 1);
-qed "HInfinite_add_lt_zero";
-
-Goal "[|a: HFinite; b: HFinite; c: HFinite|] \ 
-\     ==> a*a + b*b + c*c : HFinite";
-by (auto_tac (claset() addIs [HFinite_mult,HFinite_add], simpset()));
-qed "HFinite_sum_squares";
-
-Goal "x ~: Infinitesimal ==> x ~= 0";
-by Auto_tac;
-qed "not_Infinitesimal_not_zero";
-
-Goal "x: HFinite - Infinitesimal ==> x ~= 0";
-by Auto_tac;
-qed "not_Infinitesimal_not_zero2";
-
-Goal "x : Infinitesimal ==> abs x : Infinitesimal";
-by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1);
-by (auto_tac (claset(),simpset() addsimps [Infinitesimal_minus_iff RS sym]));
-qed "Infinitesimal_hrabs";
-
-Goal "x ~: Infinitesimal ==> abs x ~: Infinitesimal";
-by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1);
-by (auto_tac (claset(),simpset() addsimps [Infinitesimal_minus_iff RS sym]));
-qed "not_Infinitesimal_hrabs";
-
-Goal "abs x : Infinitesimal ==> x : Infinitesimal";
-by (rtac ccontr 1);
-by (blast_tac (claset() addDs [not_Infinitesimal_hrabs]) 1);
-qed "hrabs_Infinitesimal_Infinitesimal";
-
-Goal "x : HFinite - Infinitesimal ==> abs x : HFinite - Infinitesimal";
-by (fast_tac (set_cs addSEs [HFinite_hrabs,not_Infinitesimal_hrabs]) 1);
-qed "HFinite_diff_Infinitesimal_hrabs";
-
-Goalw [Infinitesimal_def] 
-      "[| e : Infinitesimal; abs x < e |] ==> x : Infinitesimal";
-by (auto_tac (claset() addSDs [bspec],simpset()));
-by (dres_inst_tac [("i","e")] (hrabs_ge_self RS hypreal_le_less_trans) 1);
-by (fast_tac (claset() addIs [hypreal_less_trans]) 1);
-qed "hrabs_less_Infinitesimal";
-
-Goal 
- "[| e : Infinitesimal; abs x <= e |] ==> x : Infinitesimal";
-by (blast_tac (claset() addDs [hypreal_le_imp_less_or_eq] addIs
-    [hrabs_Infinitesimal_Infinitesimal,hrabs_less_Infinitesimal]) 1);
-qed "hrabs_le_Infinitesimal";
-
-Goalw [Infinitesimal_def] 
-      "[| e : Infinitesimal; \
-\              e' : Infinitesimal; \
-\              e' < x ; x < e |] ==> x : Infinitesimal";
-by (auto_tac (claset() addSDs [bspec],simpset()));
-by (dres_inst_tac [("i","e")] (hrabs_ge_self RS hypreal_le_less_trans) 1);
-by (dtac (hrabs_interval_iff RS iffD1 RS conjunct1) 1);
-by (fast_tac (claset() addIs [hypreal_less_trans,hrabs_interval_iff RS iffD2]) 1);
-qed "Infinitesimal_interval";
-
-Goal "[| e : Infinitesimal; e' : Infinitesimal; \
-\       e' <= x ; x <= e |] ==> x : Infinitesimal";
-by (auto_tac (claset() addIs [Infinitesimal_interval],
-    simpset() addsimps [hypreal_le_eq_less_or_eq]));
-qed "Infinitesimal_interval2";
-
-Goalw [Infinitesimal_def] 
-      "[| x ~: Infinitesimal; \
-\                 y ~: Infinitesimal|] \
-\               ==> (x*y) ~:Infinitesimal";
-by (Clarify_tac 1);
-by (eres_inst_tac [("x","r*ra")] ballE 1);
-by (fast_tac (claset() addIs [SReal_mult]) 2);
-by (auto_tac (claset(),simpset() addsimps [hrabs_mult]));
-by (blast_tac (claset() addSDs [hypreal_mult_order]) 1);
-by (REPEAT(dtac hypreal_leI 1));
-by (dtac hypreal_mult_le_ge_zero 1);
-by (dres_inst_tac [("j","r*ra")] hypreal_less_le_trans 4);
-by (auto_tac (claset(),simpset() addsimps [hypreal_less_not_refl]));
-qed "not_Infinitesimal_mult";
-
-Goal "x*y : Infinitesimal ==> x : Infinitesimal | y : Infinitesimal";
-by (rtac ccontr 1);
-by (dtac (de_Morgan_disj RS iffD1) 1);
-by (fast_tac (claset() addDs [not_Infinitesimal_mult]) 1);
-qed "Infinitesimal_mult_disj";
-
-Goal "x: HFinite-Infinitesimal ==> x ~= 0";
-by (Blast_tac 1);
-qed "HFinite_Infinitesimal_not_zero";
-
-Goal "[| x : HFinite - Infinitesimal; \
-\                  y : HFinite - Infinitesimal \
-\               |] ==> (x*y) : HFinite - Infinitesimal";
-by (Clarify_tac 1);
-by (blast_tac (claset() addDs [HFinite_mult,not_Infinitesimal_mult]) 1);
-qed "HFinite_Infinitesimal_diff_mult";
-
-Goalw [Infinitesimal_def,HFinite_def]  
-      "Infinitesimal <= HFinite";
-by (blast_tac (claset() addSIs [SReal_one] 
-    addSEs [(hypreal_zero_less_one RSN (2,impE))]) 1);
-qed "Infinitesimal_subset_HFinite";
-
-Goal "x: Infinitesimal ==> x * hypreal_of_real r : Infinitesimal";
-by (etac (HFinite_hypreal_of_real RSN 
-          (2,Infinitesimal_HFinite_mult)) 1);
-qed "Infinitesimal_hypreal_of_real_mult";
-
-Goal "x: Infinitesimal ==> hypreal_of_real r * x: Infinitesimal";
-by (etac (HFinite_hypreal_of_real RSN 
-          (2,Infinitesimal_HFinite_mult2)) 1);
-qed "Infinitesimal_hypreal_of_real_mult2";
-
-(*----------------------------------------------------------------------
-                   Infinitely close relation @=
- ----------------------------------------------------------------------*)
-
-Goalw [Infinitesimal_def,inf_close_def] 
-        "x:Infinitesimal = (x @= 0)";
-by (Simp_tac 1);
-qed "mem_infmal_iff";
-
-Goalw [inf_close_def]" (x @= y) = (x + -y @= 0)";
-by (Simp_tac 1);
-qed "inf_close_minus_iff";
-
-Goalw [inf_close_def]" (x @= y) = (-y + x @= 0)";
-by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
-qed "inf_close_minus_iff2";
-
-Goalw [inf_close_def,Infinitesimal_def]  "x @= x";
-by (Simp_tac 1);
-qed "inf_close_refl";
-
-Goalw [inf_close_def]  "x @= y ==> y @= x";
-by (rtac (hypreal_minus_distrib1 RS subst) 1);
-by (etac (Infinitesimal_minus_iff RS iffD1) 1);
-qed "inf_close_sym";
-
-val prem1::prem2::rest = 
-goalw thy [inf_close_def]  "[| x @= y; y @= z |] ==> x @= z";
-by (res_inst_tac [("y1","y")] (hypreal_add_minus_cancel1 RS subst) 1);
-by (simp_tac (simpset() addsimps [([prem1,prem2] MRS Infinitesimal_add)]) 1);
-qed "inf_close_trans";
-
-val prem1::prem2::rest = goal thy "[| r @= x; s @= x |] ==> r @= s";
-by (rtac ([prem1,prem2 RS inf_close_sym] MRS inf_close_trans) 1);
-qed "inf_close_trans2";
-
-val prem1::prem2::rest = goal thy "[| x @= r; x @= s|] ==> r @= s";
-by (rtac ([prem1 RS inf_close_sym,prem2] MRS inf_close_trans) 1);
-qed "inf_close_trans3";
-
-Goal "(x + -y : Infinitesimal) = (x @= y)";
-by (auto_tac (claset(),simpset() addsimps 
-    [inf_close_minus_iff RS sym,mem_infmal_iff]));
-qed "Infinitesimal_inf_close_minus";
-
-Goal "(x + -y : Infinitesimal) = (y @= x)";
-by (auto_tac (claset() addIs [inf_close_sym],
-    simpset() addsimps [inf_close_minus_iff RS sym,
-    mem_infmal_iff]));
-qed "Infinitesimal_inf_close_minus2";
-
-Goalw [monad_def] "(x @= y) = (monad(x)=monad(y))";
-by (auto_tac (claset() addDs [inf_close_sym] 
-    addSEs [inf_close_trans,equalityCE],
-    simpset() addsimps [inf_close_refl]));
-qed "inf_close_monad_iff";
-
-Goal "[| x: Infinitesimal; y: Infinitesimal |] ==> x @= y";
-by (asm_full_simp_tac (simpset() addsimps [mem_infmal_iff]) 1);
-by (fast_tac (claset() addIs [inf_close_trans2]) 1);
-qed "Infinitesimal_inf_close";
-
-val prem1::prem2::rest = 
-goalw thy [inf_close_def] "[| a @= b; c @= d |] ==> a+c @= b+d";
-by (rtac (hypreal_minus_add_distrib RS ssubst) 1);
-by (rtac (hypreal_add_assoc RS ssubst) 1);
-by (res_inst_tac [("y1","c")] (hypreal_add_left_commute RS subst) 1);
-by (rtac (hypreal_add_assoc RS subst) 1);
-by (rtac ([prem1,prem2] MRS Infinitesimal_add) 1);
-qed "inf_close_add";
-
-Goal "a @= b ==> -a @= -b";
-by (rtac ((inf_close_minus_iff RS iffD2) RS inf_close_sym) 1);
-by (dtac (inf_close_minus_iff RS iffD1) 1);
-by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
-qed "inf_close_minus";
-
-Goal "-a @= -b ==> a @= b";
-by (auto_tac (claset() addDs [inf_close_minus], simpset()));
-qed "inf_close_minus2";
-
-Goal "(-a @= -b) = (a @= b)";
-by (blast_tac (claset() addIs [inf_close_minus,inf_close_minus2]) 1);
-qed "inf_close_minus_cancel";
-
-Addsimps [inf_close_minus_cancel];
-
-Goal "[| a @= b; c @= d |] ==> a + -c @= b + -d";
-by (blast_tac (claset() addSIs [inf_close_add,inf_close_minus]) 1);
-qed "inf_close_add_minus";
-
-Goalw [inf_close_def] "[| a @= b; c: HFinite|] ==> a*c @= b*c"; 
-by (asm_full_simp_tac (simpset() addsimps [Infinitesimal_HFinite_mult,
-    hypreal_minus_mult_eq1,hypreal_add_mult_distrib RS sym] 
-    delsimps [hypreal_minus_mult_eq1 RS sym]) 1);
-qed "inf_close_mult1";
-
-Goal "[|a @= b; c: HFinite|] ==> c*a @= c*b"; 
-by (asm_simp_tac (simpset() addsimps [inf_close_mult1,hypreal_mult_commute]) 1);
-qed "inf_close_mult2";
-
-Goal "[|u @= v*x; x @= y; v: HFinite|] ==> u @= v*y";
-by (fast_tac (claset() addIs [inf_close_mult2,inf_close_trans]) 1);
-qed "inf_close_mult_subst";
-
-Goal "[| u @= x*v; x @= y; v: HFinite |] ==> u @= y*v";
-by (fast_tac (claset() addIs [inf_close_mult1,inf_close_trans]) 1);
-qed "inf_close_mult_subst2";
-
-Goal "[| u @= x*hypreal_of_real v; x @= y |] ==> u @= y*hypreal_of_real v";
-by (auto_tac (claset() addIs [inf_close_mult_subst2],simpset()));
-qed "inf_close_mult_subst_SReal";
-
-Goalw [inf_close_def]  "a = b ==> a @= b";
-by (Asm_simp_tac 1);
-qed "inf_close_eq_imp";
-
-Goal "x: Infinitesimal ==> -x @= x"; 
-by (fast_tac (HOL_cs addIs [Infinitesimal_minus_iff RS iffD1,
-    mem_infmal_iff RS iffD1,inf_close_trans2]) 1);
-qed "Infinitesimal_minus_inf_close";
-
-Goalw [inf_close_def]  "(EX y: Infinitesimal. x + -z = y) = (x @= z)";
-by (Blast_tac 1);
-qed "bex_Infinitesimal_iff";
-
-Goal "(EX y: Infinitesimal. x = z + y) = (x @= z)";
-by (asm_full_simp_tac (simpset() addsimps [bex_Infinitesimal_iff RS sym]) 1);
-by (Force_tac 1);
-qed "bex_Infinitesimal_iff2";
-
-Goal "[| y: Infinitesimal; x + y = z |] ==> x @= z";
-by (rtac (bex_Infinitesimal_iff RS iffD1) 1);
-by (dtac (Infinitesimal_minus_iff RS iffD1) 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_minus_add_distrib,
-    hypreal_add_assoc RS sym]));
-qed "Infinitesimal_add_inf_close";
-
-Goal "y: Infinitesimal ==> x @= x + y";
-by (rtac (bex_Infinitesimal_iff RS iffD1) 1);
-by (dtac (Infinitesimal_minus_iff RS iffD1) 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_minus_add_distrib,
-    hypreal_add_assoc RS sym]));
-qed "Infinitesimal_add_inf_close_self";
-
-Goal "y: Infinitesimal ==> x @= y + x";
-by (auto_tac (claset() addDs [Infinitesimal_add_inf_close_self],
-    simpset() addsimps [hypreal_add_commute]));
-qed "Infinitesimal_add_inf_close_self2";
-
-Goal "y: Infinitesimal ==> x @= x + -y";
-by (blast_tac (claset() addSIs [Infinitesimal_add_inf_close_self,
-    Infinitesimal_minus_iff RS iffD1]) 1);
-qed "Infinitesimal_add_minus_inf_close_self";
-
-Goal "[| y: Infinitesimal; x+y @= z|] ==> x @= z";
-by (dres_inst_tac [("x","x")] (Infinitesimal_add_inf_close_self RS inf_close_sym) 1);
-by (etac (inf_close_trans3 RS inf_close_sym) 1);
-by (assume_tac 1);
-qed "Infinitesimal_add_cancel";
-
-Goal "[| y: Infinitesimal; x @= z + y|] ==> x @= z";
-by (dres_inst_tac [("x","z")] (Infinitesimal_add_inf_close_self2  RS inf_close_sym) 1);
-by (etac (inf_close_trans3 RS inf_close_sym) 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
-by (etac inf_close_sym 1);
-qed "Infinitesimal_add_right_cancel";
-
-Goal "d + b  @= d + c ==> b @= c";
-by (dtac (inf_close_minus_iff RS iffD1) 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypreal_minus_add_distrib,inf_close_minus_iff RS sym] 
-    @ hypreal_add_ac) 1);
-qed "inf_close_add_left_cancel";
-
-Goal "b + d @= c + d ==> b @= c";
-by (rtac inf_close_add_left_cancel 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypreal_add_commute]) 1);
-qed "inf_close_add_right_cancel";
-
-Goal "b @= c ==> d + b @= d + c";
-by (rtac (inf_close_minus_iff RS iffD2) 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypreal_minus_add_distrib,inf_close_minus_iff RS sym] 
-    @ hypreal_add_ac) 1);
-qed "inf_close_add_mono1";
-
-Goal "b @= c ==> b + a @= c + a";
-by (asm_simp_tac (simpset() addsimps 
-    [hypreal_add_commute,inf_close_add_mono1]) 1);
-qed "inf_close_add_mono2";
-
-Goal "(a + b @= a + c) = (b @= c)";
-by (fast_tac (claset() addEs [inf_close_add_left_cancel,
-    inf_close_add_mono1]) 1);
-qed "inf_close_add_left_iff";
-
-Addsimps [inf_close_add_left_iff];
-
-Goal "(b + a @= c + a) = (b @= c)";
-by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
-qed "inf_close_add_right_iff";
-
-Addsimps [inf_close_add_right_iff];
-
-Goal "[| x: HFinite; x @= y |] ==> y: HFinite";
-by (dtac (bex_Infinitesimal_iff2 RS iffD2) 1);
-by (Step_tac 1);
-by (dtac (Infinitesimal_subset_HFinite RS subsetD 
-          RS (HFinite_minus_iff RS iffD1)) 1);
-by (dtac HFinite_add 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_add_assoc]));
-qed "inf_close_HFinite";
-
-Goal "x @= hypreal_of_real D ==> x: HFinite";
-by (rtac (inf_close_sym RSN (2,inf_close_HFinite)) 1);
-by Auto_tac;
-qed "inf_close_hypreal_of_real_HFinite";
-
-Goal "[|a @= b; c @= d; b: HFinite; d: HFinite|] ==> a*c @= b*d";
-by (rtac inf_close_trans 1); 
-by (rtac inf_close_mult2 2); 
-by (rtac inf_close_mult1 1);
-by (blast_tac (claset() addIs [inf_close_HFinite, inf_close_sym]) 2);  
-by Auto_tac;  
-qed "inf_close_mult_HFinite";
-
-Goal "[|a @= hypreal_of_real b; c @= hypreal_of_real d |] \
-\     ==> a*c @= hypreal_of_real b*hypreal_of_real d";
-by (blast_tac (claset() addSIs [inf_close_mult_HFinite,
-            inf_close_hypreal_of_real_HFinite,HFinite_hypreal_of_real]) 1);
-qed "inf_close_mult_hypreal_of_real";
-
-Goal "[| a: SReal; a ~= 0; a*x @= 0 |] ==> x @= 0";
-by (dtac (SReal_inverse RS (SReal_subset_HFinite RS subsetD)) 1);
-by (auto_tac (claset() addDs [inf_close_mult2],
-    simpset() addsimps [hypreal_mult_assoc RS sym]));
-qed "inf_close_SReal_mult_cancel_zero";
-
-(* REM comments: newly added *)
-Goal "[| a: SReal; x @= 0 |] ==> x*a @= 0";
-by (auto_tac (claset() addDs [(SReal_subset_HFinite RS subsetD),
-              inf_close_mult1],simpset()));
-qed "inf_close_mult_SReal1";
-
-Goal "[| a: SReal; x @= 0 |] ==> a*x @= 0";
-by (auto_tac (claset() addDs [(SReal_subset_HFinite RS subsetD),
-              inf_close_mult2],simpset()));
-qed "inf_close_mult_SReal2";
-
-Goal "[|a : SReal; a ~= 0 |] ==> (a*x @= 0) = (x @= 0)";
-by (blast_tac (claset() addIs [inf_close_SReal_mult_cancel_zero,
-    inf_close_mult_SReal2]) 1);
-qed "inf_close_mult_SReal_zero_cancel_iff";
-Addsimps [inf_close_mult_SReal_zero_cancel_iff];
-
-Goal "[| a: SReal; a ~= 0; a* w @= a*z |] ==> w @= z";
-by (dtac (SReal_inverse RS (SReal_subset_HFinite RS subsetD)) 1);
-by (auto_tac (claset() addDs [inf_close_mult2],
-    simpset() addsimps [hypreal_mult_assoc RS sym]));
-qed "inf_close_SReal_mult_cancel";
-
-Goal "[| a: SReal; a ~= 0|] ==> (a* w @= a*z) = (w @= z)";
-by (auto_tac (claset() addSIs [inf_close_mult2,SReal_subset_HFinite RS subsetD] 
-    addIs [inf_close_SReal_mult_cancel],simpset()));
-qed "inf_close_SReal_mult_cancel_iff1";
-Addsimps [inf_close_SReal_mult_cancel_iff1];
-
-Goal "[| z <= f; f @= g; g <= z |] ==> f @= z";
-by (dtac (bex_Infinitesimal_iff2 RS iffD2) 1);
-by Auto_tac;
-by (dres_inst_tac [("x","-g")] hypreal_add_left_le_mono1 1);
-by (dres_inst_tac [("x","-g")] hypreal_add_le_mono1 1);
-by (res_inst_tac [("y","-y")] Infinitesimal_add_cancel 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_add_assoc RS sym,
-    Infinitesimal_minus_iff RS sym]));
-by (rtac inf_close_sym 1);
-by (simp_tac (simpset() addsimps [hypreal_add_assoc]) 1);
-by (rtac (inf_close_minus_iff RS iffD2) 1);
-by (auto_tac (claset(),simpset() addsimps [mem_infmal_iff RS sym,
-    hypreal_add_commute]));
-by (rtac Infinitesimal_interval2 1 THEN Auto_tac);
-qed "inf_close_le_bound";
-
-Goal "[| z <= f; f @= g; g <= z |] ==> g @= z";
-by (rtac inf_close_trans3 1);
-by (auto_tac (claset() addIs [inf_close_le_bound],simpset()));
-qed "inf_close_le_bound2";
-
-(*-----------------------------------------------------------------
-    Zero is the only infinitesimal that is also a real
- -----------------------------------------------------------------*)
-
-Goalw [Infinitesimal_def] 
-      "[| x: SReal; y: Infinitesimal; 0 < x |] ==> y < x";
-by (rtac (hrabs_ge_self RS hypreal_le_less_trans) 1);
-by (Blast_tac 1);
-qed "Infinitesimal_less_SReal";
-
-Goal "y: Infinitesimal ==> ALL r: SReal. 0 < r --> y < r";
-by (blast_tac (claset() addIs [Infinitesimal_less_SReal]) 1);
-qed "Infinitesimal_less_SReal2";
-
-Goalw [Infinitesimal_def] 
-      "[| y: SReal; 0 < y|] ==> y ~: Infinitesimal";
-by (auto_tac (claset() addSDs [bspec] addSEs [hypreal_less_irrefl],
-    simpset() addsimps [hrabs_eqI2]));
-qed "SReal_not_Infinitesimal";
-
-(* [| y : SReal; 0 < y; y : Infinitesimal |] ==> R *)
-bind_thm("SReal_not_InfinitesimalE", (SReal_not_Infinitesimal RS notE));
-
-Goal "[| y : SReal; y < 0 |] ==> y ~:Infinitesimal";
-by (blast_tac (claset() addDs [(hypreal_minus_zero_less_iff RS iffD2),
-    SReal_minus,(Infinitesimal_minus_iff RS iffD1),SReal_not_Infinitesimal]) 1);
-qed "SReal_minus_not_Infinitesimal";
-
-(* [| y : SReal; y < 0; y : Infinitesimal |] ==> R *)
-bind_thm("SReal_minus_not_InfinitesimalE", (SReal_minus_not_Infinitesimal RS notE));
-
-Goal "SReal Int Infinitesimal = {0}";
-by (auto_tac (claset(),simpset() addsimps [SReal_zero]));
-by (cut_inst_tac [("x","x"),("y","0")] hypreal_linear 1);
-by (blast_tac (claset() addIs [SReal_not_InfinitesimalE,
-    SReal_minus_not_InfinitesimalE]) 1);
-qed "SReal_Int_Infinitesimal_zero";
-
-Goal "0 : (SReal Int Infinitesimal)";
-by (simp_tac (simpset() addsimps [SReal_zero]) 1);
-qed "zero_mem_SReal_Int_Infinitesimal";
-
-Goal "[| x: SReal; x: Infinitesimal|] ==> x = 0";
-by (cut_facts_tac [SReal_Int_Infinitesimal_zero] 1);
-by (Blast_tac 1);
-qed "SReal_Infinitesimal_zero";
-
-Goal "[| x : SReal; x ~= 0 |] ==> x : HFinite - Infinitesimal";
-by (auto_tac (claset() addDs [SReal_Infinitesimal_zero,
-                              SReal_subset_HFinite RS subsetD], simpset()));
-qed "SReal_HFinite_diff_Infinitesimal";
-
-Goal "hypreal_of_real x ~= 0 ==> hypreal_of_real x : HFinite - Infinitesimal";
-by (rtac SReal_HFinite_diff_Infinitesimal 1);
-by Auto_tac;
-qed "hypreal_of_real_HFinite_diff_Infinitesimal";
-
-Goal "(hypreal_of_real x : Infinitesimal) = (x=0)";
-by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_zero]));  
-by (rtac ccontr 1); 
-by (rtac (hypreal_of_real_HFinite_diff_Infinitesimal RS DiffD2) 1); 
-by Auto_tac;  
-qed "hypreal_of_real_Infinitesimal_iff_0";
-AddIffs [hypreal_of_real_Infinitesimal_iff_0];
-
-Goal "1hr+1hr ~: Infinitesimal";
-by (fast_tac (claset() addDs [SReal_two RS SReal_Infinitesimal_zero]
-                       addSEs [hypreal_two_not_zero RS notE]) 1);
-qed "two_not_Infinitesimal";
-
-Goal "1hr ~: Infinitesimal";
-by (auto_tac (claset() addSDs [SReal_one RS SReal_Infinitesimal_zero],
-                simpset() addsimps [hypreal_zero_not_eq_one RS not_sym]));
-qed "hypreal_one_not_Infinitesimal";
-Addsimps [hypreal_one_not_Infinitesimal];
-
-Goal "[| y: SReal; x @= y; y~= 0 |] ==> x ~= 0";
-by (cut_inst_tac [("x","y")] hypreal_trichotomy 1);
-by (blast_tac (claset() addDs 
-		[inf_close_sym RS (mem_infmal_iff RS iffD2),
-		 SReal_not_Infinitesimal,SReal_minus_not_Infinitesimal]) 1);
-qed "inf_close_SReal_not_zero";
-
-Goal "[| x @= hypreal_of_real y; hypreal_of_real y ~= 0 |] ==> x ~= 0";
-by (rtac inf_close_SReal_not_zero 1);
-by Auto_tac;
-qed "inf_close_hypreal_of_real_not_zero";
-
-Goal "[| x @= y; y : HFinite - Infinitesimal |] \
-\     ==> x : HFinite - Infinitesimal";
-by (auto_tac (claset() addIs [inf_close_sym RSN (2,inf_close_HFinite)],
-              simpset() addsimps [mem_infmal_iff]));
-by (dtac inf_close_trans3 1 THEN assume_tac 1);
-by (blast_tac (claset() addDs [inf_close_sym]) 1);
-qed "HFinite_diff_Infinitesimal_inf_close";
-
-(*The premise y~=0 is essential; otherwise x/y =0 and we lose the 
-  HFinite premise.*)
-Goal "[| y ~= 0;  y: Infinitesimal;  x/y : HFinite |] ==> x : Infinitesimal";
-by (dtac Infinitesimal_HFinite_mult2 1);
-by (assume_tac 1);
-by (asm_full_simp_tac 
-    (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 1);
-qed "Infinitesimal_ratio";
-
-(*------------------------------------------------------------------
-       Standard Part Theorem: Every finite x: R* is infinitely 
-       close to a unique real number (i.e a member of SReal)
- ------------------------------------------------------------------*)
-(*------------------------------------------------------------------
-         Uniqueness: Two infinitely close reals are equal
- ------------------------------------------------------------------*)
-
-Goal "[|x: SReal; y: SReal|] ==> (x @= y) = (x = y)"; 
-by (auto_tac (claset(),simpset() addsimps [inf_close_refl]));
-by (rewrite_goals_tac [inf_close_def]);
-by (dres_inst_tac [("x","y")] SReal_minus 1);
-by (dtac SReal_add 1 THEN assume_tac 1);
-by (dtac SReal_Infinitesimal_zero 1 THEN assume_tac 1);
-by (dtac sym 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_eq_minus_iff RS sym]) 1);
-qed "SReal_inf_close_iff";
-
-Goal "(hypreal_of_real k @= hypreal_of_real m) = (k = m)";
-by (auto_tac (claset(),simpset() addsimps [inf_close_refl,
-    SReal_inf_close_iff,inj_hypreal_of_real RS injD]));
-qed "hypreal_of_real_inf_close_iff";
- 
-Addsimps [hypreal_of_real_inf_close_iff];
-
-Goal "[| r: SReal; s: SReal; r @= x; s @= x|] ==> r = s";
-by (blast_tac (claset() addIs [(SReal_inf_close_iff RS iffD1),
-               inf_close_trans2]) 1);
-qed "inf_close_unique_real";
-
-(*------------------------------------------------------------------
-       Existence of unique real infinitely close
- ------------------------------------------------------------------*)
-(* lemma about lubs *)
-Goal "!!(x::hypreal). [| isLub R S x; isLub R S y |] ==> x = y";
-by (forward_tac [isLub_isUb] 1);
-by (forw_inst_tac [("x","y")] isLub_isUb 1);
-by (blast_tac (claset() addSIs [hypreal_le_anti_sym]
-                addSDs [isLub_le_isUb]) 1);
-qed "hypreal_isLub_unique";
-
-Goal "x: HFinite ==> EX u. isUb SReal {s. s: SReal & s < x} u";
-by (dtac HFiniteD 1 THEN Step_tac 1);
-by (rtac exI 1 THEN rtac isUbI 1 THEN assume_tac 2);
-by (auto_tac (claset() addIs [hypreal_less_imp_le,setleI,isUbI,
-    hypreal_less_trans],simpset() addsimps [hrabs_interval_iff]));
-qed "lemma_st_part_ub";
-
-Goal "x: HFinite ==> EX y. y: {s. s: SReal & s < x}";
-by (dtac HFiniteD 1 THEN Step_tac 1);
-by (dtac SReal_minus 1);
-by (auto_tac (claset(),simpset() addsimps [hrabs_interval_iff]));
-qed "lemma_st_part_nonempty";
-
-Goal "{s. s: SReal & s < x} <= SReal";
-by Auto_tac;
-qed "lemma_st_part_subset";
-
-Goal "x: HFinite ==> EX t. isLub SReal {s. s: SReal & s < x} t";
-by (blast_tac (claset() addSIs [SReal_complete,lemma_st_part_ub,
-    lemma_st_part_nonempty, lemma_st_part_subset]) 1);
-qed "lemma_st_part_lub";
-
-Goal "((t::hypreal) + r <= t) = (r <= 0)";
-by (Step_tac 1);
-by (dres_inst_tac [("x","-t")] hypreal_add_left_le_mono1 1);
-by (dres_inst_tac [("x","t")] hypreal_add_left_le_mono1 2);
-by (auto_tac (claset(),simpset() addsimps [hypreal_add_assoc RS sym]));
-qed "lemma_hypreal_le_left_cancel";
-
-Goal "[| x: HFinite; \
-\                 isLub SReal {s. s: SReal & s < x} t; \
-\                 r: SReal; 0 < r \
-\              |] ==> x <= t + r";
-by (forward_tac [isLubD1a] 1);
-by (rtac ccontr 1 THEN dtac not_hypreal_leE 1);
-by (dres_inst_tac [("x","t")] SReal_add 1 THEN assume_tac 1);
-by (dres_inst_tac [("y","t + r")] (isLubD1 RS setleD) 1);
-by (auto_tac (claset(),simpset() addsimps [lemma_hypreal_le_left_cancel]));
-by (dtac hypreal_less_le_trans 1 THEN assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_less_not_refl]) 1);
-qed "lemma_st_part_le1";
-
-Goalw [setle_def] "!!x::hypreal. [| S *<= x; x < y |] ==> S *<= y";
-by (auto_tac (claset() addSDs [bspec,hypreal_le_less_trans]
-    addIs [hypreal_less_imp_le],simpset()));
-qed "hypreal_setle_less_trans";
-
-Goalw [isUb_def] 
-     "!!x::hypreal. [| isUb R S x; x < y; y: R |] ==> isUb R S y";
-by (blast_tac (claset() addIs [hypreal_setle_less_trans]) 1);
-qed "hypreal_gt_isUb";
-
-Goal "[| x: HFinite; x < y; y: SReal |] \
-\              ==> isUb SReal {s. s: SReal & s < x} y";
-by (auto_tac (claset() addDs [hypreal_less_trans]
-    addIs [hypreal_less_imp_le] addSIs [isUbI,setleI],simpset()));
-qed "lemma_st_part_gt_ub";
-
-Goal "t <= t + -r ==> r <= (0::hypreal)";
-by (dres_inst_tac [("x","-t")] hypreal_add_left_le_mono1 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_add_assoc RS sym]));
-qed "lemma_minus_le_zero";
-
-Goal "[| x: HFinite; \
-\        isLub SReal {s. s: SReal & s < x} t; \
-\        r: SReal; 0 < r |] \
-\     ==> t + -r <= x";
-by (forward_tac [isLubD1a] 1);
-by (rtac ccontr 1 THEN dtac not_hypreal_leE 1);
-by (dtac SReal_minus 1 THEN dres_inst_tac [("x","t")] 
-    SReal_add 1 THEN assume_tac 1);
-by (dtac lemma_st_part_gt_ub 1 THEN REPEAT(assume_tac 1));
-by (dtac isLub_le_isUb 1 THEN assume_tac 1);
-by (dtac lemma_minus_le_zero 1);
-by (auto_tac (claset() addDs [hypreal_less_le_trans],
-    simpset() addsimps [hypreal_less_not_refl]));
-qed "lemma_st_part_le2";
-
-Goal "((x::hypreal) <= t + r) = (x + -t <= r)";
-by (Step_tac 1);
-by (dres_inst_tac [("x","-t")] hypreal_add_left_le_mono1 1);
-by (dres_inst_tac [("x","t")] hypreal_add_le_mono1 2);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc]) 2);
-by (auto_tac (claset(),simpset() addsimps [hypreal_add_commute]));
-qed "lemma_hypreal_le_swap";
-
-Goal "[| x: HFinite; \
-\        isLub SReal {s. s: SReal & s < x} t; \
-\        r: SReal; 0 < r |] \
-\     ==> x + -t <= r";
-by (blast_tac (claset() addSIs [lemma_hypreal_le_swap RS iffD1,
-                lemma_st_part_le1]) 1);
-qed "lemma_st_part1a";
-
-Goal "(t + -r <= x) = (-(x + -t) <= (r::hypreal))";
-by (auto_tac (claset(),simpset() addsimps [hypreal_minus_add_distrib,
-    hypreal_add_commute,lemma_hypreal_le_swap  RS sym]));
-qed "lemma_hypreal_le_swap2";
-
-Goal "[| x: HFinite; \
-\        isLub SReal {s. s: SReal & s < x} t; \
-\        r: SReal; 0 < r |] \
-\     ==> -(x + -t) <= r";
-by (blast_tac (claset() addSIs [lemma_hypreal_le_swap2 RS iffD1,
-                lemma_st_part_le2]) 1);
-qed "lemma_st_part2a";
-
-Goal "x: SReal ==> isUb SReal {s. s: SReal & s < x} x";
-by (auto_tac (claset() addIs [isUbI, setleI,hypreal_less_imp_le],simpset()));
-qed "lemma_SReal_ub";
-
-Goal "x: SReal ==> isLub SReal {s. s: SReal & s < x} x";
-by (auto_tac (claset() addSIs [isLubI2,lemma_SReal_ub,setgeI],simpset()));
-by (forward_tac [isUbD2a] 1);
-by (res_inst_tac [("x","x"),("y","y")] hypreal_linear_less2 1);
-by (auto_tac (claset() addSIs [hypreal_less_imp_le,hypreal_le_refl],simpset()));
-by (EVERY1[dtac SReal_dense, assume_tac, assume_tac, Step_tac]);
-by (dres_inst_tac [("y","r")] isUbD 1);
-by (auto_tac (claset() addDs [hypreal_less_le_trans],
-    simpset() addsimps [hypreal_less_not_refl]));
-qed "lemma_SReal_lub";
-
-Goal "[| x: HFinite; \
-\        isLub SReal {s. s: SReal & s < x} t; \
-\        r: SReal; 0 < r |] \
-\     ==> x + -t ~= r";
-by Auto_tac;
-by (forward_tac [isLubD1a RS SReal_minus] 1);
-by (dtac SReal_add_cancel 1 THEN assume_tac 1);
-by (dres_inst_tac [("x","x")] lemma_SReal_lub 1);
-by (dtac hypreal_isLub_unique 1 THEN assume_tac 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_less_not_refl]));
-qed "lemma_st_part_not_eq1";
-
-Goal "[| x: HFinite; \
-\        isLub SReal {s. s: SReal & s < x} t; \
-\        r: SReal; 0 < r |] \
-\     ==> -(x + -t) ~= r";
-by (auto_tac (claset(),simpset() addsimps [hypreal_minus_add_distrib]));
-by (forward_tac [isLubD1a] 1);
-by (dtac SReal_add_cancel 1 THEN assume_tac 1);
-by (dres_inst_tac [("x","-x")] SReal_minus 1);
-by (Asm_full_simp_tac 1);
-by (dres_inst_tac [("x","x")] lemma_SReal_lub 1);
-by (dtac hypreal_isLub_unique 1 THEN assume_tac 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_less_not_refl]));
-qed "lemma_st_part_not_eq2";
-
-Goal "[| x: HFinite; \
-\        isLub SReal {s. s: SReal & s < x} t; \
-\        r: SReal; 0 < r |] \
-\     ==> abs (x + -t) < r";
-by (forward_tac [lemma_st_part1a] 1);
-by (forward_tac [lemma_st_part2a] 4);
-by Auto_tac;
-by (REPEAT(dtac hypreal_le_imp_less_or_eq 1));
-by (auto_tac (claset() addDs [lemma_st_part_not_eq1,
-    lemma_st_part_not_eq2],simpset() addsimps [hrabs_interval_iff2]));
-qed "lemma_st_part_major";
-
-Goal "[| x: HFinite; \
-\        isLub SReal {s. s: SReal & s < x} t |] \
-\     ==> ALL r: SReal. 0 < r --> abs (x + -t) < r";
-by (blast_tac (claset() addSDs [lemma_st_part_major]) 1);
-qed "lemma_st_part_major2";
-
-(*----------------------------------------------
-  Existence of real and Standard Part Theorem
- ----------------------------------------------*)
-Goal "x: HFinite ==> \
-\     EX t: SReal. ALL r: SReal. 0 < r --> abs (x + -t) < r";
-by (forward_tac [lemma_st_part_lub] 1 THEN Step_tac 1);
-by (forward_tac [isLubD1a] 1);
-by (blast_tac (claset() addDs [lemma_st_part_major2]) 1);
-qed "lemma_st_part_Ex";
-
-Goalw [inf_close_def,Infinitesimal_def] 
-          "x:HFinite ==> EX t: SReal. x @= t";
-by (blast_tac (claset() addSDs [lemma_st_part_Ex]) 1);
-qed "st_part_Ex";
-
-(*--------------------------------
-  Unique real infinitely close
- -------------------------------*)
-Goal "x:HFinite ==> EX! t. t: SReal & x @= t";
-by (dtac st_part_Ex 1 THEN Step_tac 1);
-by (dtac inf_close_sym 2 THEN dtac inf_close_sym 2 
-    THEN dtac inf_close_sym 2);
-by (auto_tac (claset() addSIs [inf_close_unique_real],simpset()));
-qed "st_part_Ex1";
-
-(*------------------------------------------------------------------
-       Finite and Infinite --- more theorems
- ------------------------------------------------------------------*)
-
-Goalw [HFinite_def,HInfinite_def] "HFinite Int HInfinite = {}";
-by (auto_tac (claset() addIs [hypreal_less_irrefl] 
-              addDs [hypreal_less_trans,bspec],
-              simpset()));
-qed "HFinite_Int_HInfinite_empty";
-Addsimps [HFinite_Int_HInfinite_empty];
-
-Goal "x: HFinite ==> x ~: HInfinite";
-by (EVERY1[Step_tac, dtac IntI, assume_tac]);
-by Auto_tac;
-qed "HFinite_not_HInfinite";
-
-val [prem] = goalw thy [HInfinite_def] "x~: HFinite ==> x: HInfinite";
-by (cut_facts_tac [prem] 1);
-by (Clarify_tac 1);
-by (auto_tac (claset() addSDs [spec],simpset() addsimps [HFinite_def,Bex_def]));
-by (dtac (hypreal_leI RS hypreal_le_imp_less_or_eq) 1);
-by (auto_tac (claset() addSDs [(SReal_subset_HFinite RS subsetD)],
-    simpset() addsimps [prem,hrabs_idempotent,prem RS not_HFinite_hrabs]));
-qed "not_HFinite_HInfinite";
-
-Goal "x : HInfinite | x : HFinite";
-by (blast_tac (claset() addIs [not_HFinite_HInfinite]) 1);
-qed "HInfinite_HFinite_disj";
-
-Goal "(x : HInfinite) = (x ~: HFinite)";
-by (blast_tac (claset() addDs [HFinite_not_HInfinite,
-               not_HFinite_HInfinite]) 1);
-qed "HInfinite_HFinite_iff";
-
-Goal "(x : HFinite) = (x ~: HInfinite)";
-by (simp_tac (simpset() addsimps [HInfinite_HFinite_iff]) 1);
-qed "HFinite_HInfinite_iff";
-
-(*------------------------------------------------------------------
-       Finite, Infinite and Infinitesimal --- more theorems
- ------------------------------------------------------------------*)
-
-Goal "x ~: Infinitesimal ==> x : HInfinite | x : HFinite - Infinitesimal";
-by (fast_tac (claset() addIs [not_HFinite_HInfinite]) 1);
-qed "HInfinite_diff_HFinite_Infinitesimal_disj";
-
-Goal "[| x : HFinite; x ~: Infinitesimal |] ==> inverse x : HFinite";
-by (cut_inst_tac [("x","inverse x")] HInfinite_HFinite_disj 1);
-by (auto_tac (claset() addSDs [HInfinite_inverse_Infinitesimal],simpset()));
-qed "HFinite_inverse";
-
-Goal "x : HFinite - Infinitesimal ==> inverse x : HFinite";
-by (blast_tac (claset() addIs [HFinite_inverse]) 1);
-qed "HFinite_inverse2";
-
-(* stronger statement possible in fact *)
-Goal "x ~: Infinitesimal ==> inverse(x) : HFinite";
-by (dtac HInfinite_diff_HFinite_Infinitesimal_disj 1);
-by (blast_tac (claset() addIs [HFinite_inverse,
-                 HInfinite_inverse_Infinitesimal,
-                 Infinitesimal_subset_HFinite RS subsetD]) 1);
-qed "Infinitesimal_inverse_HFinite";
-
-Goal "x : HFinite - Infinitesimal ==> inverse x : HFinite - Infinitesimal";
-by (auto_tac (claset() addIs [Infinitesimal_inverse_HFinite],simpset()));
-by (dtac Infinitesimal_HFinite_mult2 1);
-by (auto_tac (claset() addSDs [not_Infinitesimal_not_zero, 
-                               hypreal_mult_inverse],
-              simpset()));
-qed "HFinite_not_Infinitesimal_inverse";
-
-Goal "[| x @= y; y :  HFinite - Infinitesimal |] \
-\     ==> inverse x @= inverse y";
-by (forward_tac [HFinite_diff_Infinitesimal_inf_close] 1);
-by (assume_tac 1);
-by (forward_tac [not_Infinitesimal_not_zero2] 1);
-by (forw_inst_tac [("x","x")] not_Infinitesimal_not_zero2 1);
-by (REPEAT(dtac HFinite_inverse2 1));
-by (dtac inf_close_mult2 1 THEN assume_tac 1);
-by Auto_tac;
-by (dres_inst_tac [("c","inverse x")] inf_close_mult1 1 
-    THEN assume_tac 1);
-by (auto_tac (claset() addIs [inf_close_sym],
-    simpset() addsimps [hypreal_mult_assoc]));
-qed "inf_close_inverse";
-
-(*Used for NSLIM_inverse, NSLIMSEQ_inverse*)
-bind_thm ("hypreal_of_real_inf_close_inverse",
-       hypreal_of_real_HFinite_diff_Infinitesimal RSN (2, inf_close_inverse));
-
-Goal "[| x: HFinite - Infinitesimal; \
-\        h : Infinitesimal |] ==> inverse(x + h) @= inverse x";
-by (auto_tac (claset() addIs [inf_close_inverse, inf_close_sym, 
-                              Infinitesimal_add_inf_close_self], 
-              simpset()));
-qed "inverse_add_Infinitesimal_inf_close";
-
-Goal "[| x: HFinite - Infinitesimal; \
-\                     h : Infinitesimal |] ==> inverse(h + x) @= inverse x";
-by (rtac (hypreal_add_commute RS subst) 1);
-by (blast_tac (claset() addIs [inverse_add_Infinitesimal_inf_close]) 1);
-qed "inverse_add_Infinitesimal_inf_close2";
-
-Goal 
-     "[| x: HFinite - Infinitesimal; \
-\             h : Infinitesimal |] ==> inverse(x + h) + -inverse x @= h"; 
-by (rtac inf_close_trans2 1);
-by (auto_tac (claset() addIs [inverse_add_Infinitesimal_inf_close,
-    inf_close_minus_iff RS iffD1],simpset() addsimps [mem_infmal_iff RS sym]));
-qed "inverse_add_Infinitesimal_inf_close_Infinitesimal";
-
-Goal "(x : Infinitesimal) = (x*x : Infinitesimal)";
-by (auto_tac (claset() addIs [Infinitesimal_mult],simpset()));
-by (rtac ccontr 1 THEN forward_tac [Infinitesimal_inverse_HFinite] 1);
-by (forward_tac [not_Infinitesimal_not_zero] 1);
-by (auto_tac (claset() addDs [Infinitesimal_HFinite_mult],
-    simpset() addsimps [hypreal_mult_assoc]));
-qed "Infinitesimal_square_iff";
-Addsimps [Infinitesimal_square_iff RS sym];
-
-Goal "(x*x : HFinite) = (x : HFinite)";
-by (auto_tac (claset() addIs [HFinite_mult],simpset()));
-by (auto_tac (claset() addDs [HInfinite_mult],
-    simpset() addsimps [HFinite_HInfinite_iff]));
-qed "HFinite_square_iff";
-Addsimps [HFinite_square_iff];
-
-Goal "(x*x : HInfinite) = (x : HInfinite)";
-by (auto_tac (claset(),simpset() addsimps 
-    [HInfinite_HFinite_iff]));
-qed "HInfinite_square_iff";
-Addsimps [HInfinite_square_iff];
-
-Goal "[| a: HFinite-Infinitesimal; a* w @= a*z |] ==> w @= z";
-by (Step_tac 1);
-by (ftac HFinite_inverse 1 THEN assume_tac 1);
-by (dtac not_Infinitesimal_not_zero 1);
-by (auto_tac (claset() addDs [inf_close_mult2],
-    simpset() addsimps [hypreal_mult_assoc RS sym]));
-qed "inf_close_HFinite_mult_cancel";
-
-Goal "a: HFinite-Infinitesimal ==> (a * w @= a * z) = (w @= z)";
-by (auto_tac (claset() addIs [inf_close_mult2,
-    inf_close_HFinite_mult_cancel],simpset()));
-qed "inf_close_HFinite_mult_cancel_iff1";
-
-(*------------------------------------------------------------------
-                  more about absolute value (hrabs)
- ------------------------------------------------------------------*)
-
-Goal "abs x @= x | abs x @= -x";
-by (cut_inst_tac [("x","x")] hrabs_disj 1);
-by (auto_tac (claset(),simpset() addsimps [inf_close_refl]));
-qed "inf_close_hrabs_disj";
-
-(*------------------------------------------------------------------
-                  Theorems about monads
- ------------------------------------------------------------------*)
-
-Goal "monad (abs x) <= monad(x) Un monad(-x)";
-by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1);
-by Auto_tac;
-qed "monad_hrabs_Un_subset";
-
-Goal "e : Infinitesimal ==> monad (x+e) = monad x";
-by (fast_tac (claset() addSIs [Infinitesimal_add_inf_close_self RS inf_close_sym,
-    inf_close_monad_iff RS iffD1]) 1);
-qed "Infinitesimal_monad_eq";
-
-Goalw [monad_def] "(u:monad x) = (-u:monad (-x))";
-by Auto_tac;
-qed "mem_monad_iff";
-
-Goalw [monad_def] "(x:Infinitesimal) = (x:monad 0)";
-by (auto_tac (claset() addIs [inf_close_sym],
-    simpset() addsimps [mem_infmal_iff]));
-qed "Infinitesimal_monad_zero_iff";
-
-Goal "(x:monad 0) = (-x:monad 0)";
-by (simp_tac (simpset() addsimps [Infinitesimal_monad_zero_iff RS sym,
-    Infinitesimal_minus_iff RS sym]) 1);
-qed "monad_zero_minus_iff";
-
-Goal "(x:monad 0) = (abs x:monad 0)";
-by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1);
-by (auto_tac (claset(),simpset() addsimps [monad_zero_minus_iff RS sym]));
-qed "monad_zero_hrabs_iff";
-
-Goalw [monad_def] "x:monad x";
-by (simp_tac (simpset() addsimps [inf_close_refl]) 1);
-qed "mem_monad_self";
-Addsimps [mem_monad_self];
-
-(*------------------------------------------------------------------
-         Proof that x @= y ==> abs x @= abs y
- ------------------------------------------------------------------*)
-Goal "x @= y ==> {x,y}<=monad x";
-by (Simp_tac 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [inf_close_monad_iff]) 1);
-qed "inf_close_subset_monad";
-
-Goal "x @= y ==> {x,y}<=monad y";
-by (dtac inf_close_sym 1);
-by (fast_tac (claset() addDs [inf_close_subset_monad]) 1);
-qed "inf_close_subset_monad2";
-
-Goalw [monad_def] "u:monad x ==> x @= u";
-by (Fast_tac 1);
-qed "mem_monad_inf_close";
-
-Goalw [monad_def] "x @= u ==> u:monad x";
-by (Fast_tac 1);
-qed "inf_close_mem_monad";
-
-Goalw [monad_def] "x @= u ==> x:monad u";
-by (blast_tac (claset() addSIs [inf_close_sym]) 1);
-qed "inf_close_mem_monad2";
-
-Goal "[| x @= y;x:monad 0 |] ==> y:monad 0";
-by (dtac mem_monad_inf_close 1);
-by (fast_tac (claset() addIs [inf_close_mem_monad,inf_close_trans]) 1);
-qed "inf_close_mem_monad_zero";
-
-Goal "[| x @= y; x: Infinitesimal |] ==> abs x @= abs y";
-by (fast_tac (claset() addIs [(Infinitesimal_monad_zero_iff RS iffD1), 
-    inf_close_mem_monad_zero,(monad_zero_hrabs_iff RS iffD1),
-    mem_monad_inf_close,inf_close_trans3]) 1);
-qed "Infinitesimal_inf_close_hrabs";
-
-Goal "[| 0 < x; x ~:Infinitesimal |] \
-\     ==> ALL e: Infinitesimal. e < x";
-by (Step_tac 1 THEN rtac ccontr 1);
-by (auto_tac (claset() addIs [Infinitesimal_zero RSN (2, Infinitesimal_interval)] 
-    addSDs [hypreal_leI RS hypreal_le_imp_less_or_eq],simpset()));
-qed "Ball_Infinitesimal_less";
-
-Goal "[| 0 < x; x ~: Infinitesimal|] \
-\     ==> ALL u: monad x. 0 < u";
-by (Step_tac 1);
-by (dtac (mem_monad_inf_close RS inf_close_sym) 1);
-by (etac (bex_Infinitesimal_iff2 RS iffD2 RS bexE) 1);
-by (eres_inst_tac [("x","-xa")] (Ball_Infinitesimal_less RS ballE) 1);
-by (dtac (hypreal_less_minus_iff RS iffD1) 2);
-by (auto_tac (claset(),simpset() addsimps [Infinitesimal_minus_iff RS sym]));
-qed "Ball_mem_monad_gt_zero";
-
-Goal "[| x < 0; x ~: Infinitesimal|] \
-\     ==> ALL u: monad x. u < 0";
-by (Step_tac 1);
-by (dtac (mem_monad_inf_close RS inf_close_sym) 1);
-by (etac (bex_Infinitesimal_iff RS iffD2 RS bexE) 1);
-by (dtac (hypreal_minus_zero_less_iff RS iffD2) 1);
-by (eres_inst_tac [("x","xa")] (Ball_Infinitesimal_less RS ballE) 1);
-by (dtac (hypreal_less_minus_iff2 RS iffD1) 2);
-by (auto_tac (claset(),simpset() addsimps [Infinitesimal_minus_iff RS sym] @ hypreal_add_ac));
-qed "Ball_mem_monad_less_zero";
-
-Goal "[|0 < x; x ~: Infinitesimal; x @= y|] ==> 0 < y";
-by (blast_tac (claset() addDs [Ball_mem_monad_gt_zero,
-    inf_close_subset_monad]) 1);
-qed "lemma_inf_close_gt_zero";
-
-Goal "[|x < 0; x ~: Infinitesimal; x @= y|] ==> y < 0";
-by (blast_tac (claset() addDs [Ball_mem_monad_less_zero,
-    inf_close_subset_monad]) 1);
-qed "lemma_inf_close_less_zero";
-
-Goal "[| x @= y; x < 0; x ~: Infinitesimal |] \
-\     ==> abs x @= abs y";
-by (forward_tac [lemma_inf_close_less_zero] 1);
-by (REPEAT(assume_tac 1));
-by (REPEAT(dtac hrabs_minus_eqI2 1));
-by Auto_tac;
-qed "inf_close_hrabs1";
-
-Goal "[| x @= y; 0 < x; x ~: Infinitesimal |] \
-\     ==> abs x @= abs y";
-by (forward_tac [lemma_inf_close_gt_zero] 1);
-by (REPEAT(assume_tac 1));
-by (REPEAT(dtac hrabs_eqI2 1));
-by Auto_tac;
-qed "inf_close_hrabs2";
-
-Goal "x @= y ==> abs x @= abs y";
-by (res_inst_tac [("Q","x:Infinitesimal")] (excluded_middle RS disjE) 1);
-by (res_inst_tac [("x1","x"),("y1","0")] (hypreal_linear RS disjE) 1);
-by (auto_tac (claset() addIs [inf_close_hrabs1,inf_close_hrabs2,
-    Infinitesimal_inf_close_hrabs],simpset()));
-qed "inf_close_hrabs";
-
-Goal "abs(x) @= 0 ==> x @= 0";
-by (cut_inst_tac [("x","x")] hrabs_disj 1);
-by (auto_tac (claset() addDs [inf_close_minus],simpset()));
-qed "inf_close_hrabs_zero_cancel";
-
-Goal "e: Infinitesimal ==> abs x @= abs(x+e)";
-by (fast_tac (claset() addIs [inf_close_hrabs,
-       Infinitesimal_add_inf_close_self]) 1);
-qed "inf_close_hrabs_add_Infinitesimal";
-
-Goal "e: Infinitesimal ==> abs x @= abs(x + -e)";
-by (fast_tac (claset() addIs [inf_close_hrabs,
-    Infinitesimal_add_minus_inf_close_self]) 1);
-qed "inf_close_hrabs_add_minus_Infinitesimal";
-
-Goal "[| e: Infinitesimal; e': Infinitesimal; \
-\        abs(x+e) = abs(y+e')|] ==> abs x @= abs y";
-by (dres_inst_tac [("x","x")] inf_close_hrabs_add_Infinitesimal 1);
-by (dres_inst_tac [("x","y")] inf_close_hrabs_add_Infinitesimal 1);
-by (auto_tac (claset() addIs [inf_close_trans2],simpset()));
-qed "hrabs_add_Infinitesimal_cancel";
-
-Goal "[| e: Infinitesimal; e': Infinitesimal; \
-\        abs(x + -e) = abs(y + -e')|] ==> abs x @= abs y";
-by (dres_inst_tac [("x","x")] inf_close_hrabs_add_minus_Infinitesimal 1);
-by (dres_inst_tac [("x","y")] inf_close_hrabs_add_minus_Infinitesimal 1);
-by (auto_tac (claset() addIs [inf_close_trans2],simpset()));
-qed "hrabs_add_minus_Infinitesimal_cancel";
-
-(* interesting slightly counterintuitive theorem: necessary 
-   for proving that an open interval is an NS open set 
-*)
-Goalw [Infinitesimal_def] 
-      "[| xa: Infinitesimal; hypreal_of_real x < hypreal_of_real y |] \
-\      ==> hypreal_of_real x + xa < hypreal_of_real y";
-by (dtac (hypreal_less_minus_iff RS iffD1) 1 THEN Step_tac 1);
-by (rtac (hypreal_less_minus_iff RS iffD2) 1);
-by (dres_inst_tac [("x","hypreal_of_real y + -hypreal_of_real x")] bspec 1);
-by (auto_tac (claset(),
-              simpset() addsimps [hypreal_add_commute, 
-                                  hrabs_interval_iff,
-                                  SReal_add, SReal_minus]));
-by (dres_inst_tac [("x1","xa")] (hypreal_less_minus_iff RS iffD1) 1);
-by (auto_tac (claset(),
-        simpset() addsimps [hypreal_minus_add_distrib] @ hypreal_add_ac));
-qed "Infinitesimal_add_hypreal_of_real_less";
-
-Goal "[| x: Infinitesimal; abs(hypreal_of_real r) < hypreal_of_real y |] \
-\     ==> abs (hypreal_of_real r + x) < hypreal_of_real y";
-by (dres_inst_tac [("x","hypreal_of_real r")] 
-    inf_close_hrabs_add_Infinitesimal 1);
-by (dtac (inf_close_sym RS (bex_Infinitesimal_iff2 RS iffD2)) 1);
-by (auto_tac (claset() addSIs [Infinitesimal_add_hypreal_of_real_less],
-        simpset() addsimps [hypreal_of_real_hrabs]));
-qed "Infinitesimal_add_hrabs_hypreal_of_real_less";
-
-Goal "[| x: Infinitesimal; abs(hypreal_of_real r) < hypreal_of_real y |] \
-\     ==> abs (x + hypreal_of_real r) < hypreal_of_real y";
-by (rtac (hypreal_add_commute RS subst) 1);
-by (etac Infinitesimal_add_hrabs_hypreal_of_real_less 1);
-by (assume_tac 1);
-qed "Infinitesimal_add_hrabs_hypreal_of_real_less2";
-
-Goalw [hypreal_le_def]
-      "[| xa: Infinitesimal; hypreal_of_real x + xa <= hypreal_of_real y |] \
-\      ==> hypreal_of_real x <= hypreal_of_real y";
-by (EVERY1[rtac notI, rtac contrapos_np, assume_tac]);
-by (res_inst_tac [("C","-xa")] hypreal_less_add_right_cancel 1);
-by (auto_tac (claset() addDs [Infinitesimal_minus_iff RS iffD1]
-    addIs [Infinitesimal_add_hypreal_of_real_less],
-    simpset() addsimps [hypreal_add_assoc]));
-qed "Infinitesimal_add_cancel_hypreal_of_real_le";
-
-Goal "[| xa: Infinitesimal; hypreal_of_real x + xa <= hypreal_of_real y |] \
-\                   ==> x <= y";
-by (blast_tac (claset() addSIs [hypreal_of_real_le_iff RS iffD2,
-               Infinitesimal_add_cancel_hypreal_of_real_le]) 1);
-qed "Infinitesimal_add_cancel_real_le";
-
-Goalw [hypreal_le_def]
-     "[| xa: Infinitesimal; ya: Infinitesimal; \
-\        hypreal_of_real x + xa <= hypreal_of_real y + ya |] \
-\     ==> hypreal_of_real x <= hypreal_of_real y";
-by (EVERY1[rtac notI, rtac contrapos_np, assume_tac]);
-by (res_inst_tac [("C","-xa")] hypreal_less_add_right_cancel 1);
-by (dtac (Infinitesimal_minus_iff RS iffD1) 1);
-by (dtac Infinitesimal_add 1 THEN assume_tac 1);
-by (auto_tac (claset() addIs [Infinitesimal_add_hypreal_of_real_less],
-              simpset() addsimps [hypreal_add_assoc]));
-qed "hypreal_of_real_le_add_Infininitesimal_cancel";
-
-Goal "[| xa: Infinitesimal; ya: Infinitesimal; \
-\        hypreal_of_real x + xa <= hypreal_of_real y + ya |] \
-\     ==> x <= y";
-by (blast_tac (claset() addSIs [hypreal_of_real_le_iff RS iffD2,
-                          hypreal_of_real_le_add_Infininitesimal_cancel]) 1);
-qed "hypreal_of_real_le_add_Infininitesimal_cancel2";
-
-Goal "[| hypreal_of_real x < e; e: Infinitesimal |] ==> hypreal_of_real x <= 0";
-by (rtac hypreal_leI 1 THEN Step_tac 1);
-by (dtac Infinitesimal_interval 1);
-by (dtac (SReal_hypreal_of_real RS SReal_Infinitesimal_zero) 4);
-by (auto_tac (claset(), 
-           simpset() addsimps [hypreal_of_real_zero, hypreal_less_not_refl]));
-qed "hypreal_of_real_less_Infinitesimal_le_zero";
-
-(*used once, in NSDERIV_inverse*)
-Goal "[| h: Infinitesimal; x ~= 0 |] ==> hypreal_of_real x + h ~= 0";
-by (res_inst_tac [("t","h")] (hypreal_minus_minus RS subst) 1);
-by (rtac (not_sym RS (hypreal_not_eq_minus_iff RS iffD1)) 1);
-by (dtac (Infinitesimal_minus_iff RS iffD1) 1);
-by Auto_tac;  
-qed "Infinitesimal_add_not_zero";
-
-Goal "x*x + y*y : Infinitesimal ==> x*x : Infinitesimal";
-by (blast_tac (claset() addIs [hypreal_self_le_add_pos,
-                 Infinitesimal_interval2, hypreal_le_square]) 1);
-qed "Infinitesimal_square_cancel";
-Addsimps [Infinitesimal_square_cancel];
-
-Goal "x*x + y*y : HFinite ==> x*x : HFinite";
-by (blast_tac (claset() addIs [hypreal_self_le_add_pos,
-            HFinite_bounded,hypreal_le_square, HFinite_zero]) 1);
-qed "HFinite_square_cancel";
-Addsimps [HFinite_square_cancel];
-
-Goal "x*x + y*y : Infinitesimal ==> y*y : Infinitesimal";
-by (rtac Infinitesimal_square_cancel 1);
-by (rtac (hypreal_add_commute RS subst) 1);
-by (Simp_tac 1);
-qed "Infinitesimal_square_cancel2";
-Addsimps [Infinitesimal_square_cancel2];
-
-Goal "x*x + y*y : HFinite ==> y*y : HFinite";
-by (rtac HFinite_square_cancel 1);
-by (rtac (hypreal_add_commute RS subst) 1);
-by (Simp_tac 1);
-qed "HFinite_square_cancel2";
-Addsimps [HFinite_square_cancel2];
-
-Goal "x*x + y*y + z*z : Infinitesimal ==> x*x : Infinitesimal";
-by (blast_tac (claset() addIs [hypreal_self_le_add_pos2,
-    Infinitesimal_interval2,hypreal_le_square]) 1);
-qed "Infinitesimal_sum_square_cancel";
-Addsimps [Infinitesimal_sum_square_cancel];
-
-Goal "x*x + y*y + z*z : HFinite ==> x*x : HFinite";
-by (blast_tac (claset() addIs [hypreal_self_le_add_pos2,
-    HFinite_bounded,hypreal_le_square,
-    HFinite_zero]) 1);
-qed "HFinite_sum_square_cancel";
-Addsimps [HFinite_sum_square_cancel];
-
-Goal "y*y + x*x + z*z : Infinitesimal ==> x*x : Infinitesimal";
-by (rtac Infinitesimal_sum_square_cancel 1);
-by (asm_full_simp_tac (simpset() addsimps hypreal_add_ac) 1);
-qed "Infinitesimal_sum_square_cancel2";
-Addsimps [Infinitesimal_sum_square_cancel2];
-
-Goal "y*y + x*x + z*z : HFinite ==> x*x : HFinite";
-by (rtac HFinite_sum_square_cancel 1);
-by (asm_full_simp_tac (simpset() addsimps hypreal_add_ac) 1);
-qed "HFinite_sum_square_cancel2";
-Addsimps [HFinite_sum_square_cancel2];
-
-Goal "z*z + y*y + x*x : Infinitesimal ==> x*x : Infinitesimal";
-by (rtac Infinitesimal_sum_square_cancel 1);
-by (asm_full_simp_tac (simpset() addsimps hypreal_add_ac) 1);
-qed "Infinitesimal_sum_square_cancel3";
-Addsimps [Infinitesimal_sum_square_cancel3];
-
-Goal "z*z + y*y + x*x : HFinite ==> x*x : HFinite";
-by (rtac HFinite_sum_square_cancel 1);
-by (asm_full_simp_tac (simpset() addsimps hypreal_add_ac) 1);
-qed "HFinite_sum_square_cancel3";
-Addsimps [HFinite_sum_square_cancel3];
-
-Goal "[| y: monad x; 0 < hypreal_of_real e |] \
-\     ==> abs (y + -x) < hypreal_of_real e";
-by (dtac (mem_monad_inf_close RS inf_close_sym) 1);
-by (dtac (bex_Infinitesimal_iff RS iffD2) 1);
-by (auto_tac (claset() addSDs [InfinitesimalD],simpset()));
-qed "monad_hrabs_less";
-
-Goal "x: monad (hypreal_of_real  a) ==> x: HFinite";
-by (dtac (mem_monad_inf_close RS inf_close_sym) 1);
-by (dtac (bex_Infinitesimal_iff2 RS iffD2) 1);
-by (step_tac (claset() addSDs 
-       [Infinitesimal_subset_HFinite RS subsetD]) 1);
-by (etac (SReal_hypreal_of_real RS (SReal_subset_HFinite 
-         RS subsetD) RS HFinite_add) 1);
-qed "mem_monad_SReal_HFinite";
-
-(*------------------------------------------------------------------
-              Theorems about standard part
- ------------------------------------------------------------------*)
-
-Goalw [st_def] "x: HFinite ==> st x @= x";
-by (forward_tac [st_part_Ex] 1 THEN Step_tac 1);
-by (rtac someI2 1);
-by (auto_tac (claset() addIs [inf_close_sym],simpset()));
-qed "st_inf_close_self";
-
-Goalw [st_def] "x: HFinite ==> st x: SReal";
-by (forward_tac [st_part_Ex] 1 THEN Step_tac 1);
-by (rtac someI2 1);
-by (auto_tac (claset() addIs [inf_close_sym],simpset()));
-qed "st_SReal";
-
-Goal "x: HFinite ==> st x: HFinite";
-by (etac (st_SReal RS (SReal_subset_HFinite RS subsetD)) 1);
-qed "st_HFinite";
-
-Goalw [st_def] "x: SReal ==> st x = x";
-by (rtac some_equality 1);
-by (fast_tac (claset() addIs [(SReal_subset_HFinite RS subsetD), 
-    inf_close_refl]) 1);
-by (blast_tac (claset() addDs [SReal_inf_close_iff RS iffD1]) 1);
-qed "st_SReal_eq";
-
-(* should be added to simpset *)
-Goal "st (hypreal_of_real x) = hypreal_of_real x";
-by (rtac (SReal_hypreal_of_real RS st_SReal_eq) 1);
-qed "st_hypreal_of_real";
-
-Goal "[| x: HFinite; y: HFinite; st x = st y |] ==> x @= y";
-by (auto_tac (claset() addSDs [st_inf_close_self] 
-              addSEs [inf_close_trans3],simpset()));
-qed "st_eq_inf_close";
-
-Goal "[| x: HFinite; y: HFinite; x @= y |] ==> st x = st y";
-by (EVERY1 [forward_tac [st_inf_close_self],
-    forw_inst_tac [("x","y")] st_inf_close_self,
-    dtac st_SReal,dtac st_SReal]);
-by (fast_tac (claset() addEs [inf_close_trans,
-    inf_close_trans2,SReal_inf_close_iff RS iffD1]) 1);
-qed "inf_close_st_eq";
-
-Goal "[| x: HFinite; y: HFinite|] \
-\                  ==> (x @= y) = (st x = st y)";
-by (blast_tac (claset() addIs [inf_close_st_eq,
-               st_eq_inf_close]) 1);
-qed "st_eq_inf_close_iff";
-
-Goal "[| x: SReal; e: Infinitesimal |] ==> st(x + e) = x";
-by (forward_tac [st_SReal_eq RS subst] 1);
-by (assume_tac 2);
-by (forward_tac [SReal_subset_HFinite RS subsetD] 1);
-by (forward_tac [Infinitesimal_subset_HFinite RS subsetD] 1);
-by (dtac st_SReal_eq 1);
-by (rtac inf_close_st_eq 1);
-by (auto_tac (claset() addIs  [HFinite_add],
-    simpset() addsimps [Infinitesimal_add_inf_close_self 
-    RS inf_close_sym]));
-qed "st_Infinitesimal_add_SReal";
-
-Goal "[| x: SReal; e: Infinitesimal |] \
-\     ==> st(e + x) = x";
-by (rtac (hypreal_add_commute RS subst) 1);
-by (blast_tac (claset() addSIs [st_Infinitesimal_add_SReal]) 1);
-qed "st_Infinitesimal_add_SReal2";
-
-Goal "x: HFinite ==> \
-\     EX e: Infinitesimal. x = st(x) + e";
-by (blast_tac (claset() addSDs [(st_inf_close_self RS 
-    inf_close_sym),bex_Infinitesimal_iff2 RS iffD2]) 1);
-qed "HFinite_st_Infinitesimal_add";
-
-Goal "[| x: HFinite; y: HFinite |] \
-\     ==> st (x + y) = st(x) + st(y)";
-by (forward_tac [HFinite_st_Infinitesimal_add] 1);
-by (forw_inst_tac [("x","y")] HFinite_st_Infinitesimal_add 1);
-by (Step_tac 1);
-by (subgoal_tac "st (x + y) = st ((st x + e) + (st y + ea))" 1);
-by (dtac sym 2 THEN dtac sym 2);
-by (Asm_full_simp_tac 2);
-by (asm_simp_tac (simpset() addsimps hypreal_add_ac) 1);
-by (REPEAT(dtac st_SReal 1));
-by (dtac SReal_add 1 THEN assume_tac 1);
-by (dtac Infinitesimal_add 1 THEN assume_tac 1);
-by (rtac (hypreal_add_assoc RS subst) 1);
-by (blast_tac (claset() addSIs [st_Infinitesimal_add_SReal2]) 1);
-qed "st_add";
-
-Goal "st 0 = 0";
-by (rtac (SReal_zero RS st_SReal_eq) 1);
-qed "st_zero";
-
-Goal "st 1hr = 1hr";
-by (rtac (SReal_one RS st_SReal_eq) 1);
-qed "st_one";
-
-Addsimps [st_zero,st_one];
-
-Goal "y: HFinite ==> st(-y) = -st(y)";
-by (forward_tac [HFinite_minus_iff RS iffD1] 1);
-by (rtac hypreal_add_minus_eq_minus 1);
-by (dtac (st_add RS sym) 1 THEN assume_tac 1);
-by (auto_tac (claset(),simpset() addsimps hypreal_add_ac));
-qed "st_minus";
-
-Goal "[| x: HFinite; y: HFinite |] \
-\              ==> st (x + -y) = st(x) + -st(y)";
-by (forw_inst_tac [("y1","y")] (st_minus RS sym) 1);
-by (dres_inst_tac [("x1","y")] (HFinite_minus_iff RS iffD1) 1);
-by (asm_simp_tac (simpset() addsimps [st_add]) 1);
-qed "st_add_minus";
-
-(* lemma *)
-Goal "[| x: HFinite; y: HFinite; \
-\                 e: Infinitesimal; \
-\                 ea : Infinitesimal \
-\              |] ==> e*y + x*ea + e*ea: Infinitesimal";
-by (forw_inst_tac [("x","e"),("y","y")] Infinitesimal_HFinite_mult 1);
-by (forw_inst_tac [("x","ea"),("y","x")] Infinitesimal_HFinite_mult 2);
-by (dtac Infinitesimal_mult 3);
-by (auto_tac (claset() addIs [Infinitesimal_add],
-    simpset() addsimps hypreal_add_ac @ hypreal_mult_ac));
-qed "lemma_st_mult";
-
-Goal "[| x: HFinite; y: HFinite |] \
-\              ==> st (x * y) = st(x) * st(y)";
-by (forward_tac [HFinite_st_Infinitesimal_add] 1);
-by (forw_inst_tac [("x","y")] HFinite_st_Infinitesimal_add 1);
-by (Step_tac 1);
-by (subgoal_tac "st (x * y) = st ((st x + e) * (st y + ea))" 1);
-by (dtac sym 2 THEN dtac sym 2);
-by (Asm_full_simp_tac 2);
-by (thin_tac "x = st x + e" 1);
-by (thin_tac "y = st y + ea" 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypreal_add_mult_distrib,hypreal_add_mult_distrib2]) 1);
-by (REPEAT(dtac st_SReal 1));
-by (full_simp_tac (simpset() addsimps [hypreal_add_assoc]) 1);
-by (rtac st_Infinitesimal_add_SReal 1);
-by (blast_tac (claset() addSIs [SReal_mult]) 1);
-by (REPEAT(dtac (SReal_subset_HFinite RS subsetD) 1));
-by (rtac (hypreal_add_assoc RS subst) 1);
-by (blast_tac (claset() addSIs [lemma_st_mult]) 1);
-qed "st_mult";
-
-Goal "st(x) ~= 0 ==> x ~=0";
-by (fast_tac (claset() addIs [st_zero]) 1);
-qed "st_not_zero";
-
-Goal "x: Infinitesimal ==> st x = 0";
-by (rtac (st_zero RS subst) 1);
-by (rtac inf_close_st_eq 1);
-by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite 
-    RS subsetD],simpset() addsimps [mem_infmal_iff RS sym]));
-qed "st_Infinitesimal";
-
-Goal "st(x) ~= 0 ==> x ~: Infinitesimal";
-by (fast_tac (claset() addIs [st_Infinitesimal]) 1);
-qed "st_not_Infinitesimal";
-
-val [prem1,prem2] = 
-goal thy "[| x: HFinite; st x ~= 0 |] \
-\         ==> st(inverse x) = inverse (st x)";
-by (rtac ((prem2 RS hypreal_mult_left_cancel) RS iffD1) 1);
-by (auto_tac (claset(),simpset() addsimps [st_not_zero,
-    st_mult RS sym,prem2,st_not_Infinitesimal,HFinite_inverse,prem1]));
-qed "st_inverse";
-
-val [prem1,prem2,prem3] = 
-goal thy "[| x: HFinite; y: HFinite; st y ~= 0 |] \
-\                 ==> st(x * inverse y) = (st x) * inverse (st y)";
-by (auto_tac (claset(),simpset() addsimps [st_not_zero,prem3,
-    st_mult,prem2,st_not_Infinitesimal,HFinite_inverse,prem1,st_inverse]));
-qed "st_mult_inverse";
-
-Goal "x: HFinite ==> st(st(x)) = st(x)";
-by (blast_tac (claset() addIs [st_HFinite,
-    st_inf_close_self,inf_close_st_eq]) 1);
-qed "st_idempotent";
-
-(*** lemmas ***)
-Goal "[| x: HFinite; y: HFinite; \
-\        xa: Infinitesimal; st x < st y \
-\     |] ==> st x + xa < st y";
-by (REPEAT(dtac st_SReal 1));
-by (auto_tac (claset() addSIs 
-    [Infinitesimal_add_hypreal_of_real_less],
-    simpset() addsimps [SReal_iff]));
-qed "Infinitesimal_add_st_less";
-
-Goalw [hypreal_le_def]
-     "[| x: HFinite; y: HFinite; \
-\        xa: Infinitesimal; st x <= st y + xa\
-\     |] ==> st x <= st y";
-by (auto_tac (claset() addDs [Infinitesimal_add_st_less],
-    simpset()));
-qed "Infinitesimal_add_st_le_cancel";
-
-Goal "[| x: HFinite; y: HFinite; x <= y |] \
-\     ==> st(x) <= st(y)";
-by (forward_tac [HFinite_st_Infinitesimal_add] 1);
-by (rotate_tac 1 1);
-by (forward_tac [HFinite_st_Infinitesimal_add] 1);
-by (Step_tac 1);
-by (rtac Infinitesimal_add_st_le_cancel 1);
-by (res_inst_tac [("x","ea"),("y","e")] 
-             Infinitesimal_add_minus 3);
-by (auto_tac (claset(),simpset() addsimps 
-    [hypreal_add_assoc RS sym]));
-by (res_inst_tac [("C","e")] hypreal_le_add_right_cancel 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [hypreal_add_assoc]) 1);
-qed "st_le";
-
-Goal "[| x: HFinite; 0 <= x |] ==> 0 <= st x";
-by (rtac (st_zero RS subst) 1);
-by (auto_tac (claset() addIs [st_le],simpset() 
-    delsimps [st_zero]));
-qed "st_zero_le";
-
-Goal "[| x: HFinite; x <= 0 |] ==> st x <= 0";
-by (rtac (st_zero RS subst) 1);
-by (auto_tac (claset() addIs [st_le],simpset() 
-    delsimps [st_zero]));
-qed "st_zero_ge";
-
-Goal "x: HFinite ==> abs(st x) = st(abs x)";
-by (case_tac "0 <= x" 1);
-by (auto_tac (claset() addSDs [not_hypreal_leE,
-    hypreal_less_imp_le],simpset() addsimps 
-    [st_zero_le,hrabs_eqI1, hrabs_minus_eqI1,
-     st_zero_ge,st_minus]));
-qed "st_hrabs";
-
-(*--------------------------------------------------------------------
-   Alternative definitions for HFinite using Free ultrafilter
- --------------------------------------------------------------------*)
-
-Goal "[| X: Rep_hypreal x; Y: Rep_hypreal x |] \
-\              ==> {n. X n = Y n} : FreeUltrafilterNat";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by Auto_tac;
-by (Ultra_tac 1);
-qed "FreeUltrafilterNat_Rep_hypreal";
-
-Goal "{n. Yb n < Y n} Int {n. -y = Yb n} <= {n. -y < Y n}";
-by Auto_tac;
-qed "lemma_free1";
-
-Goal "{n. Xa n < Yc n} Int {n. y = Yc n} <= {n. Xa n < y}";
-by Auto_tac;
-qed "lemma_free2";
-
-Goalw [HFinite_def] 
-    "x : HFinite ==> EX X: Rep_hypreal x. \
-\    EX u. {n. abs (X n) < u}:  FreeUltrafilterNat";
-by (auto_tac (claset(),simpset() addsimps 
-    [hrabs_interval_iff]));
-by (auto_tac (claset(),simpset() addsimps 
-    [hypreal_less_def,SReal_iff,hypreal_minus,
-     hypreal_of_real_def]));
-by (dtac FreeUltrafilterNat_Rep_hypreal 1 THEN assume_tac 1);
-by (res_inst_tac [("x","Y")] bexI 1 THEN assume_tac 2);
-by (res_inst_tac [("x","y")] exI 1);
-by (Ultra_tac 1 THEN arith_tac 1);
-qed "HFinite_FreeUltrafilterNat";
-
-Goalw [HFinite_def] 
-     "EX X: Rep_hypreal x. \
-\      EX u. {n. abs (X n) < u}: FreeUltrafilterNat\
-\      ==>  x : HFinite";
-by (auto_tac (claset(),simpset() addsimps 
-    [hrabs_interval_iff]));
-by (res_inst_tac [("x","hypreal_of_real u")] bexI 1);
-by (auto_tac (claset() addSIs [exI],simpset() addsimps 
-    [hypreal_less_def,SReal_iff,hypreal_minus,
-     hypreal_of_real_def]));
-by (ALLGOALS(Ultra_tac THEN' arith_tac));
-qed "FreeUltrafilterNat_HFinite";
-
-Goal "(x : HFinite) = (EX X: Rep_hypreal x. \
-\          EX u. {n. abs (X n) < u}:  FreeUltrafilterNat)";
-by (blast_tac (claset() addSIs [HFinite_FreeUltrafilterNat,
-               FreeUltrafilterNat_HFinite]) 1);
-qed "HFinite_FreeUltrafilterNat_iff";
-
-(*--------------------------------------------------------------------
-   Alternative definitions for HInfinite using Free ultrafilter
- --------------------------------------------------------------------*)
-Goal "- {n. (u::real) < abs (xa n)} = {n. abs (xa n) <= u}";
-by Auto_tac;
-qed "lemma_Compl_eq";
-
-Goal "- {n. abs (xa n) < (u::real)} = {n. u <= abs (xa n)}";
-by Auto_tac;
-qed "lemma_Compl_eq2";
-
-Goal "{n. abs (xa n) <= (u::real)} Int {n. u <= abs (xa n)} \
-\         = {n. abs(xa n) = u}";
-by Auto_tac;
-qed "lemma_Int_eq1";
-
-Goal "{n. abs (xa n) = u} <= {n. abs (xa n) < u + (#1::real)}";
-by Auto_tac;
-qed "lemma_FreeUltrafilterNat_one";
-
-(*-------------------------------------
-  Exclude this type of sets from free 
-  ultrafilter for Infinite numbers!
- -------------------------------------*)
-Goal "[| xa: Rep_hypreal x; \
-\                 {n. abs (xa n) = u} : FreeUltrafilterNat \
-\              |] ==> x: HFinite";
-by (rtac FreeUltrafilterNat_HFinite 1);
-by (res_inst_tac [("x","xa")] bexI 1);
-by (res_inst_tac [("x","u + #1")] exI 1);
-by (Ultra_tac 1 THEN assume_tac 1);
-qed "FreeUltrafilterNat_const_Finite";
-
-val [prem] = goal thy "x : HInfinite ==> EX X: Rep_hypreal x. \
-\          ALL u. {n. u < abs (X n)}:  FreeUltrafilterNat";
-by (cut_facts_tac [(prem RS (HInfinite_HFinite_iff RS iffD1))] 1);
-by (cut_inst_tac [("x","x")] Rep_hypreal_nonempty 1);
-by (auto_tac (claset(),simpset() delsimps [Rep_hypreal_nonempty] 
-    addsimps [HFinite_FreeUltrafilterNat_iff,Bex_def]));
-by (REPEAT(dtac spec 1));
-by Auto_tac;
-by (dres_inst_tac [("x","u")] spec 1 THEN 
-    REPEAT(dtac FreeUltrafilterNat_Compl_mem 1));
-by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
-
-
-by (asm_full_simp_tac (simpset() addsimps [lemma_Compl_eq,
-    lemma_Compl_eq2,lemma_Int_eq1]) 1);
-by (auto_tac (claset() addDs [FreeUltrafilterNat_const_Finite],
-    simpset() addsimps [(prem RS (HInfinite_HFinite_iff RS iffD1))]));
-qed "HInfinite_FreeUltrafilterNat";
-
-(* yet more lemmas! *)
-Goal "{n. abs (Xa n) < u} Int {n. X n = Xa n} \
-\         <= {n. abs (X n) < (u::real)}";
-by Auto_tac;
-qed "lemma_Int_HI";
-
-Goal "{n. u < abs (X n)} Int {n. abs (X n) < (u::real)} = {}";
-by (auto_tac (claset() addIs [real_less_asym],simpset()));
-qed "lemma_Int_HIa";
-
-Goal "EX X: Rep_hypreal x. ALL u. \
-\              {n. u < abs (X n)}: FreeUltrafilterNat\
-\              ==>  x : HInfinite";
-by (rtac (HInfinite_HFinite_iff RS iffD2) 1);
-by (Step_tac 1 THEN dtac HFinite_FreeUltrafilterNat 1);
-by Auto_tac;
-by (dres_inst_tac [("x","u")] spec 1);
-by (dtac FreeUltrafilterNat_Rep_hypreal 1 THEN assume_tac 1);
-by (dres_inst_tac [("Y","{n. X n = Xa n}")] FreeUltrafilterNat_Int 1);
-by (dtac (lemma_Int_HI RSN (2,FreeUltrafilterNat_subset)) 2);
-by (dres_inst_tac [("Y","{n. abs (X n) < u}")] FreeUltrafilterNat_Int 2);
-by (auto_tac (claset(),simpset() addsimps [lemma_Int_HIa,
-              FreeUltrafilterNat_empty]));
-qed "FreeUltrafilterNat_HInfinite";
-
-Goal "(x : HInfinite) = (EX X: Rep_hypreal x. \
-\          ALL u. {n. u < abs (X n)}:  FreeUltrafilterNat)";
-by (blast_tac (claset() addSIs [HInfinite_FreeUltrafilterNat,
-               FreeUltrafilterNat_HInfinite]) 1);
-qed "HInfinite_FreeUltrafilterNat_iff";
-
-(*--------------------------------------------------------------------
-   Alternative definitions for Infinitesimal using Free ultrafilter
- --------------------------------------------------------------------*)
-
-Goal "{n. - u < Yd n} Int {n. xa n = Yd n} <= {n. -u < xa n}";
-by Auto_tac;
-qed "lemma_free4";
-
-Goal "{n. Yb n < u} Int {n. xa n = Yb n} <= {n. xa n < u}";
-by Auto_tac;
-qed "lemma_free5";
-
-Goalw [Infinitesimal_def] 
-          "x : Infinitesimal ==> EX X: Rep_hypreal x. \
-\          ALL u. 0 < u --> {n. abs (X n) < u}:  FreeUltrafilterNat";
-by (auto_tac (claset(),simpset() addsimps [hrabs_interval_iff]));
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (EVERY[Auto_tac, rtac bexI 1, rtac lemma_hyprel_refl 2, Step_tac 1]);
-by (dtac (hypreal_of_real_less_iff RS iffD1) 1);
-by (dres_inst_tac [("x","hypreal_of_real u")] bspec 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_zero]));
-by (thin_tac "0 < hypreal_of_real  u" 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_less_def,       
-     hypreal_minus,hypreal_of_real_def,hypreal_of_real_zero]));
-by (Ultra_tac 1 THEN arith_tac 1);
-qed "Infinitesimal_FreeUltrafilterNat";
-
-Goalw [Infinitesimal_def] 
-          "EX X: Rep_hypreal x. \
-\               ALL u. 0 < u --> \
-\               {n. abs (X n) < u}:  FreeUltrafilterNat \
-\               ==> x : Infinitesimal";
-by (auto_tac (claset(),simpset() addsimps 
-    [hrabs_interval_iff,abs_interval_iff]));
-by (auto_tac (claset(),simpset() addsimps [SReal_iff,
-    hypreal_of_real_zero RS sym]));
-by (auto_tac (claset() addSIs [exI] 
-    addIs [FreeUltrafilterNat_subset],
-    simpset() addsimps [hypreal_less_def,
-    hypreal_minus,hypreal_of_real_def]));
-qed "FreeUltrafilterNat_Infinitesimal";
-
-Goal "(x : Infinitesimal) = (EX X: Rep_hypreal x. \
-\          ALL u. 0 < u --> {n. abs (X n) < u}:  FreeUltrafilterNat)";
-by (blast_tac (claset() addSIs [Infinitesimal_FreeUltrafilterNat,
-               FreeUltrafilterNat_Infinitesimal]) 1);
-qed "Infinitesimal_FreeUltrafilterNat_iff";
-
-(*------------------------------------------------------------------------
-         Infinitesimals as smaller than 1/n for all n::nat (> 0)
- ------------------------------------------------------------------------*)
-Goal "(ALL r. 0 < r --> x < r) = (ALL n. x < inverse(real_of_posnat n))";
-by (auto_tac (claset(),
-        simpset() addsimps [rename_numerals real_of_posnat_gt_zero]));
-by (blast_tac (claset() addSDs [rename_numerals reals_Archimedean] 
-                        addIs [order_less_trans]) 1);
-qed "lemma_Infinitesimal";
-
-Goal "(ALL r: SReal. 0 < r --> x < r) = \
-\     (ALL n. x < inverse(hypreal_of_posnat n))";
-by (Step_tac 1);
-by (dres_inst_tac [("x","inverse (hypreal_of_real(real_of_posnat n))")] 
-    bspec 1);
-by (full_simp_tac (simpset() addsimps [SReal_inverse]) 1); 
-by (rtac (real_of_posnat_gt_zero RS real_inverse_gt_zero RS 
-          (hypreal_of_real_less_iff RS iffD1) RSN(2,impE)) 1);
-by (assume_tac 2);
-by (asm_full_simp_tac (simpset() addsimps 
-       [rename_numerals real_of_posnat_gt_zero, hypreal_of_real_zero,
-        hypreal_of_real_of_posnat]) 1);
-by (auto_tac (claset() addSDs [reals_Archimedean],
-              simpset() addsimps [SReal_iff,hypreal_of_real_zero RS sym]));
-by (dtac (hypreal_of_real_less_iff RS iffD1) 1);
-by (asm_full_simp_tac (simpset() addsimps 
-       [rename_numerals real_of_posnat_gt_zero,hypreal_of_real_of_posnat])1);
-by (blast_tac (claset() addIs [hypreal_less_trans]) 1);
-qed "lemma_Infinitesimal2";
-
-Goalw [Infinitesimal_def] 
-     "Infinitesimal = {x. ALL n. abs x < inverse (hypreal_of_posnat n)}";
-by (auto_tac (claset(), simpset() addsimps [lemma_Infinitesimal2]));
-qed "Infinitesimal_hypreal_of_posnat_iff";
-
-(*-----------------------------------------------------------------------------
-       Actual proof that omega (whr) is an infinite number and 
-       hence that epsilon (ehr) is an infinitesimal number.
- -----------------------------------------------------------------------------*)
-Goal "{n. n < Suc m} = {n. n < m} Un {n. n = m}";
-by (auto_tac (claset(),simpset() addsimps [less_Suc_eq]));
-qed "Suc_Un_eq";
-
-(*-------------------------------------------
-  Prove that any segment is finite and 
-  hence cannot belong to FreeUltrafilterNat
- -------------------------------------------*)
-Goal "finite {n::nat. n < m}";
-by (nat_ind_tac "m" 1);
-by (auto_tac (claset(),simpset() addsimps [Suc_Un_eq]));
-qed "finite_nat_segment";
-
-Goal "finite {n::nat. real_of_posnat n < real_of_posnat m}";
-by (auto_tac (claset() addIs [finite_nat_segment],simpset()));
-qed "finite_real_of_posnat_segment";
-
-Goal "finite {n. real_of_posnat n < u}";
-by (cut_inst_tac [("x","u")] reals_Archimedean2 1);
-by (Step_tac 1);
-by (rtac (finite_real_of_posnat_segment RSN (2,finite_subset)) 1);
-by (auto_tac (claset() addDs [order_less_trans],simpset()));
-qed "finite_real_of_posnat_less_real";
-
-Goal "{n. real_of_posnat n <= u} = \
-\     {n. real_of_posnat n < u} Un {n. u = real_of_posnat n}";
-by (auto_tac (claset() addDs [real_le_imp_less_or_eq],
-    simpset() addsimps [real_le_refl,real_less_imp_le]));
-qed "lemma_real_le_Un_eq";
-
-Goal "finite {n. real_of_posnat n <= u}";
-by (auto_tac (claset(),simpset() addsimps 
-    [lemma_real_le_Un_eq,lemma_finite_omega_set,
-     finite_real_of_posnat_less_real]));
-qed "finite_real_of_posnat_le_real";
-
-Goal "finite {n. abs(real_of_posnat n) <= u}";
-by (full_simp_tac (simpset() addsimps [rename_numerals 
-    real_of_posnat_gt_zero,abs_eqI2,
-    finite_real_of_posnat_le_real]) 1);
-qed "finite_rabs_real_of_posnat_le_real";
-
-Goal "{n. abs(real_of_posnat n) <= u} ~: FreeUltrafilterNat";
-by (blast_tac (claset() addSIs [FreeUltrafilterNat_finite,
-    finite_rabs_real_of_posnat_le_real]) 1);
-qed "rabs_real_of_posnat_le_real_FreeUltrafilterNat";
-
-Goal "{n. u < real_of_posnat n} : FreeUltrafilterNat";
-by (rtac ccontr 1 THEN dtac FreeUltrafilterNat_Compl_mem 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [CLAIM_SIMP "- {n. u < real_of_posnat  n} = \
-\                {n. real_of_posnat n <= u}" 
-     [real_le_def],finite_real_of_posnat_le_real 
-                   RS FreeUltrafilterNat_finite]));
-qed "FreeUltrafilterNat_nat_gt_real";
-
-(*--------------------------------------------------------------
- The complement of {n. abs(real_of_posnat n) <= u} = 
- {n. u < abs (real_of_posnat n)} is in FreeUltrafilterNat 
- by property of (free) ultrafilters
- --------------------------------------------------------------*)
-Goal "- {n. abs (real_of_posnat  n) <= u} = \
-\     {n. u < abs (real_of_posnat  n)}";
-by (auto_tac (claset() addSDs [real_le_less_trans],
-   simpset() addsimps [not_real_leE,real_less_not_refl]));
-val lemma = result();
-
-Goal "{n. u < abs (real_of_posnat n)} : FreeUltrafilterNat";
-by (cut_inst_tac [("u","u")] rabs_real_of_posnat_le_real_FreeUltrafilterNat 1);
-by (auto_tac (claset() addDs [FreeUltrafilterNat_Compl_mem],
-    simpset() addsimps [lemma]));
-qed "FreeUltrafilterNat_omega";
-
-(*-----------------------------------------------
-       Omega is a member of HInfinite
- -----------------------------------------------*)
-Goalw [omega_def] "whr: HInfinite";
-by (auto_tac (claset() addSIs [FreeUltrafilterNat_HInfinite,
-    lemma_hyprel_refl,FreeUltrafilterNat_omega],simpset()));
-qed "HInfinite_omega";
-
-(*-----------------------------------------------
-       Epsilon is a member of Infinitesimal
- -----------------------------------------------*)
-
-Goal "ehr : Infinitesimal";
-by (auto_tac (claset() addSIs [HInfinite_inverse_Infinitesimal,HInfinite_omega],
-    simpset() addsimps [hypreal_epsilon_inverse_omega]));
-qed "Infinitesimal_epsilon";
-Addsimps [Infinitesimal_epsilon];
-
-Goal "ehr : HFinite";
-by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],
-    simpset()));
-qed "HFinite_epsilon";
-Addsimps [HFinite_epsilon];
-
-Goal "ehr @= 0";
-by (simp_tac (simpset() addsimps [mem_infmal_iff RS sym]) 1);
-qed "epsilon_inf_close_zero";
-Addsimps [epsilon_inf_close_zero];
-
-(*------------------------------------------------------------------------
-  Needed for proof that we define a hyperreal [<X(n)] @= hypreal_of_real a given 
-  that ALL n. |X n - a| < 1/n. Used in proof of NSLIM => LIM.
- -----------------------------------------------------------------------*)
-
-Goal "0 < u ==> finite {n. u < inverse(real_of_posnat n)}";
-by (asm_simp_tac (simpset() addsimps [real_of_posnat_less_inverse_iff]) 1);
-by (rtac finite_real_of_posnat_less_real 1);
-qed "finite_inverse_real_of_posnat_gt_real";
-
-Goal "{n. u <= inverse(real_of_posnat n)} = \
-\      {n. u < inverse(real_of_posnat n)} Un {n. u = inverse(real_of_posnat n)}";
-by (auto_tac (claset() addDs [real_le_imp_less_or_eq],
-    simpset() addsimps [real_le_refl,real_less_imp_le]));
-qed "lemma_real_le_Un_eq2";
-
-Goal "0 < u ==> finite {n::nat. u = inverse(real_of_posnat  n)}";
-by (asm_simp_tac (simpset() addsimps [real_of_posnat_inverse_eq_iff]) 1);
-by (auto_tac (claset() addIs [lemma_finite_omega_set RSN 
-    (2,finite_subset)],simpset()));
-qed "lemma_finite_omega_set2";
-
-Goal "0 < u ==> finite {n. u <= inverse(real_of_posnat n)}";
-by (auto_tac (claset(),simpset() addsimps 
-    [lemma_real_le_Un_eq2,lemma_finite_omega_set2,
-     finite_inverse_real_of_posnat_gt_real]));
-qed "finite_inverse_real_of_posnat_ge_real";
-
-Goal "0 < u ==> \
-\    {n. u <= inverse(real_of_posnat n)} ~: FreeUltrafilterNat";
-by (blast_tac (claset() addSIs [FreeUltrafilterNat_finite,
-    finite_inverse_real_of_posnat_ge_real]) 1);
-qed "inverse_real_of_posnat_ge_real_FreeUltrafilterNat";
-
-(*--------------------------------------------------------------
-    The complement of  {n. u <= inverse(real_of_posnat n)} =
-    {n. inverse(real_of_posnat n) < u} is in FreeUltrafilterNat 
-    by property of (free) ultrafilters
- --------------------------------------------------------------*)
-Goal "- {n. u <= inverse(real_of_posnat n)} = \
-\     {n. inverse(real_of_posnat n) < u}";
-by (auto_tac (claset() addSDs [real_le_less_trans],
-   simpset() addsimps [not_real_leE,real_less_not_refl]));
-val lemma = result();
-
-Goal "0 < u ==> \
-\     {n. inverse(real_of_posnat n) < u} : FreeUltrafilterNat";
-by (cut_inst_tac [("u","u")]  inverse_real_of_posnat_ge_real_FreeUltrafilterNat 1);
-by (auto_tac (claset() addDs [FreeUltrafilterNat_Compl_mem],
-    simpset() addsimps [lemma]));
-qed "FreeUltrafilterNat_inverse_real_of_posnat";
-
-(*--------------------------------------------------------------
-      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 (the theorem was previously in REALTOPOS.thy).
- -------------------------------------------------------------*)
-(*-----------------------------------------------------
-    |X(n) - x| < 1/n ==> [<X n>] - hypreal_of_real x|: Infinitesimal 
- -----------------------------------------------------*)
-Goal "ALL n. abs(X n + -x) < inverse(real_of_posnat n) \ 
-\    ==> Abs_hypreal(hyprel^^{X}) + -hypreal_of_real x : Infinitesimal";
-by (auto_tac (claset() addSIs [bexI] 
-           addDs [rename_numerals FreeUltrafilterNat_inverse_real_of_posnat,
-                  FreeUltrafilterNat_all,FreeUltrafilterNat_Int] 
-           addIs [order_less_trans, FreeUltrafilterNat_subset],
-      simpset() addsimps [hypreal_minus, 
-                          hypreal_of_real_def,hypreal_add,
-                          Infinitesimal_FreeUltrafilterNat_iff,hypreal_inverse,
-                          hypreal_of_real_of_posnat]));
-qed "real_seq_to_hypreal_Infinitesimal";
-
-Goal "ALL n. abs(X n + -x) < inverse(real_of_posnat n) \ 
-\     ==> Abs_hypreal(hyprel^^{X}) @= hypreal_of_real x";
-by (rtac (inf_close_minus_iff RS ssubst) 1);
-by (rtac (mem_infmal_iff RS subst) 1);
-by (etac real_seq_to_hypreal_Infinitesimal 1);
-qed "real_seq_to_hypreal_inf_close";
-
-Goal "ALL n. abs(x + -X n) < inverse(real_of_posnat n) \ 
-\              ==> Abs_hypreal(hyprel^^{X}) @= hypreal_of_real x";
-by (asm_full_simp_tac (simpset() addsimps [abs_minus_add_cancel,
-        real_seq_to_hypreal_inf_close]) 1);
-qed "real_seq_to_hypreal_inf_close2";
-
-Goal "ALL n. abs(X n + -Y n) < inverse(real_of_posnat n) \ 
-\     ==> Abs_hypreal(hyprel^^{X}) + \
-\         -Abs_hypreal(hyprel^^{Y}) : Infinitesimal";
-by (auto_tac (claset() addSIs [bexI] 
-                  addDs [rename_numerals 
-                         FreeUltrafilterNat_inverse_real_of_posnat,
-                         FreeUltrafilterNat_all,FreeUltrafilterNat_Int] 
-        addIs [order_less_trans, FreeUltrafilterNat_subset],
-     simpset() addsimps 
-            [Infinitesimal_FreeUltrafilterNat_iff,hypreal_minus,hypreal_add, 
-             hypreal_inverse,hypreal_of_real_of_posnat]));
-qed "real_seq_to_hypreal_Infinitesimal2";
--- a/src/HOL/Real/Hyperreal/NSA.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,49 +0,0 @@
-(*  Title       : NSA.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Infinite numbers, Infinitesimals,
-                  infinitely close relation etc.
-*) 
-
-NSA = HRealAbs + RComplete +
-
-constdefs
-
-   (* standard real numbers reagarded as *)
-   (* an embedded subset of hyperreals   *)
-   SReal  :: "hypreal set"
-   "SReal == {x. EX r. x = hypreal_of_real r}"
-
-   Infinitesimal  :: "hypreal set"
-   "Infinitesimal == {x. ALL r: SReal. 0 < r --> abs x < r}"
-
-   HFinite :: "hypreal set"
-   "HFinite == {x. EX r: SReal. abs x < r}"
-
-   HInfinite :: "hypreal set"
-   "HInfinite == {x. ALL r: SReal. r < abs x}"
-
-consts   
-
-    (* standard part map *)
-    st       :: hypreal => hypreal
-
-    (* infinitely close *)
-    "@="     :: [hypreal,hypreal] => bool  (infixl 50)  
-
-    monad    :: hypreal => hypreal set
-    galaxy   :: hypreal => hypreal set
-
-defs  
-
-   inf_close_def  "x @= y       == (x + -y) : Infinitesimal"     
-   st_def         "st           == (%x. @r. x : HFinite & r:SReal & r @= x)"
-
-   monad_def      "monad x      == {y. x @= y}"
-   galaxy_def     "galaxy x     == {y. (x + -y) : HFinite}"
- 
-end
-
-
-
-
--- a/src/HOL/Real/Hyperreal/NatStar.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,490 +0,0 @@
-(*  Title       : NatStar.ML
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : *-transforms in NSA which extends 
-                  sets of reals, and nat=>real, 
-                  nat=>nat functions
-*) 
-
-Goalw [starsetNat_def] 
-      "*sNat*(UNIV::nat set) = (UNIV::hypnat set)";
-by (auto_tac (claset(), simpset() addsimps [FreeUltrafilterNat_Nat_set]));
-qed "NatStar_real_set";
-
-Goalw [starsetNat_def] "*sNat* {} = {}";
-by (Step_tac 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (dres_inst_tac [("x","%n. xa n")] bspec 1);
-by (auto_tac (claset(), simpset() addsimps [FreeUltrafilterNat_empty]));
-qed "NatStar_empty_set";
-
-Addsimps [NatStar_empty_set];
-
-Goalw [starsetNat_def] 
-      "*sNat* (A Un B) = *sNat* A Un *sNat* B";
-by (Auto_tac);
-by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2));
-by (dtac FreeUltrafilterNat_Compl_mem 1);
-by (dtac bspec 1 THEN assume_tac 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (Auto_tac);
-by (Fuf_tac 1);
-qed "NatStar_Un";
-
-Goalw [starsetNat_n_def] 
-      "*sNatn* (%n. (A n) Un (B n)) = *sNatn* A Un *sNatn* B";
-by Auto_tac;
-by (dres_inst_tac [("x","Xa")] bspec 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 2);
-by (auto_tac (claset() addSDs [bspec], simpset()));
-by (TRYALL(Ultra_tac));
-qed "starsetNat_n_Un";
-
-Goalw [InternalNatSets_def]
-     "[| X : InternalNatSets; Y : InternalNatSets |] \
-\     ==> (X Un Y) : InternalNatSets";
-by (auto_tac (claset(),
-         simpset() addsimps [starsetNat_n_Un RS sym]));
-qed "InternalNatSets_Un";
-
-Goalw [starsetNat_def] "*sNat* (A Int B) = *sNat* A Int *sNat* B";
-by (Auto_tac);
-by (blast_tac (claset() addIs [FreeUltrafilterNat_Int,
-    FreeUltrafilterNat_subset]) 3);
-by (REPEAT(blast_tac (claset() addIs 
-    [FreeUltrafilterNat_subset]) 1));
-qed "NatStar_Int";
-
-Goalw [starsetNat_n_def] 
-      "*sNatn* (%n. (A n) Int (B n)) = *sNatn* A Int *sNatn* B";
-by (Auto_tac);
-by (auto_tac (claset() addSDs [bspec],
-         simpset()));
-by (TRYALL(Ultra_tac));
-qed "starsetNat_n_Int";
-
-Goalw [InternalNatSets_def]
-     "[| X : InternalNatSets; Y : InternalNatSets |] \
-\     ==> (X Int Y) : InternalNatSets";
-by (auto_tac (claset(),
-         simpset() addsimps [starsetNat_n_Int RS sym]));
-qed "InternalNatSets_Int";
-
-Goalw [starsetNat_def] "*sNat* (-A) = -(*sNat* A)";
-by (Auto_tac);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 2);
-by (REPEAT(Step_tac 1) THEN Auto_tac);
-by (TRYALL(Ultra_tac));
-qed "NatStar_Compl";
-
-Goalw [starsetNat_n_def] "*sNatn* ((%n. - A n)) = -(*sNatn* A)";
-by (Auto_tac);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 2);
-by (REPEAT(Step_tac 1) THEN Auto_tac);
-by (TRYALL(Ultra_tac));
-qed "starsetNat_n_Compl";
-
-Goalw [InternalNatSets_def]
-     "X :InternalNatSets ==> -X : InternalNatSets";
-by (auto_tac (claset(),
-         simpset() addsimps [starsetNat_n_Compl RS sym]));
-qed "InternalNatSets_Compl";
-
-Goalw [starsetNat_n_def] 
-      "*sNatn* (%n. (A n) - (B n)) = *sNatn* A - *sNatn* B";
-by (Auto_tac);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 2);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 3);
-by (auto_tac (claset() addSDs [bspec], simpset()));
-by (TRYALL(Ultra_tac));
-qed "starsetNat_n_diff";
-
-Goalw [InternalNatSets_def]
-     "[| X : InternalNatSets; Y : InternalNatSets |] \
-\     ==> (X - Y) : InternalNatSets";
-by (auto_tac (claset(), simpset() addsimps [starsetNat_n_diff RS sym]));
-qed "InternalNatSets_diff";
-
-Goalw [starsetNat_def] "A <= B ==> *sNat* A <= *sNat* B";
-by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1));
-qed "NatStar_subset";
-
-Goalw [starsetNat_def,hypnat_of_nat_def] 
-          "a : A ==> hypnat_of_nat a : *sNat* A";
-by (auto_tac (claset() addIs [FreeUltrafilterNat_subset],
-         simpset()));
-qed "NatStar_mem";
-
-Goalw [starsetNat_def] "hypnat_of_nat `` A <= *sNat* A";
-by (auto_tac (claset(), simpset() addsimps [hypnat_of_nat_def]));
-by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1);
-qed "NatStar_hypreal_of_real_image_subset";
-
-Goal "SHNat <= *sNat* (UNIV:: nat set)";
-by (simp_tac (simpset() addsimps [SHNat_hypnat_of_nat_iff,
-    NatStar_hypreal_of_real_image_subset]) 1);
-qed "NatStar_SHNat_subset";
-
-Goalw [starsetNat_def] 
-      "*sNat* X Int SHNat = hypnat_of_nat `` X";
-by (auto_tac (claset(),
-         simpset() addsimps 
-           [hypnat_of_nat_def,SHNat_def]));
-by (fold_tac [hypnat_of_nat_def]);
-by (rtac imageI 1 THEN rtac ccontr 1);
-by (dtac bspec 1);
-by (rtac lemma_hypnatrel_refl 1);
-by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2);
-by (Auto_tac);
-qed "NatStar_hypreal_of_real_Int";
-
-Goal "x ~: hypnat_of_nat `` A ==> ALL y: A. x ~= hypnat_of_nat y";
-by (Auto_tac);
-qed "lemma_not_hypnatA";
-
-Goalw [starsetNat_n_def,starsetNat_def] "*sNat* X = *sNatn* (%n. X)";
-by Auto_tac;
-qed "starsetNat_starsetNat_n_eq";
-
-Goalw [InternalNatSets_def] "(*sNat* X) : InternalNatSets";
-by (auto_tac (claset(),
-         simpset() addsimps [starsetNat_starsetNat_n_eq]));
-qed "InternalNatSets_starsetNat_n";
-Addsimps [InternalNatSets_starsetNat_n];
-
-Goal "X : InternalNatSets ==> UNIV - X : InternalNatSets";
-by (auto_tac (claset() addIs [InternalNatSets_Compl], simpset()));
-qed "InternalNatSets_UNIV_diff";
-
-(*------------------------------------------------------------------ 
-   Nonstandard extension of a set (defined using a constant 
-   sequence) as a special case of an internal set
- -----------------------------------------------------------------*)
-
-Goalw [starsetNat_n_def,starsetNat_def] 
-     "ALL n. (As n = A) ==> *sNatn* As = *sNat* A";
-by (Auto_tac);
-qed "starsetNat_n_starsetNat";
-
-(*------------------------------------------------------
-   Theorems about nonstandard extensions of functions
- ------------------------------------------------------*)
-
-(*------------------------------------------------------------------ 
-   Nonstandard extension of a function (defined using a constant 
-   sequence) as a special case of an internal function
- -----------------------------------------------------------------*)
-
-Goalw [starfunNat_n_def,starfunNat_def] 
-     "ALL n. (F n = f) ==> *fNatn* F = *fNat* f";
-by (Auto_tac);
-qed "starfunNat_n_starfunNat";
-
-Goalw [starfunNat2_n_def,starfunNat2_def] 
-     "ALL n. (F n = f) ==> *fNat2n* F = *fNat2* f";
-by (Auto_tac);
-qed "starfunNat2_n_starfunNat2";
-
-Goalw [congruent_def] 
-      "congruent hypnatrel (%X. hypnatrel^^{%n. f (X n)})";
-by (safe_tac (claset()));
-by (ALLGOALS(Fuf_tac));
-qed "starfunNat_congruent";
-
-(* f::nat=>real *)
-Goalw [starfunNat_def]
-      "(*fNat* f) (Abs_hypnat(hypnatrel^^{%n. X n})) = \
-\      Abs_hypreal(hyprel ^^ {%n. f (X n)})";
-by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
-by (simp_tac (simpset() addsimps 
-   [hyprel_in_hypreal RS Abs_hypreal_inverse]) 1);
-by (Auto_tac THEN Fuf_tac 1);
-qed "starfunNat";
-
-(* f::nat=>nat *)
-Goalw [starfunNat2_def]
-      "(*fNat2* f) (Abs_hypnat(hypnatrel^^{%n. X n})) = \
-\      Abs_hypnat(hypnatrel ^^ {%n. f (X n)})";
-by (res_inst_tac [("f","Abs_hypnat")] arg_cong 1);
-by (simp_tac (simpset() addsimps 
-   [hypnatrel_in_hypnat RS Abs_hypnat_inverse,
-    [equiv_hypnatrel, starfunNat_congruent] MRS UN_equiv_class]) 1);
-qed "starfunNat2";
-
-(*---------------------------------------------
-  multiplication: ( *f ) x ( *g ) = *(f x g)  
- ---------------------------------------------*)
-Goal "(*fNat* f) xa * (*fNat* g) xa = (*fNat* (%x. f x * g x)) xa";
-by (res_inst_tac [("z","xa")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat,hypreal_mult]));
-qed "starfunNat_mult";
-
-Goal "(*fNat2* f) xa * (*fNat2* g) xa = (*fNat2* (%x. f x * g x)) xa";
-by (res_inst_tac [("z","xa")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat2,hypnat_mult]));
-qed "starfunNat2_mult";
-
-(*---------------------------------------
-  addition: ( *f ) + ( *g ) = *(f + g)  
- ---------------------------------------*)
-Goal "(*fNat* f) xa + (*fNat* g) xa = (*fNat* (%x. f x + g x)) xa";
-by (res_inst_tac [("z","xa")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat,hypreal_add]));
-qed "starfunNat_add";
-
-Goal "(*fNat2* f) xa + (*fNat2* g) xa = (*fNat2* (%x. f x + g x)) xa";
-by (res_inst_tac [("z","xa")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat2,hypnat_add]));
-qed "starfunNat2_add";
-
-Goal "(*fNat2* f) xa - (*fNat2* g) xa = (*fNat2* (%x. f x - g x)) xa";
-by (res_inst_tac [("z","xa")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat2, hypnat_minus]));
-qed "starfunNat2_minus";
-
-(*--------------------------------------
-  composition: ( *f ) o ( *g ) = *(f o g)  
- ---------------------------------------*)
-(***** ( *f::nat=>real ) o ( *g::nat=>nat ) = *(f o g) *****)
- 
-Goal "(*fNat* f) o (*fNat2* g) = (*fNat* (f o g))";
-by (rtac ext 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat2, starfunNat]));
-qed "starfunNatNat2_o";
-
-Goal "(%x. (*fNat* f) ((*fNat2* g) x)) = (*fNat* (%x. f(g x)))";
-by (rtac ( simplify (simpset() addsimps [o_def]) starfunNatNat2_o) 1);
-qed "starfunNatNat2_o2";
-
-(***** ( *f::nat=>nat ) o ( *g::nat=>nat ) = *(f o g) *****)
-Goal "(*fNat2* f) o (*fNat2* g) = (*fNat2* (f o g))";
-by (rtac ext 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat2]));
-qed "starfunNat2_o";
-
-(***** ( *f::real=>real ) o ( *g::nat=>real ) = *(f o g) *****)
-
-Goal "(*f* f) o (*fNat* g) = (*fNat* (f o g))"; 
-by (rtac ext 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat,starfun]));
-qed "starfun_stafunNat_o";
-
-Goal "(%x. (*f* f) ((*fNat* g) x)) = (*fNat* (%x. f (g x)))"; 
-by (rtac ( simplify (simpset() addsimps [o_def]) starfun_stafunNat_o) 1);
-qed "starfun_stafunNat_o2";
-
-(*--------------------------------------
-  NS extension of constant function
- --------------------------------------*)
-Goal "(*fNat* (%x. k)) xa = hypreal_of_real k";
-by (res_inst_tac [("z","xa")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat, hypreal_of_real_def]));
-qed "starfunNat_const_fun";
-Addsimps [starfunNat_const_fun];
-
-Goal "(*fNat2* (%x. k)) xa = hypnat_of_nat  k";
-by (res_inst_tac [("z","xa")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat2, hypnat_of_nat_def]));
-qed "starfunNat2_const_fun";
-
-Addsimps [starfunNat2_const_fun];
-
-Goal "- (*fNat* f) x = (*fNat* (%x. - f x)) x";
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat, hypreal_minus]));
-qed "starfunNat_minus";
-
-Goal "inverse ((*fNat* f) x) = (*fNat* (%x. inverse (f x))) x";
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat, hypreal_inverse]));
-qed "starfunNat_inverse";
-
-(*--------------------------------------------------------
-   extented function has same solution as its standard
-   version for natural arguments. i.e they are the same
-   for all natural arguments (c.f. Hoskins pg. 107- SEQ)
- -------------------------------------------------------*)
-
-Goal "(*fNat* f) (hypnat_of_nat a) = hypreal_of_real (f a)";
-by (auto_tac (claset(),
-      simpset() addsimps [starfunNat,hypnat_of_nat_def,hypreal_of_real_def]));
-qed "starfunNat_eq";
-
-Addsimps [starfunNat_eq];
-
-Goal "(*fNat2* f) (hypnat_of_nat a) = hypnat_of_nat (f a)";
-by (auto_tac (claset(), simpset() addsimps [starfunNat2,hypnat_of_nat_def]));
-qed "starfunNat2_eq";
-
-Addsimps [starfunNat2_eq];
-
-Goal "(*fNat* f) (hypnat_of_nat a) @= hypreal_of_real (f a)";
-by (Auto_tac);
-qed "starfunNat_inf_close";
-
-
-(*-----------------------------------------------------------------
-    Example of transfer of a property from reals to hyperreals
-    --- used for limit comparison of sequences
- ----------------------------------------------------------------*)
-Goal "ALL n. N <= n --> f n <= g n \
-\         ==> ALL n. hypnat_of_nat N <= n --> (*fNat* f) n <= (*fNat* g) n";
-by (Step_tac 1);
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),
-         simpset() addsimps [starfunNat, hypnat_of_nat_def,hypreal_le,
-                             hypreal_less, hypnat_le,hypnat_less]));
-by (Ultra_tac 1);
-by Auto_tac;
-qed "starfun_le_mono";
-
-(*****----- and another -----*****) 
-Goal "ALL n. N <= n --> f n < g n \
-\         ==> ALL n. hypnat_of_nat N <= n --> (*fNat* f) n < (*fNat* g) n";
-by (Step_tac 1);
-by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),
-         simpset() addsimps [starfunNat, hypnat_of_nat_def,hypreal_le,
-                             hypreal_less, hypnat_le,hypnat_less]));
-by (Ultra_tac 1);
-by Auto_tac;
-qed "starfun_less_mono";
-
-(*----------------------------------------------------------------
-            NS extension when we displace argument by one
- ---------------------------------------------------------------*)
-Goal "(*fNat* (%n. f (Suc n))) N = (*fNat* f) (N + 1hn)";
-by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),
-         simpset() addsimps [starfunNat, hypnat_one_def,hypnat_add]));
-qed "starfunNat_shift_one";
-
-(*----------------------------------------------------------------
-                 NS extension with rabs    
- ---------------------------------------------------------------*)
-Goal "(*fNat* (%n. abs (f n))) N = abs((*fNat* f) N)";
-by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat, hypreal_hrabs]));
-qed "starfunNat_rabs";
-
-(*----------------------------------------------------------------
-       The hyperpow function as a NS extension of realpow
- ----------------------------------------------------------------*)
-Goal "(*fNat* (%n. r ^ n)) N = (hypreal_of_real r) pow N";
-by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),
-         simpset() addsimps [hyperpow, hypreal_of_real_def,starfunNat]));
-qed "starfunNat_pow";
-
-Goal "(*fNat* (%n. (X n) ^ m)) N = (*fNat* X) N pow hypnat_of_nat m";
-by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),
-         simpset() addsimps [hyperpow, hypnat_of_nat_def,starfunNat]));
-qed "starfunNat_pow2";
-
-Goalw [hypnat_of_nat_def] "(*f* (%r. r ^ n)) R = (R) pow hypnat_of_nat n";
-by (res_inst_tac [("z","R")] eq_Abs_hypreal 1);
-by (auto_tac (claset(), simpset() addsimps [hyperpow,starfun]));
-qed "starfun_pow";
-
-(*----------------------------------------------------- 
-   hypreal_of_hypnat as NS extension of real_of_nat! 
--------------------------------------------------------*)
-Goal "(*fNat* real_of_nat) = hypreal_of_hypnat";
-by (rtac ext 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [hypreal_of_hypnat,starfunNat]));
-qed "starfunNat_real_of_nat";
-
-Goal "N : HNatInfinite \
-\     ==> (*fNat* (%x. inverse (real_of_nat x))) N = inverse(hypreal_of_hypnat N)";
-by (res_inst_tac [("f1","inverse")]  (starfun_stafunNat_o2 RS subst) 1);
-by (subgoal_tac "hypreal_of_hypnat N ~= 0" 1);
-by (auto_tac (claset(), 
-       simpset() addsimps [starfunNat_real_of_nat, starfun_inverse_inverse]));
-qed "starfunNat_inverse_real_of_nat_eq";
-
-(*----------------------------------------------------------
-     Internal functions - some redundancy with *fNat* now
- ---------------------------------------------------------*)
-Goalw [congruent_def] 
-      "congruent hypnatrel (%X. hypnatrel^^{%n. f n (X n)})";
-by (safe_tac (claset()));
-by (ALLGOALS(Fuf_tac));
-qed "starfunNat_n_congruent";
-
-Goalw [starfunNat_n_def]
-      "(*fNatn* f) (Abs_hypnat(hypnatrel^^{%n. X n})) = \
-\      Abs_hypreal(hyprel ^^ {%n. f n (X n)})";
-by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
-by Auto_tac;
-by (Ultra_tac 1);
-qed "starfunNat_n";
-
-(*-------------------------------------------------
-  multiplication: ( *fn ) x ( *gn ) = *(fn x gn)  
- -------------------------------------------------*)
-Goal "(*fNatn* f) xa * (*fNatn* g) xa = \
-\         (*fNatn* (% i x. f i x * g i x)) xa";
-by (res_inst_tac [("z","xa")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat_n,hypreal_mult]));
-qed "starfunNat_n_mult";
-
-(*-----------------------------------------------
-  addition: ( *fn ) + ( *gn ) = *(fn + gn)  
- -----------------------------------------------*)
-Goal "(*fNatn* f) xa + (*fNatn* g) xa = \
-\         (*fNatn* (%i x. f i x + g i x)) xa";
-by (res_inst_tac [("z","xa")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat_n,hypreal_add]));
-qed "starfunNat_n_add";
-
-(*-------------------------------------------------
-  subtraction: ( *fn ) + -( *gn ) = *(fn + -gn)  
- -------------------------------------------------*)
-Goal "(*fNatn* f) xa + -(*fNatn* g) xa = \
-\         (*fNatn* (%i x. f i x + -g i x)) xa";
-by (res_inst_tac [("z","xa")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), 
-          simpset() addsimps [starfunNat_n, hypreal_minus,hypreal_add]));
-qed "starfunNat_n_add_minus";
-
-(*--------------------------------------------------
-  composition: ( *fn ) o ( *gn ) = *(fn o gn)  
- -------------------------------------------------*)
- 
-Goal "(*fNatn* (%i x. k)) xa = hypreal_of_real  k";
-by (res_inst_tac [("z","xa")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), 
-       simpset() addsimps [starfunNat_n, hypreal_of_real_def]));
-qed "starfunNat_n_const_fun";
-
-Addsimps [starfunNat_n_const_fun];
-
-Goal "- (*fNatn* f) x = (*fNatn* (%i x. - (f i) x)) x";
-by (res_inst_tac [("z","x")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat_n, hypreal_minus]));
-qed "starfunNat_n_minus";
-
-Goal "(*fNatn* f) (hypnat_of_nat n) = \
-\         Abs_hypreal(hyprel ^^ {%i. f i n})";
-by (auto_tac (claset(), simpset() addsimps [starfunNat_n,hypnat_of_nat_def]));
-qed "starfunNat_n_eq";
-Addsimps [starfunNat_n_eq];
-
-Goal "((*fNat* f) = (*fNat* g)) = (f = g)";
-by Auto_tac;
-by (rtac ext 1 THEN rtac ccontr 1);
-by (dres_inst_tac [("x","hypnat_of_nat(x)")] fun_cong 1);
-by (auto_tac (claset(), simpset() addsimps [starfunNat,hypnat_of_nat_def]));
-qed "starfun_eq_iff";
-
-
-
-
-
--- a/src/HOL/Real/Hyperreal/NatStar.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,46 +0,0 @@
-(*  Title       : NatStar.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : defining *-transforms in NSA which extends 
-                  sets of reals, and nat=>real, nat=>nat functions
-*) 
-
-NatStar = HyperNat + RealPow + HyperPow + 
-
-constdefs
-
-  (* internal sets and nonstandard extensions -- see Star.thy as well *)
-
-    starsetNat :: nat set => hypnat set          ("*sNat* _" [80] 80)
-    "*sNat* A  == {x. ALL X: Rep_hypnat(x). {n::nat. X n : A}: FreeUltrafilterNat}"
-
-    starsetNat_n :: (nat => nat set) => hypnat set        ("*sNatn* _" [80] 80)
-    "*sNatn* As  == {x. ALL X: Rep_hypnat(x). {n::nat. X n : (As n)}: FreeUltrafilterNat}"   
-
-    InternalNatSets :: "hypnat set set"
-    "InternalNatSets == {X. EX As. X = *sNatn* As}"
-
-    (* star transform of functions f:Nat --> Real *)
-
-    starfunNat :: (nat => real) => hypnat => hypreal        ("*fNat* _" [80] 80)
-    "*fNat* f  == (%x. Abs_hypreal(UN X: Rep_hypnat(x). hyprel^^{%n. f (X n)}))" 
-
-    starfunNat_n :: (nat => (nat => real)) => hypnat => hypreal        ("*fNatn* _" [80] 80)
-    "*fNatn* F  == (%x. Abs_hypreal(UN X: Rep_hypnat(x). hyprel^^{%n. (F n)(X n)}))" 
-
-    InternalNatFuns :: (hypnat => hypreal) set
-    "InternalNatFuns == {X. EX F. X = *fNatn* F}"
-
-    (* star transform of functions f:Nat --> Nat *)
-
-    starfunNat2 :: (nat => nat) => hypnat => hypnat        ("*fNat2* _" [80] 80)
-    "*fNat2* f  == (%x. Abs_hypnat(UN X: Rep_hypnat(x). hypnatrel^^{%n. f (X n)}))" 
-
-    starfunNat2_n :: (nat => (nat => nat)) => hypnat => hypnat        ("*fNat2n* _" [80] 80)
-    "*fNat2n* F  == (%x. Abs_hypnat(UN X: Rep_hypnat(x). hypnatrel^^{%n. (F n)(X n)}))" 
-
-    InternalNatFuns2 :: (hypnat => hypnat) set
-    "InternalNatFuns2 == {X. EX F. X = *fNat2n* F}"
-end
-
-
--- a/src/HOL/Real/Hyperreal/SEQ.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1379 +0,0 @@
-(*  Title       : SEQ.ML
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Theory of sequence and series of real numbers
-*) 
-
-(*---------------------------------------------------------------------------
-   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 
- -------------------------------------------------------------------------- *)
-
-Goalw [hypnat_omega_def] 
-      "(*fNat* (%n::nat. inverse(real_of_posnat n))) whn : Infinitesimal";
-by (auto_tac (claset(),simpset() addsimps 
-    [Infinitesimal_FreeUltrafilterNat_iff,starfunNat]));
-by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
-by (auto_tac (claset(),simpset() addsimps (map rename_numerals) 
-    [real_of_posnat_gt_zero,real_inverse_gt_zero,abs_eqI2,
-     FreeUltrafilterNat_inverse_real_of_posnat]));
-qed "SEQ_Infinitesimal";
-
-(*--------------------------------------------------------------------------
-                  Rules for LIMSEQ and NSLIMSEQ etc.
- --------------------------------------------------------------------------*)
-
-(*** LIMSEQ ***)
-Goalw [LIMSEQ_def] 
-      "X ----> L ==> \
-\      ALL r. #0 < r --> (EX no. ALL n. no <= n --> abs(X n + -L) < r)";
-by (Asm_simp_tac 1);
-qed "LIMSEQD1";
-
-Goalw [LIMSEQ_def] 
-      "[| X ----> L; #0 < r|] ==> \
-\      EX no. ALL n. no <= n --> abs(X n + -L) < r";
-by (Asm_simp_tac 1);
-qed "LIMSEQD2";
-
-Goalw [LIMSEQ_def] 
-      "ALL r. #0 < r --> (EX no. ALL n. \
-\      no <= n --> abs(X n + -L) < r) ==> X ----> L";
-by (Asm_simp_tac 1);
-qed "LIMSEQI";
-
-Goalw [LIMSEQ_def] 
-      "(X ----> L) = \
-\      (ALL r. #0 <r --> (EX no. ALL n. no <= n --> abs(X n + -L) < r))";
-by (Simp_tac 1);
-qed "LIMSEQ_iff";
-
-(*** NSLIMSEQ ***)
-Goalw [NSLIMSEQ_def] 
-      "X ----NS> L ==> ALL N: HNatInfinite. (*fNat* X) N @= hypreal_of_real L";
-by (Asm_simp_tac 1);
-qed "NSLIMSEQD1";
-
-Goalw [NSLIMSEQ_def] 
-      "[| X ----NS> L; N: HNatInfinite |] ==> (*fNat* X) N @= hypreal_of_real L";
-by (Asm_simp_tac 1);
-qed "NSLIMSEQD2";
-
-Goalw [NSLIMSEQ_def] 
-      "ALL N: HNatInfinite. (*fNat* X) N @= hypreal_of_real L ==> X ----NS> L";
-by (Asm_simp_tac 1);
-qed "NSLIMSEQI";
-
-Goalw [NSLIMSEQ_def] 
-      "(X ----NS> L) = (ALL N: HNatInfinite. (*fNat* X) N @= hypreal_of_real L)";
-by (Simp_tac 1);
-qed "NSLIMSEQ_iff";
-
-(*----------------------------------------
-          LIMSEQ ==> NSLIMSEQ
- ---------------------------------------*)
-Goalw [LIMSEQ_def,NSLIMSEQ_def] 
-      "X ----> L ==> X ----NS> L";
-by (auto_tac (claset(),simpset() addsimps 
-    [HNatInfinite_FreeUltrafilterNat_iff]));
-by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
-by (rtac (inf_close_minus_iff RS iffD2) 1);
-by (auto_tac (claset(),simpset() addsimps [starfunNat,
-    mem_infmal_iff RS sym,hypreal_of_real_def,
-    hypreal_minus,hypreal_add,
-    Infinitesimal_FreeUltrafilterNat_iff]));
-by (EVERY[rtac bexI 1, rtac lemma_hyprel_refl 2, Step_tac 1]);
-by (dres_inst_tac [("x","u")] spec 1 THEN Step_tac 1);
-by (dres_inst_tac [("x","no")] spec 1);
-by (Fuf_tac 1);
-by (blast_tac (claset() addDs [less_imp_le]) 1);
-qed "LIMSEQ_NSLIMSEQ";
-
-(*-------------------------------------------------------------
-          NSLIMSEQ ==> LIMSEQ
-    proving NS def ==> Standard def is trickier as usual 
- -------------------------------------------------------------*)
-(* the following sequence f(n) defines a hypernatural *)
-(* lemmas etc. first *)
-Goal "!!(f::nat=>nat). ALL n. n <= f n \
-\          ==> {n. f n = 0} = {0} | {n. f n = 0} = {}";
-by (Auto_tac);
-by (dres_inst_tac [("x","xa")] spec 1);
-by (dres_inst_tac [("x","x")] spec 2);
-by (Auto_tac);
-val lemma_NSLIMSEQ1 = result();
-
-Goal "{n. f n <= Suc u} = {n. f n <= u} Un {n. f n = Suc u}";
-by (auto_tac (claset(),simpset() addsimps [le_Suc_eq]));
-val lemma_NSLIMSEQ2 = result();
-
-Goal "!!(f::nat=>nat). ALL n. n <= f n \
-\          ==> {n. f n = Suc u} <= {n. n <= Suc u}";
-by (Auto_tac);
-by (dres_inst_tac [("x","x")] spec 1);
-by (Auto_tac);
-val lemma_NSLIMSEQ3 = result();
-
-Goal "!!(f::nat=>nat). ALL n. n <= f n \ 
-\         ==> finite {n. f n <= u}";
-by (induct_tac "u" 1);
-by (auto_tac (claset(),simpset() addsimps [lemma_NSLIMSEQ2]));
-by (auto_tac (claset() addIs [(lemma_NSLIMSEQ3 RS finite_subset),
-    finite_nat_le_segment], simpset()));
-by (dtac lemma_NSLIMSEQ1 1 THEN Step_tac 1);
-by (ALLGOALS(Asm_simp_tac));
-qed "NSLIMSEQ_finite_set";
-
-Goal "- {n. u < (f::nat=>nat) n} = {n. f n <= u}";
-by (auto_tac (claset() addDs [less_le_trans],
-    simpset() addsimps [le_def]));
-qed "Compl_less_set";
-
-(* the index set is in the free ultrafilter *)
-Goal "!!(f::nat=>nat). ALL n. n <= f n \ 
-\         ==> {n. u < f n} : FreeUltrafilterNat";
-by (rtac (FreeUltrafilterNat_Compl_iff2 RS iffD2) 1);
-by (rtac FreeUltrafilterNat_finite 1);
-by (auto_tac (claset() addDs [NSLIMSEQ_finite_set],
-    simpset() addsimps [Compl_less_set]));
-qed "FreeUltrafilterNat_NSLIMSEQ";
-
-(* thus, the sequence defines an infinite hypernatural! *)
-Goal "ALL n. n <= f n \
-\         ==> Abs_hypnat (hypnatrel ^^ {f}) : HNatInfinite";
-by (auto_tac (claset(),simpset() addsimps [HNatInfinite_FreeUltrafilterNat_iff]));
-by (EVERY[rtac bexI 1, rtac lemma_hypnatrel_refl 2, Step_tac 1]);
-by (etac FreeUltrafilterNat_NSLIMSEQ 1);
-qed "HNatInfinite_NSLIMSEQ";
-
-val lemmaLIM = CLAIM  "{n. X (f n) + - L = Y n} Int {n. abs (Y n) < r} <= \
-\         {n. abs (X (f n) + - L) < r}";
-
-Goal "{n. abs (X (f n) + - L) < r} Int \
-\         {n. r <= abs (X (f n) + - (L::real))} = {}";
-by (auto_tac (claset() addDs [real_less_le_trans] 
-    addIs [real_less_irrefl], simpset()));
-val lemmaLIM2 = result();
-
-Goal "[| #0 < r; ALL n. r <= abs (X (f n) + - L); \
-\          (*fNat* X) (Abs_hypnat (hypnatrel ^^ {f})) + \
-\          - hypreal_of_real  L @= 0 |] ==> False";
-by (auto_tac (claset(),simpset() addsimps [starfunNat,
-    mem_infmal_iff RS sym,hypreal_of_real_def,
-    hypreal_minus,hypreal_add,
-    Infinitesimal_FreeUltrafilterNat_iff]));
-by (dres_inst_tac [("x","r")] spec 1 THEN Step_tac 1);
-by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
-by (dtac (lemmaLIM RSN (2,FreeUltrafilterNat_subset)) 1);
-by (dtac FreeUltrafilterNat_all 1);
-by (thin_tac "{n. abs (Y n) < r} : FreeUltrafilterNat" 1);
-by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [lemmaLIM2,
-    FreeUltrafilterNat_empty]) 1);
-val lemmaLIM3 = result();
-
-Goalw [LIMSEQ_def,NSLIMSEQ_def] 
-      "X ----NS> L ==> X ----> L";
-by (rtac ccontr 1 THEN Asm_full_simp_tac 1);
-by (Step_tac 1);
-(* skolemization step *)
-by (dtac choice 1 THEN Step_tac 1);
-by (dres_inst_tac [("x","Abs_hypnat(hypnatrel^^{f})")] bspec 1);
-by (dtac (inf_close_minus_iff RS iffD1) 2);
-by (fold_tac [real_le_def]);
-by (blast_tac (claset() addIs [HNatInfinite_NSLIMSEQ]) 1);
-by (blast_tac (claset() addIs [rename_numerals lemmaLIM3]) 1);
-qed "NSLIMSEQ_LIMSEQ";
-
-(* Now the all important result is trivially proved! *)
-Goal "(f ----> L) = (f ----NS> L)";
-by (blast_tac (claset() addIs [LIMSEQ_NSLIMSEQ,NSLIMSEQ_LIMSEQ]) 1);
-qed "LIMSEQ_NSLIMSEQ_iff";
-
-(*-------------------------------------------------------------------
-                   Theorems about sequences
- ------------------------------------------------------------------*)
-Goalw [NSLIMSEQ_def] "(%n. k) ----NS> k";
-by (Auto_tac);
-qed "NSLIMSEQ_const";
-
-Goalw [LIMSEQ_def] "(%n. k) ----> k";
-by (Auto_tac);
-qed "LIMSEQ_const";
-
-Goalw [NSLIMSEQ_def]
-      "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n + Y n) ----NS> a + b";
-by (auto_tac (claset() addIs [inf_close_add],
-    simpset() addsimps [starfunNat_add RS sym]));
-qed "NSLIMSEQ_add";
-
-Goal "[| X ----> a; Y ----> b |] ==> (%n. X n + Y n) ----> a + b";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
-                                           NSLIMSEQ_add]) 1);
-qed "LIMSEQ_add";
-
-Goalw [NSLIMSEQ_def]
-      "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n * Y n) ----NS> a * b";
-by (auto_tac (claset() addSIs [inf_close_mult_HFinite],
-    simpset() addsimps [hypreal_of_real_mult, starfunNat_mult RS sym]));
-qed "NSLIMSEQ_mult";
-
-Goal "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
-                                           NSLIMSEQ_mult]) 1);
-qed "LIMSEQ_mult";
-
-Goalw [NSLIMSEQ_def] 
-     "X ----NS> a ==> (%n. -(X n)) ----NS> -a";
-by (auto_tac (claset(), simpset() addsimps [starfunNat_minus RS sym]));
-qed "NSLIMSEQ_minus";
-
-Goal "X ----> a ==> (%n. -(X n)) ----> -a";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
-                                           NSLIMSEQ_minus]) 1);
-qed "LIMSEQ_minus";
-
-Goal "(%n. -(X n)) ----> -a ==> X ----> a";
-by (dtac LIMSEQ_minus 1);
-by (Asm_full_simp_tac 1);
-qed "LIMSEQ_minus_cancel";
-
-Goal "(%n. -(X n)) ----NS> -a ==> X ----NS> a";
-by (dtac NSLIMSEQ_minus 1);
-by (Asm_full_simp_tac 1);
-qed "NSLIMSEQ_minus_cancel";
-
-Goal "[| X ----NS> a; Y ----NS> b |] \
-\               ==> (%n. X n + -Y n) ----NS> a + -b";
-by (dres_inst_tac [("X","Y")] NSLIMSEQ_minus 1);
-by (auto_tac (claset(),simpset() addsimps [NSLIMSEQ_add]));
-qed "NSLIMSEQ_add_minus";
-
-Goal "[| X ----> a; Y ----> b |] \
-\               ==> (%n. X n + -Y n) ----> a + -b";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
-    NSLIMSEQ_add_minus]) 1);
-qed "LIMSEQ_add_minus";
-
-Goalw [real_diff_def] 
-     "[| X ----> a; Y ----> b |] ==> (%n. X n - Y n) ----> a - b";
-by (blast_tac (claset() addIs [LIMSEQ_add_minus]) 1);
-qed "LIMSEQ_diff";
-
-Goalw [real_diff_def] 
-     "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n - Y n) ----NS> a - b";
-by (blast_tac (claset() addIs [NSLIMSEQ_add_minus]) 1);
-qed "NSLIMSEQ_diff";
-
-(*---------------------------------------------------------------
-    Proof is like that of NSLIM_inverse.
- --------------------------------------------------------------*)
-Goalw [NSLIMSEQ_def] 
-     "[| X ----NS> a;  a ~= #0 |] ==> (%n. inverse(X n)) ----NS> inverse(a)";
-by (Clarify_tac 1);
-by (dtac bspec 1);
-by (auto_tac (claset(), 
-              simpset() addsimps [starfunNat_inverse RS sym, 
-                                  hypreal_of_real_inf_close_inverse]));  
-qed "NSLIMSEQ_inverse";
-
-
-(*------ Standard version of theorem -------*)
-Goal "[| X ----> a; a ~= #0 |] ==> (%n. inverse(X n)) ----> inverse(a)";
-by (asm_full_simp_tac (simpset() addsimps [NSLIMSEQ_inverse,
-    LIMSEQ_NSLIMSEQ_iff]) 1);
-qed "LIMSEQ_inverse";
-
-Goal "[| X ----NS> a;  Y ----NS> b;  b ~= #0 |] \
-\     ==> (%n. X n / Y n) ----NS> a/b";
-by (asm_full_simp_tac (simpset() addsimps [NSLIMSEQ_mult, NSLIMSEQ_inverse, 
-                                           real_divide_def]) 1);
-qed "NSLIMSEQ_mult_inverse";
-
-Goal "[| X ----> a;  Y ----> b;  b ~= #0 |] ==> (%n. X n / Y n) ----> a/b";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_mult, LIMSEQ_inverse, 
-                                           real_divide_def]) 1);
-qed "LIMSEQ_divide";
-
-(*-----------------------------------------------
-            Uniqueness of limit
- ----------------------------------------------*)
-Goalw [NSLIMSEQ_def] 
-     "[| X ----NS> a; X ----NS> b |] ==> a = b";
-by (REPEAT(dtac (HNatInfinite_whn RSN (2,bspec)) 1));
-by (auto_tac (claset() addDs [inf_close_trans3], simpset()));
-qed "NSLIMSEQ_unique";
-
-Goal "[| X ----> a; X ----> b |] ==> a = b";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
-    NSLIMSEQ_unique]) 1);
-qed "LIMSEQ_unique";
-
-(*-----------------------------------------------------------------
-    theorems about nslim and lim
- ----------------------------------------------------------------*)
-Goalw [lim_def] "X ----> L ==> lim X = L";
-by (blast_tac (claset() addIs [LIMSEQ_unique]) 1);
-qed "limI";
-
-Goalw [nslim_def] "X ----NS> L ==> nslim X = L";
-by (blast_tac (claset() addIs [NSLIMSEQ_unique]) 1);
-qed "nslimI";
-
-Goalw [lim_def,nslim_def] "lim X = nslim X";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff]) 1);
-qed "lim_nslim_iff";
-
-(*------------------------------------------------------------------
-                      Convergence
- -----------------------------------------------------------------*)
-Goalw [convergent_def]
-     "convergent X ==> EX L. (X ----> L)";
-by (assume_tac 1);
-qed "convergentD";
-
-Goalw [convergent_def]
-     "(X ----> L) ==> convergent X";
-by (Blast_tac 1);
-qed "convergentI";
-
-Goalw [NSconvergent_def]
-     "NSconvergent X ==> EX L. (X ----NS> L)";
-by (assume_tac 1);
-qed "NSconvergentD";
-
-Goalw [NSconvergent_def]
-     "(X ----NS> L) ==> NSconvergent X";
-by (Blast_tac 1);
-qed "NSconvergentI";
-
-Goalw [convergent_def,NSconvergent_def] 
-     "convergent X = NSconvergent X";
-by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff]) 1);
-qed "convergent_NSconvergent_iff";
-
-Goalw [NSconvergent_def,nslim_def] 
-     "NSconvergent X = (X ----NS> nslim X)";
-by (auto_tac (claset() addIs [someI], simpset()));
-qed "NSconvergent_NSLIMSEQ_iff";
-
-Goalw [convergent_def,lim_def] 
-     "convergent X = (X ----> lim X)";
-by (auto_tac (claset() addIs [someI], simpset()));
-qed "convergent_LIMSEQ_iff";
-
-(*-------------------------------------------------------------------
-         Subsequence (alternative definition) (e.g. Hoskins)
- ------------------------------------------------------------------*)
-Goalw [subseq_def] "subseq f = (ALL n. (f n) < (f (Suc n)))";
-by (auto_tac (claset() addSDs [less_imp_Suc_add], simpset()));
-by (nat_ind_tac "k" 1);
-by (auto_tac (claset() addIs [less_trans], simpset()));
-qed "subseq_Suc_iff";
-
-(*-------------------------------------------------------------------
-                   Monotonicity
- ------------------------------------------------------------------*)
-
-Goalw [monoseq_def]
-   "monoseq X = ((ALL n. X n <= X (Suc n)) \
-\                | (ALL n. X (Suc n) <= X n))";
-by (auto_tac (claset () addSDs [le_imp_less_or_eq],
-    simpset() addsimps [real_le_refl]));
-by (auto_tac (claset() addSIs [lessI RS less_imp_le]
-                       addSDs [less_imp_Suc_add], 
-    simpset()));
-by (induct_tac "ka" 1);
-by (auto_tac (claset() addIs [real_le_trans], simpset()));
-by (EVERY1[rtac ccontr, rtac swap, Simp_tac]);
-by (induct_tac "k" 1);
-by (auto_tac (claset() addIs [real_le_trans], simpset()));
-qed "monoseq_Suc";
-
-Goalw [monoseq_def] 
-       "ALL m n. m <= n --> X m <= X n ==> monoseq X";
-by (Blast_tac 1);
-qed "monoI1";
-
-Goalw [monoseq_def] 
-       "ALL m n. m <= n --> X n <= X m ==> monoseq X";
-by (Blast_tac 1);
-qed "monoI2";
-
-Goal "ALL n. X n <= X (Suc n) ==> monoseq X";
-by (asm_simp_tac (simpset() addsimps [monoseq_Suc]) 1);
-qed "mono_SucI1";
-
-Goal "ALL n. X (Suc n) <= X n ==> monoseq X";
-by (asm_simp_tac (simpset() addsimps [monoseq_Suc]) 1);
-qed "mono_SucI2";
-
-(*-------------------------------------------------------------------
-                  Bounded Sequence
- ------------------------------------------------------------------*)
-Goalw [Bseq_def] 
-      "Bseq X ==> EX K. #0 < K & (ALL n. abs(X n) <= K)";
-by (assume_tac 1);
-qed "BseqD";
-
-Goalw [Bseq_def]
-      "[| #0 < K; ALL n. abs(X n) <= K |] \
-\           ==> Bseq X";
-by (Blast_tac 1);
-qed "BseqI";
-
-Goal "(EX K. #0 < K & (ALL n. abs(X n) <= K)) = \
-\         (EX N. ALL n. abs(X n) <= real_of_posnat N)";
-by (auto_tac (claset(),simpset() addsimps 
-    (map rename_numerals) [real_gt_zero_preal_Ex,real_of_posnat_gt_zero]));
-by (cut_inst_tac [("x","real_of_preal y")] reals_Archimedean2 1);
-by (blast_tac (claset() addIs [real_le_less_trans,
-    real_less_imp_le]) 1);
-by (auto_tac (claset(),simpset() addsimps [real_of_posnat_def]));
-qed "lemma_NBseq_def";
-
-(* alternative definition for Bseq *)
-Goalw [Bseq_def] 
-      "Bseq X = (EX N. ALL n. abs(X n) <= real_of_posnat N)";
-by (simp_tac (simpset() addsimps [lemma_NBseq_def]) 1);
-qed "Bseq_iff";
-
-Goal "(EX K. #0 < K & (ALL n. abs(X n) <= K)) = \
-\         (EX N. ALL n. abs(X n) < real_of_posnat N)";
-by (auto_tac (claset(),simpset() addsimps 
-    (map rename_numerals) [real_gt_zero_preal_Ex,real_of_posnat_gt_zero]));
-by (cut_inst_tac [("x","real_of_preal y")] reals_Archimedean2 1);
-by (blast_tac (claset() addIs [real_less_trans,
-    real_le_less_trans]) 1);
-by (auto_tac (claset() addIs [real_less_imp_le],
-    simpset() addsimps [real_of_posnat_def]));
-qed "lemma_NBseq_def2";
-
-(* yet another definition for Bseq *)
-Goalw [Bseq_def] 
-      "Bseq X = (EX N. ALL n. abs(X n) < real_of_posnat N)";
-by (simp_tac (simpset() addsimps [lemma_NBseq_def2]) 1);
-qed "Bseq_iff1a";
-
-Goalw [NSBseq_def]
-      "[| NSBseq X; N: HNatInfinite |] \
-\           ==> (*fNat* X) N : HFinite";
-by (Blast_tac 1);
-qed "NSBseqD";
-
-Goalw [NSBseq_def]
-      "ALL N: HNatInfinite. (*fNat* X) N : HFinite \
-\           ==> NSBseq X";
-by (assume_tac 1);
-qed "NSBseqI";
-
-(*-----------------------------------------------------------
-       Standard definition ==> NS definition
- ----------------------------------------------------------*)
-(* a few lemmas *)
-Goal "ALL n. abs(X n) <= K ==> \
-\     ALL n. abs(X((f::nat=>nat) n)) <= K";
-by (Auto_tac);
-val lemma_Bseq = result();
-
-Goalw [Bseq_def,NSBseq_def] "Bseq X ==> NSBseq X";
-by (Step_tac 1);
-by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
-by (auto_tac (claset(),simpset() addsimps [starfunNat,    
-    HFinite_FreeUltrafilterNat_iff,
-    HNatInfinite_FreeUltrafilterNat_iff]));
-by (EVERY[rtac bexI 1, rtac lemma_hyprel_refl 2]);
-by (dres_inst_tac [("f","Xa")] lemma_Bseq 1); 
-by (res_inst_tac [("x","K+#1")] exI 1);
-by (rotate_tac 2 1 THEN dtac FreeUltrafilterNat_all 1);
-by (Ultra_tac 1);
-qed "Bseq_NSBseq";
-
-(*---------------------------------------------------------------
-       NS  definition ==> Standard definition
- ---------------------------------------------------------------*)
-(* similar to NSLIM proof in REALTOPOS *)
-(*------------------------------------------------------------------- 
-   We need to get rid of the real variable and do so by proving the
-   following which relies on the Archimedean property of the reals
-   When we skolemize we then get the required function f::nat=>nat 
-   o/w we would be stuck with a skolem function f :: real=>nat which
-   is not what we want (read useless!)
- -------------------------------------------------------------------*)
- 
-Goal "ALL K. #0 < K --> (EX n. K < abs (X n)) \
-\          ==> ALL N. EX n. real_of_posnat  N < abs (X n)";
-by (Step_tac 1);
-by (cut_inst_tac [("n","N")] (rename_numerals real_of_posnat_gt_zero) 1);
-by (Blast_tac 1);
-val lemmaNSBseq = result();
-
-Goal "ALL K. #0 < K --> (EX n. K < abs (X n)) \
-\         ==> EX f. ALL N. real_of_posnat  N < abs (X (f N))";
-by (dtac lemmaNSBseq 1);
-by (dtac choice 1);
-by (Blast_tac 1);
-val lemmaNSBseq2 = result();
-
-Goal "ALL N. real_of_posnat  N < abs (X (f N)) \
-\         ==>  Abs_hypreal(hyprel^^{X o f}) : HInfinite";
-by (auto_tac (claset(),simpset() addsimps 
-    [HInfinite_FreeUltrafilterNat_iff,o_def]));
-by (EVERY[rtac bexI 1, rtac lemma_hyprel_refl 2, 
-    Step_tac 1]);
-by (cut_inst_tac [("u","u")] FreeUltrafilterNat_nat_gt_real 1);
-by (blast_tac (claset() addDs [FreeUltrafilterNat_all,
-    FreeUltrafilterNat_Int] addIs [real_less_trans,
-    FreeUltrafilterNat_subset]) 1);
-qed "real_seq_to_hypreal_HInfinite";
-
-(*--------------------------------------------------------------------------------
-     Now prove that we can get out an infinite hypernatural as well 
-     defined using the skolem function f::nat=>nat above
- --------------------------------------------------------------------------------*)
-
-Goal "{n. f n <= Suc u & real_of_posnat  n < abs (X (f n))} <= \
-\         {n. f n <= u & real_of_posnat  n < abs (X (f n))} \
-\         Un {n. real_of_posnat n < abs (X (Suc u))}";
-by (auto_tac (claset() addSDs [le_imp_less_or_eq] addIs [less_imp_le],
-    simpset() addsimps [less_Suc_eq]));
-val lemma_finite_NSBseq = result();
-
-Goal "finite {n. f n <= (u::nat) &  real_of_posnat n < abs(X(f n))}";
-by (induct_tac "u" 1);
-by (rtac (CLAIM "{n. f n <= (0::nat) & real_of_posnat n < abs (X (f n))} <= \
-\         {n. real_of_posnat n < abs (X 0)}"
-          RS finite_subset) 1);
-by (rtac finite_real_of_posnat_less_real 1);
-by (rtac (lemma_finite_NSBseq RS finite_subset) 1);
-by (auto_tac (claset() addIs [finite_real_of_posnat_less_real], simpset()));
-val lemma_finite_NSBseq2 = result();
-
-Goal "ALL N. real_of_posnat  N < abs (X (f N)) \
-\     ==> Abs_hypnat(hypnatrel^^{f}) : HNatInfinite";
-by (auto_tac (claset(),simpset() addsimps 
-    [HNatInfinite_FreeUltrafilterNat_iff]));
-by (EVERY[rtac bexI 1, rtac lemma_hypnatrel_refl 2, 
-    Step_tac 1]);
-by (rtac ccontr 1 THEN dtac FreeUltrafilterNat_Compl_mem 1);
-by (asm_full_simp_tac (simpset() addsimps 
-   [CLAIM_SIMP "- {n. u < (f::nat=>nat) n} \
-\   = {n. f n <= u}" [le_def]]) 1);
-by (dtac FreeUltrafilterNat_all 1);
-by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [CLAIM "({n. f n <= u} Int {n. real_of_posnat n < abs(X(f n))}) = \
-\          {n. f n <= (u::nat) &  real_of_posnat n < abs(X(f n))}",
-     lemma_finite_NSBseq2 RS FreeUltrafilterNat_finite]));
-qed "HNatInfinite_skolem_f";
-
-Goalw [Bseq_def,NSBseq_def] 
-      "NSBseq X ==> Bseq X";
-by (rtac ccontr 1);
-by (auto_tac (claset(),simpset() addsimps [real_le_def]));
-by (dtac lemmaNSBseq2 1 THEN Step_tac 1);
-by (forw_inst_tac [("X","X"),("f","f")] real_seq_to_hypreal_HInfinite 1);
-by (dtac (HNatInfinite_skolem_f RSN (2,bspec)) 1 THEN assume_tac 1);
-by (auto_tac (claset(),simpset() addsimps [starfunNat,
-    o_def,HFinite_HInfinite_iff]));
-qed "NSBseq_Bseq";
-
-(*----------------------------------------------------------------------
-  Equivalence of nonstandard and standard definitions 
-  for a bounded sequence
- -----------------------------------------------------------------------*)
-Goal "(Bseq X) = (NSBseq X)";
-by (blast_tac (claset() addSIs [NSBseq_Bseq,Bseq_NSBseq]) 1);
-qed "Bseq_NSBseq_iff";
-
-(*----------------------------------------------------------------------
-   A convergent sequence is bounded
-   (Boundedness as a necessary condition for convergence)
- -----------------------------------------------------------------------*)
-(* easier --- nonstandard version - no existential as usual *)
-Goalw [NSconvergent_def,NSBseq_def,NSLIMSEQ_def] 
-          "NSconvergent X ==> NSBseq X";
-by (blast_tac (claset() addDs [HFinite_hypreal_of_real RS 
-               (inf_close_sym RSN (2,inf_close_HFinite))]) 1);
-qed "NSconvergent_NSBseq";
-
-(* standard version - easily now proved using *)
-(* equivalence of NS and standard definitions *)
-Goal "convergent X ==> Bseq X";
-by (asm_full_simp_tac (simpset() addsimps [NSconvergent_NSBseq,
-    convergent_NSconvergent_iff,Bseq_NSBseq_iff]) 1);
-qed "convergent_Bseq";
-
-(*----------------------------------------------------------------------
-             Results about Ubs and Lubs of bounded sequences
- -----------------------------------------------------------------------*)
-Goalw [Bseq_def]
-  "!!(X::nat=>real). Bseq X ==> \
-\  EX U. isUb (UNIV::real set) {x. EX n. X n = x} U";
-by (auto_tac (claset() addIs [isUbI,setleI],
-    simpset() addsimps [abs_le_interval_iff]));
-qed "Bseq_isUb";
-
-(*----------------------------------------------------------------------
-   Use completeness of reals (supremum property) 
-   to show that any bounded sequence has a lub 
------------------------------------------------------------------------*)
-Goal
-  "!!(X::nat=>real). Bseq X ==> \
-\  EX U. isLub (UNIV::real set) {x. EX n. X n = x} U";
-by (blast_tac (claset() addIs [reals_complete,
-    Bseq_isUb]) 1);
-qed "Bseq_isLub";
-
-(* nonstandard version of premise will be *)
-(* handy when we work in NS universe      *)
-Goal   "NSBseq X ==> \
-\  EX U. isUb (UNIV::real set) {x. EX n. X n = x} U";
-by (asm_full_simp_tac (simpset() addsimps 
-    [Bseq_NSBseq_iff RS sym,Bseq_isUb]) 1);
-qed "NSBseq_isUb";
-
-Goal
-  "NSBseq X ==> \
-\  EX U. isLub (UNIV::real set) {x. EX n. X n = x} U";
-by (asm_full_simp_tac (simpset() addsimps 
-    [Bseq_NSBseq_iff RS sym,Bseq_isLub]) 1);
-qed "NSBseq_isLub";
-
-(*--------------------------------------------------------------------
-             Bounded and monotonic sequence converges              
- --------------------------------------------------------------------*)
-(* lemmas *)
-Goal 
-     "!!(X::nat=>real). [| ALL m n. m <= n -->  X m <= X n; \
-\                 isLub (UNIV::real set) {x. EX n. X n = x} (X ma) \
-\              |] ==> ALL n. ma <= n --> X n = X ma";
-by (Step_tac 1);
-by (dres_inst_tac [("y","X n")] isLubD2 1);
-by (ALLGOALS(blast_tac (claset() addDs [real_le_anti_sym])));
-val lemma_converg1 = result();
-
-(*------------------------------------------------------------------- 
-   The best of both world: Easier to prove this result as a standard
-   theorem and then use equivalence to "transfer" it into the
-   equivalent nonstandard form if needed!
- -------------------------------------------------------------------*)
-Goalw [LIMSEQ_def] 
-         "ALL n. m <= n --> X n = X m \
-\         ==> EX L. (X ----> L)";  
-by (res_inst_tac [("x","X m")] exI 1);
-by (Step_tac 1);
-by (res_inst_tac [("x","m")] exI 1);
-by (Step_tac 1);
-by (dtac spec 1 THEN etac impE 1);
-by (Auto_tac);
-qed "Bmonoseq_LIMSEQ";
-
-(* Now same theorem in terms of NS limit *)
-Goal "ALL n. m <= n --> X n = X m \
-\         ==> EX L. (X ----NS> L)";  
-by (auto_tac (claset() addSDs [Bmonoseq_LIMSEQ],
-    simpset() addsimps [LIMSEQ_NSLIMSEQ_iff]));
-qed "Bmonoseq_NSLIMSEQ";
-
-(* a few more lemmas *)
-Goal "!!(X::nat=>real). \
-\              [| ALL m. X m ~= U; \
-\                 isLub UNIV {x. EX n. X n = x} U \
-\              |] ==> ALL m. X m < U";
-by (Step_tac 1);
-by (dres_inst_tac [("y","X m")] isLubD2 1);
-by (auto_tac (claset() addSDs [real_le_imp_less_or_eq],
-              simpset()));
-val lemma_converg2 = result();
-
-Goal "!!(X ::nat=>real). ALL m. X m <= U ==> \
-\         isUb UNIV {x. EX n. X n = x} U";
-by (rtac (setleI RS isUbI) 1);
-by (Auto_tac);
-val lemma_converg3 = result();
-
-(* FIXME: U - T < U redundant *)
-Goal "!!(X::nat=> real). \
-\              [| ALL m. X m ~= U; \
-\                 isLub UNIV {x. EX n. X n = x} U; \
-\                 #0 < T; \
-\                 U + - T < U \
-\              |] ==> EX m. U + -T < X m & X m < U";
-by (dtac lemma_converg2 1 THEN assume_tac 1);
-by (rtac ccontr 1 THEN Asm_full_simp_tac 1);
-by (fold_tac [real_le_def]);
-by (dtac lemma_converg3 1);
-by (dtac isLub_le_isUb 1 THEN assume_tac 1);
-by (auto_tac (claset() addDs [real_less_le_trans],
-    simpset() addsimps [real_minus_zero_le_iff]));
-val lemma_converg4 = result();
-
-(*-------------------------------------------------------------------
-  A standard proof of the theorem for monotone increasing sequence
- ------------------------------------------------------------------*)
-
-Goalw [convergent_def] 
-     "[| Bseq X; ALL m n. m <= n --> X m <= X n |] \
-\                ==> convergent X";
-by (forward_tac [Bseq_isLub] 1);
-by (Step_tac 1);
-by (case_tac "EX m. X m = U" 1 THEN Auto_tac);
-by (blast_tac (claset() addDs [lemma_converg1,
-    Bmonoseq_LIMSEQ]) 1);
-(* second case *)
-by (res_inst_tac [("x","U")] exI 1);
-by (rtac LIMSEQI 1 THEN Step_tac 1);
-by (forward_tac [lemma_converg2] 1 THEN assume_tac 1);
-by (dtac lemma_converg4 1 THEN Auto_tac);
-by (res_inst_tac [("x","m")] exI 1 THEN Step_tac 1);
-by (subgoal_tac "X m <= X n" 1 THEN Fast_tac 2);
-by (rotate_tac 3 1 THEN dres_inst_tac [("x","n")] spec 1);
-by (arith_tac 1);
-qed "Bseq_mono_convergent";
-
-(* NS version of theorem *)
-Goalw [convergent_def] 
-     "[| NSBseq X; ALL m n. m <= n --> X m <= X n |] \
-\                ==> NSconvergent X";
-by (auto_tac (claset() addIs [Bseq_mono_convergent], 
-    simpset() addsimps [convergent_NSconvergent_iff RS sym,
-    Bseq_NSBseq_iff RS sym]));
-qed "NSBseq_mono_NSconvergent";
-
-Goalw [convergent_def] 
-      "(convergent X) = (convergent (%n. -(X n)))";
-by (auto_tac (claset() addDs [LIMSEQ_minus], simpset()));
-by (dtac LIMSEQ_minus 1 THEN Auto_tac);
-qed "convergent_minus_iff";
-
-Goalw [Bseq_def] "Bseq (%n. -(X n)) = Bseq X";
-by (Asm_full_simp_tac 1);
-qed "Bseq_minus_iff";
-
-(*--------------------------------
-   **** main mono theorem ****
- -------------------------------*)
-Goalw [monoseq_def] "[| Bseq X; monoseq X |] ==> convergent X";
-by (Step_tac 1);
-by (rtac (convergent_minus_iff RS ssubst) 2);
-by (dtac (Bseq_minus_iff RS ssubst) 2);
-by (auto_tac (claset() addSIs [Bseq_mono_convergent], simpset()));
-qed "Bseq_monoseq_convergent";
-
-(*----------------------------------------------------------------
-          A few more equivalence theorems for boundedness 
- ---------------------------------------------------------------*)
- 
-(***--- alternative formulation for boundedness---***)
-Goalw [Bseq_def] 
-   "Bseq X = (EX k x. #0 < k & (ALL n. abs(X(n) + -x) <= k))";
-by (Step_tac 1);
-by (res_inst_tac [("x","k + abs(x)")] exI 2);
-by (res_inst_tac [("x","K")] exI 1);
-by (res_inst_tac [("x","0")] exI 1);
-by (Auto_tac);
-by (ALLGOALS (dres_inst_tac [("x","n")] spec));
-by (ALLGOALS arith_tac);
-qed "Bseq_iff2";
-
-(***--- alternative formulation for boundedness ---***)
-Goal "Bseq X = (EX k N. #0 < k & (ALL n. abs(X(n) + -X(N)) <= k))";
-by (Step_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [Bseq_def]) 1);
-by (Step_tac 1);
-by (res_inst_tac [("x","K + abs(X N)")] exI 1);
-by (Auto_tac);
-by (arith_tac 1);
-by (res_inst_tac [("x","N")] exI 1);
-by (Step_tac 1);
-by (dres_inst_tac [("x","n")] spec 1);
-by (arith_tac 1);
-by (auto_tac (claset(), simpset() addsimps [Bseq_iff2]));
-qed "Bseq_iff3";
-
-Goalw [Bseq_def] "(ALL n. k <= f n & f n <= K) ==> Bseq f";
-by (res_inst_tac [("x","(abs(k) + abs(K)) + #1")] exI 1);
-by (Auto_tac);
-by (dres_inst_tac [("x","n")] spec 2);
-by (ALLGOALS arith_tac);
-qed "BseqI2";
-
-(*-------------------------------------------------------------------
-   Equivalence between NS and standard definitions of Cauchy seqs
- ------------------------------------------------------------------*)
-(*-------------------------------
-      Standard def => NS def
- -------------------------------*)
-Goal "Abs_hypnat (hypnatrel ^^ {x}) : HNatInfinite \
-\         ==> {n. M <= x n} : FreeUltrafilterNat";
-by (auto_tac (claset(),
-              simpset() addsimps [HNatInfinite_FreeUltrafilterNat_iff]));
-by (dres_inst_tac [("x","M")] spec 1);
-by (ultra_tac (claset(),simpset() addsimps [less_imp_le]) 1);
-val lemmaCauchy1 = result();
-
-Goal "{n. ALL m n. M <= m & M <= (n::nat) --> abs (X m + - X n) < u} Int \
-\     {n. M <= xa n} Int {n. M <= x n} <= \
-\     {n. abs (X (xa n) + - X (x n)) < u}";
-by (Blast_tac 1);
-val lemmaCauchy2 = result();
-
-Goalw [Cauchy_def,NSCauchy_def] 
-      "Cauchy X ==> NSCauchy X";
-by (Step_tac 1);
-by (res_inst_tac [("z","M")] eq_Abs_hypnat 1);
-by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
-by (rtac (inf_close_minus_iff RS iffD2) 1);
-by (rtac (mem_infmal_iff RS iffD1) 1);
-by (auto_tac (claset(),simpset() addsimps [starfunNat,
-    hypreal_minus,hypreal_add,Infinitesimal_FreeUltrafilterNat_iff]));
-by (EVERY[rtac bexI 1, Auto_tac]);
-by (dtac spec 1 THEN Auto_tac);
-by (dres_inst_tac [("M","M")] lemmaCauchy1 1);
-by (dres_inst_tac [("M","M")] lemmaCauchy1 1);
-by (res_inst_tac [("x1","xa")] 
-    (lemmaCauchy2 RSN (2,FreeUltrafilterNat_subset)) 1);
-by (rtac FreeUltrafilterNat_Int 1 THEN assume_tac 2);
-by (auto_tac (claset() addIs [FreeUltrafilterNat_Int,
-        FreeUltrafilterNat_Nat_set], simpset()));
-qed "Cauchy_NSCauchy";
-
-(*-----------------------------------------------
-     NS def => Standard def -- rather long but 
-     straightforward proof in this case
- ---------------------------------------------*)
-Goalw [Cauchy_def,NSCauchy_def] 
-      "NSCauchy X ==> Cauchy X";
-by (EVERY1[Step_tac, rtac ccontr,Asm_full_simp_tac]);
-by (dtac choice 1 THEN auto_tac (claset(),simpset() 
-         addsimps [all_conj_distrib]));
-by (dtac choice 1 THEN step_tac (claset() addSDs 
-         [all_conj_distrib RS iffD1]) 1);
-by (REPEAT(dtac HNatInfinite_NSLIMSEQ 1));
-by (dtac bspec 1 THEN assume_tac 1);
-by (dres_inst_tac [("x","Abs_hypnat (hypnatrel ^^ {fa})")] bspec 1 
-    THEN auto_tac (claset(),simpset() addsimps [starfunNat]));
-by (dtac (inf_close_minus_iff RS iffD1) 1);
-by (dtac (mem_infmal_iff RS iffD2) 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_minus,
-    hypreal_add,Infinitesimal_FreeUltrafilterNat_iff]));
-by (dres_inst_tac [("x","e")] spec 1 THEN Auto_tac);
-by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
-by (dtac (CLAIM "{n. X (f n) + - X (fa n) = Y n} Int \
-\         {n. abs (Y n) < e} <= \
-\         {n. abs (X (f n) + - X (fa n)) < e}" RSN 
-          (2,FreeUltrafilterNat_subset)) 1);
-by (thin_tac "{n. abs (Y n) < e} : FreeUltrafilterNat" 1);
-by (dtac FreeUltrafilterNat_all 1);
-by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [CLAIM "{n. abs (X (f n) + - X (fa n)) < e} Int \
-\         {M. ~ abs (X (f M) + - X (fa M)) < e} = {}",
-     FreeUltrafilterNat_empty]) 1);
-qed "NSCauchy_Cauchy";
-
-(*----- Equivalence -----*)
-Goal "NSCauchy X = Cauchy X";
-by (blast_tac (claset() addSIs[NSCauchy_Cauchy,
-    Cauchy_NSCauchy]) 1);
-qed "NSCauchy_Cauchy_iff";
-
-(*-------------------------------------------------------
-  Cauchy sequence is bounded -- this is the standard 
-  proof mechanization rather than the nonstandard proof 
- -------------------------------------------------------*)
-
-(***-------------  VARIOUS LEMMAS --------------***)
-Goal "ALL n. M <= n --> abs (X M + - X n) < (#1::real) \
-\         ==>  ALL n. M <= n --> abs(X n) < #1 + abs(X M)";
-by (Step_tac 1);
-by (dtac spec 1 THEN Auto_tac);
-by (arith_tac 1);
-val lemmaCauchy = result();
-
-Goal "(n < Suc M) = (n <= M)";
-by Auto_tac;
-qed "less_Suc_cancel_iff";
-
-(* FIXME: Long. Maximal element in subsequence *)
-Goal "EX m. m <= M & (ALL n. n <= M --> \
-\         abs ((X::nat=> real) n) <= abs (X m))";
-by (induct_tac "M" 1);
-by (res_inst_tac [("x","0")] exI 1);
-by (Asm_full_simp_tac 1);
-by (Step_tac 1);
-by (cut_inst_tac [("R1.0","abs (X (Suc n))"),("R2.0","abs(X m)")]
-        real_linear 1);
-by (Step_tac 1);
-by (res_inst_tac [("x","m")] exI 1);
-by (res_inst_tac [("x","m")] exI 2);
-by (res_inst_tac [("x","Suc n")] exI 3);
-by (ALLGOALS(Asm_full_simp_tac));
-by (Step_tac 1);
-by (ALLGOALS(eres_inst_tac [("m1","na")] 
-    (le_imp_less_or_eq RS disjE)));
-by (ALLGOALS(asm_full_simp_tac (simpset() addsimps 
-    [real_le_refl,less_Suc_cancel_iff,
-     real_less_imp_le])));
-by (blast_tac (claset() addIs [real_le_less_trans RS
-    real_less_imp_le]) 1);
-qed "SUP_rabs_subseq";
-
-(* lemmas to help proof - mostly trivial *)
-Goal "[| ALL m::nat. m <= M --> P M m; \
-\        ALL m. M <= m --> P M m |] \
-\     ==> ALL m. P M m";
-by (Step_tac 1);
-by (REPEAT(dres_inst_tac [("x","m")] spec 1));
-by (auto_tac (claset() addEs [less_asym],
-    simpset() addsimps [le_def]));
-val lemma_Nat_covered = result();
-
-Goal "[| ALL n. n <= M --> abs ((X::nat=>real) n) <= a; \
-\        a < b |] \
-\     ==> ALL n. n <= M --> abs(X n) <= b";
-by (blast_tac (claset() addIs [real_le_less_trans RS 
-               real_less_imp_le]) 1);
-val lemma_trans1 = result();
-
-Goal "[| ALL n. M <= n --> abs ((X::nat=>real) n) < a; \
-\        a < b |] \
-\     ==> ALL n. M <= n --> abs(X n)<= b";
-by (blast_tac (claset() addIs [real_less_trans RS 
-               real_less_imp_le]) 1);
-val lemma_trans2 = result();
-
-Goal "[| ALL n. n <= M --> abs (X n) <= a; \
-\        a = b |] \
-\     ==> ALL n. n <= M --> abs(X n) <= b";
-by (Auto_tac);
-val lemma_trans3 = result();
-
-Goal "ALL n. M <= n --> abs ((X::nat=>real) n) < a \
-\             ==>  ALL n. M <= n --> abs (X n) <= a";
-by (blast_tac (claset() addIs [real_less_imp_le]) 1);
-val lemma_trans4 = result();
-
-(*---------------------------------------------------------- 
-   Trickier than expected --- proof is more involved than
-   outlines sketched by various authors would suggest
- ---------------------------------------------------------*)
-Goalw [Cauchy_def,Bseq_def] "Cauchy X ==> Bseq X";
-by (dres_inst_tac [("x","#1")] spec 1);
-by (etac (rename_numerals real_zero_less_one RSN (2,impE)) 1);
-by (Step_tac 1);
-by (dres_inst_tac [("x","M")] spec 1);
-by (Asm_full_simp_tac 1);
-by (dtac lemmaCauchy 1);
-by (cut_inst_tac [("M","M"),("X","X")] SUP_rabs_subseq 1);
-by (Step_tac 1);
-by (cut_inst_tac [("R1.0","abs(X m)"),
-     ("R2.0","#1 + abs(X M)")] real_linear 1);
-by (Step_tac 1);
-by (dtac lemma_trans1 1 THEN assume_tac 1);
-by (dtac lemma_trans2 3 THEN assume_tac 3);
-by (dtac lemma_trans3 2 THEN assume_tac 2);
-by (dtac (abs_add_one_gt_zero RS real_less_trans) 3);
-by (dtac lemma_trans4 1);
-by (dtac lemma_trans4 2);
-by (res_inst_tac [("x","#1 + abs(X M)")] exI 1);
-by (res_inst_tac [("x","#1 + abs(X M)")] exI 2);
-by (res_inst_tac [("x","abs(X m)")] exI 3);
-by (auto_tac (claset() addSEs [lemma_Nat_covered],
-              simpset()));
-by (ALLGOALS arith_tac);
-qed "Cauchy_Bseq";
-
-(*------------------------------------------------
-  Cauchy sequence is bounded -- NSformulation
- ------------------------------------------------*)
-Goal "NSCauchy X ==> NSBseq X";
-by (asm_full_simp_tac (simpset() addsimps [Cauchy_Bseq,
-    Bseq_NSBseq_iff RS sym,NSCauchy_Cauchy_iff]) 1);
-qed "NSCauchy_NSBseq";
-
-
-(*-----------------------------------------------------------------
-          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.
- -----------------------------------------------------------------*)
-Goalw [NSconvergent_def,NSLIMSEQ_def]
-      "NSCauchy X = NSconvergent X";
-by (Step_tac 1);
-by (forward_tac [NSCauchy_NSBseq] 1);
-by (auto_tac (claset() addIs [inf_close_trans2], 
-    simpset() addsimps 
-    [NSBseq_def,NSCauchy_def]));
-by (dtac (HNatInfinite_whn RSN (2,bspec)) 1);
-by (dtac (HNatInfinite_whn RSN (2,bspec)) 1);
-by (auto_tac (claset() addSDs [st_part_Ex], simpset() 
-              addsimps [SReal_iff]));
-by (blast_tac (claset() addIs [inf_close_trans3]) 1);
-qed "NSCauchy_NSconvergent_iff";
-
-(* Standard proof for free *)
-Goal "Cauchy X = convergent X";
-by (simp_tac (simpset() addsimps [NSCauchy_Cauchy_iff RS sym,
-    convergent_NSconvergent_iff, NSCauchy_NSconvergent_iff]) 1);
-qed "Cauchy_convergent_iff";
-
-(*-----------------------------------------------------------------
-     We can now try and derive a few properties of sequences
-     starting with the limit comparison property for sequences
- -----------------------------------------------------------------*)
-Goalw [NSLIMSEQ_def]
-       "[| f ----NS> l; g ----NS> m; \
-\                  EX N. ALL n. N <= n --> f(n) <= g(n) \
-\               |] ==> l <= m";
-by (Step_tac 1);
-by (dtac starfun_le_mono 1);
-by (REPEAT(dtac (HNatInfinite_whn RSN (2,bspec)) 1));
-by (dres_inst_tac [("x","whn")] spec 1);
-by (REPEAT(dtac (bex_Infinitesimal_iff2 RS iffD2) 1));
-by Auto_tac;
-by (auto_tac (claset() addIs 
-    [hypreal_of_real_le_add_Infininitesimal_cancel2], simpset()));
-qed "NSLIMSEQ_le";
-
-(* standard version *)
-Goal "[| f ----> l; g ----> m; \
-\        EX N. ALL n. N <= n --> f(n) <= g(n) |] \
-\     ==> l <= m";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
-    NSLIMSEQ_le]) 1);
-qed "LIMSEQ_le";
-
-(*---------------
-    Also...
- --------------*)
-Goal "[| X ----> r; ALL n. a <= X n |] ==> a <= r";
-by (rtac LIMSEQ_le 1);
-by (rtac LIMSEQ_const 1);
-by (Auto_tac);
-qed "LIMSEQ_le_const";
-
-Goal "[| X ----NS> r; ALL n. a <= X n |] ==> a <= r";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
-    LIMSEQ_le_const]) 1);
-qed "NSLIMSEQ_le_const";
-
-Goal "[| X ----> r; ALL n. X n <= a |] ==> r <= a";
-by (rtac LIMSEQ_le 1);
-by (rtac LIMSEQ_const 2);
-by (Auto_tac);
-qed "LIMSEQ_le_const2";
-
-Goal "[| X ----NS> r; ALL n. X n <= a |] ==> r <= a";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
-    LIMSEQ_le_const2]) 1);
-qed "NSLIMSEQ_le_const2";
-
-(*-----------------------------------------------------
-            Shift a convergent series by 1
-  We use the fact that Cauchyness and convergence
-  are equivalent and also that the successor of an
-  infinite hypernatural is also infinite.
- -----------------------------------------------------*)
-Goal "f ----NS> l ==> (%n. f(Suc n)) ----NS> l";
-by (forward_tac [NSconvergentI RS 
-    (NSCauchy_NSconvergent_iff RS iffD2)] 1);
-by (auto_tac (claset(),simpset() addsimps [NSCauchy_def,
-    NSLIMSEQ_def,starfunNat_shift_one]));
-by (dtac bspec 1 THEN assume_tac 1);
-by (dtac bspec 1 THEN assume_tac 1);
-by (dtac (SHNat_one RSN (2,HNatInfinite_SHNat_add)) 1);
-by (blast_tac (claset() addIs [inf_close_trans3]) 1);
-qed "NSLIMSEQ_Suc";
-
-(* standard version *)
-Goal "f ----> l ==> (%n. f(Suc n)) ----> l";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
-    NSLIMSEQ_Suc]) 1);
-qed "LIMSEQ_Suc";
-
-Goal "(%n. f(Suc n)) ----NS> l ==> f ----NS> l";
-by (forward_tac [NSconvergentI RS 
-    (NSCauchy_NSconvergent_iff RS iffD2)] 1);
-by (auto_tac (claset(),simpset() addsimps [NSCauchy_def,
-    NSLIMSEQ_def,starfunNat_shift_one]));
-by (dtac bspec 1 THEN assume_tac 1);
-by (dtac bspec 1 THEN assume_tac 1);
-by (ftac (SHNat_one RSN (2,HNatInfinite_SHNat_diff)) 1);
-by (rotate_tac 2 1);
-by (auto_tac (claset() addSDs [bspec] addIs [inf_close_trans3],
-    simpset()));
-qed "NSLIMSEQ_imp_Suc";
-
-Goal "(%n. f(Suc n)) ----> l ==> f ----> l";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff]) 1);
-by (etac NSLIMSEQ_imp_Suc 1);
-qed "LIMSEQ_imp_Suc";
-
-Goal "((%n. f(Suc n)) ----> l) = (f ----> l)";
-by (blast_tac (claset() addIs [LIMSEQ_imp_Suc,LIMSEQ_Suc]) 1);
-qed "LIMSEQ_Suc_iff";
-
-Goal "((%n. f(Suc n)) ----NS> l) = (f ----NS> l)";
-by (blast_tac (claset() addIs [NSLIMSEQ_imp_Suc,NSLIMSEQ_Suc]) 1);
-qed "NSLIMSEQ_Suc_iff";
-
-(*-----------------------------------------------------
-       A sequence tends to zero iff its abs does
- ----------------------------------------------------*)
-(* we can prove this directly since proof is trivial *)
-Goalw [LIMSEQ_def] 
-      "((%n. abs(f n)) ----> #0) = (f ----> #0)";
-by (simp_tac (simpset() addsimps [abs_idempotent]) 1);
-qed "LIMSEQ_rabs_zero";
-
-(*-----------------------------------------------------*)
-(* We prove the NS version from the standard one       *)
-(* Actually pure NS proof seems more complicated       *)
-(* than the direct standard one above!                 *)
-(*-----------------------------------------------------*)
-
-Goal "((%n. abs(f n)) ----NS> #0) = (f ----NS> #0)";
-by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
-             LIMSEQ_rabs_zero]) 1);
-qed "NSLIMSEQ_rabs_zero";
-
-(*----------------------------------------
-    Also we have for a general limit 
-        (NS proof much easier)
- ---------------------------------------*)
-Goalw [NSLIMSEQ_def] 
-       "f ----NS> l ==> (%n. abs(f n)) ----NS> abs(l)";
-by (auto_tac (claset() addIs [inf_close_hrabs], simpset() 
-    addsimps [starfunNat_rabs,hypreal_of_real_hrabs RS sym]));
-qed "NSLIMSEQ_imp_rabs";
-
-(* standard version *)
-Goal "f ----> l ==> (%n. abs(f n)) ----> abs(l)";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
-    NSLIMSEQ_imp_rabs]) 1);
-qed "LIMSEQ_imp_rabs";
-
-(*-----------------------------------------------------
-       An unbounded sequence's inverse tends to 0
-  ----------------------------------------------------*)
-(* standard proof seems easier *)
-Goalw [LIMSEQ_def] 
-      "ALL y. EX N. ALL n. N <= n --> y < f(n) \
-\      ==> (%n. inverse(f n)) ----> #0";
-by (Step_tac 1 THEN Asm_full_simp_tac 1);
-by (dres_inst_tac [("x","inverse r")] spec 1 THEN Step_tac 1);
-by (res_inst_tac [("x","N")] exI 1 THEN Step_tac 1);
-by (dtac spec 1 THEN Auto_tac);
-by (forward_tac [rename_numerals real_inverse_gt_zero] 1);
-by (forward_tac [real_less_trans] 1 THEN assume_tac 1);
-by (forw_inst_tac [("x","f n")] (rename_numerals real_inverse_gt_zero) 1);
-by (asm_simp_tac (simpset() addsimps [abs_eqI2]) 1);
-by (res_inst_tac [("t","r")] (real_inverse_inverse RS subst) 1);
-by (auto_tac (claset() addIs [real_inverse_less_iff RS iffD2], 
-              simpset() delsimps [real_inverse_inverse]));
-qed "LIMSEQ_inverse_zero";
-
-Goal "ALL y. EX N. ALL n. N <= n --> y < f(n) \
-\     ==> (%n. inverse(f n)) ----NS> #0";
-by (asm_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
-                  LIMSEQ_inverse_zero]) 1);
-qed "NSLIMSEQ_inverse_zero";
-
-(*--------------------------------------------------------------
-             Sequence  1/n --> 0 as n --> infinity 
- -------------------------------------------------------------*)
-Goal "(%n. inverse(real_of_posnat n)) ----> #0";
-by (rtac LIMSEQ_inverse_zero 1 THEN Step_tac 1);
-by (cut_inst_tac [("x","y")] reals_Archimedean2 1);
-by (Step_tac 1 THEN res_inst_tac [("x","n")] exI 1);
-by (Step_tac 1 THEN etac (le_imp_less_or_eq RS disjE) 1);
-by (dtac (real_of_posnat_less_iff RS iffD2) 1);
-by (auto_tac (claset() addEs [real_less_trans], simpset()));
-qed "LIMSEQ_inverse_real_of_posnat";
-
-Goal "(%n. inverse(real_of_posnat n)) ----NS> #0";
-by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
-    LIMSEQ_inverse_real_of_posnat]) 1);
-qed "NSLIMSEQ_inverse_real_of_posnat";
-
-(*--------------------------------------------
-    Sequence  r + 1/n --> r as n --> infinity 
-    now easily proved
- --------------------------------------------*)
-Goal "(%n. r + inverse(real_of_posnat n)) ----> r";
-by (cut_facts_tac [[LIMSEQ_const,LIMSEQ_inverse_real_of_posnat]
-    MRS LIMSEQ_add] 1);
-by (Auto_tac);
-qed "LIMSEQ_inverse_real_of_posnat_add";
-
-Goal "(%n. r + inverse(real_of_posnat n)) ----NS> r";
-by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
-    LIMSEQ_inverse_real_of_posnat_add]) 1);
-qed "NSLIMSEQ_inverse_real_of_posnat_add";
-
-(*--------------
-    Also...
- --------------*)
-
-Goal "(%n. r + -inverse(real_of_posnat n)) ----> r";
-by (cut_facts_tac [[LIMSEQ_const,LIMSEQ_inverse_real_of_posnat]
-    MRS LIMSEQ_add_minus] 1);
-by (Auto_tac);
-qed "LIMSEQ_inverse_real_of_posnat_add_minus";
-
-Goal "(%n. r + -inverse(real_of_posnat n)) ----NS> r";
-by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
-    LIMSEQ_inverse_real_of_posnat_add_minus]) 1);
-qed "NSLIMSEQ_inverse_real_of_posnat_add_minus";
-
-Goal "(%n. r*( #1 + -inverse(real_of_posnat n))) ----> r";
-by (cut_inst_tac [("b","#1")] ([LIMSEQ_const,
-    LIMSEQ_inverse_real_of_posnat_add_minus] MRS LIMSEQ_mult) 1);
-by (Auto_tac);
-qed "LIMSEQ_inverse_real_of_posnat_add_minus_mult";
-
-Goal "(%n. r*( #1 + -inverse(real_of_posnat n))) ----NS> r";
-by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
-    LIMSEQ_inverse_real_of_posnat_add_minus_mult]) 1);
-qed "NSLIMSEQ_inverse_real_of_posnat_add_minus_mult";
-
-(*---------------------------------------------------------------
-                          Real Powers
- --------------------------------------------------------------*)
-Goal "(X ----NS> a) --> ((%n. (X n) ^ m) ----NS> a ^ m)";
-by (induct_tac "m" 1);
-by (auto_tac (claset() addIs [NSLIMSEQ_mult,NSLIMSEQ_const],
-    simpset()));
-qed_spec_mp "NSLIMSEQ_pow";
-
-Goal "X ----> a ==> (%n. (X n) ^ m) ----> a ^ m";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff,
-    NSLIMSEQ_pow]) 1);
-qed "LIMSEQ_pow";
-
-(*----------------------------------------------------------------
-               0 <= x < #1 ==> (x ^ n ----> 0)
-  Proof will use (NS) Cauchy equivalence for convergence and
-  also fact that bounded and monotonic sequence converges.  
- ---------------------------------------------------------------*)
-Goalw [Bseq_def] 
-      "[| #0 <= x; x < #1 |] ==> Bseq (%n. x ^ n)";
-by (res_inst_tac [("x","#1")] exI 1);
-by (auto_tac (claset() addDs [conjI RS realpow_le2] 
-    addIs [real_less_imp_le], simpset() addsimps 
-    [real_zero_less_one,abs_eqI1,realpow_abs RS sym] ));
-qed "Bseq_realpow";
-
-Goal "[| #0 <= x; x < #1 |] ==> monoseq (%n. x ^ n)";
-by (blast_tac (claset() addSIs [mono_SucI2,realpow_Suc_le3]) 1);
-qed "monoseq_realpow";
-
-Goal "[| #0 <= x; x < #1 |] ==> convergent (%n. x ^ n)";
-by (blast_tac (claset() addSIs [Bseq_monoseq_convergent,
-    Bseq_realpow,monoseq_realpow]) 1);
-qed "convergent_realpow";
-
-(* We now use NS criterion to bring proof of theorem through *)
-
-
-Goalw [NSLIMSEQ_def]
-     "[| #0 <= x; x < #1 |] ==> (%n. x ^ n) ----NS> #0";
-by (auto_tac (claset() addSDs [convergent_realpow],
-              simpset() addsimps [convergent_NSconvergent_iff]));
-by (forward_tac [NSconvergentD] 1);
-by (auto_tac (claset(),
-        simpset() addsimps [NSLIMSEQ_def, NSCauchy_NSconvergent_iff RS sym,
-                            NSCauchy_def, starfunNat_pow]));
-by (forward_tac [HNatInfinite_add_one] 1);
-by (dtac bspec 1 THEN assume_tac 1);
-by (dtac bspec 1 THEN assume_tac 1);
-by (dres_inst_tac [("x","N + 1hn")] bspec 1 THEN assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [hyperpow_add]) 1);
-by (dtac inf_close_mult_subst_SReal 1 THEN assume_tac 1);
-by (dtac inf_close_trans3 1 THEN assume_tac 1);
-by (auto_tac (claset(),
-              simpset() delsimps [hypreal_of_real_mult]
-			addsimps [hypreal_of_real_mult RS sym]));
-qed "NSLIMSEQ_realpow_zero";
-
-(*---------------  standard version ---------------*)
-Goal "[| #0 <= x; x < #1 |] ==> (%n. x ^ n) ----> #0";
-by (asm_simp_tac (simpset() addsimps [NSLIMSEQ_realpow_zero,
-                                      LIMSEQ_NSLIMSEQ_iff]) 1);
-qed "LIMSEQ_realpow_zero";
-
-Goal "#1 < x ==> (%n. a / (x ^ n)) ----> #0";
-by (cut_inst_tac [("a","a"),("x1","inverse x")] 
-    ([LIMSEQ_const, LIMSEQ_realpow_zero] MRS LIMSEQ_mult) 1);
-by (auto_tac (claset(), 
-              simpset() addsimps [real_divide_def, realpow_inverse])); 
-by (asm_simp_tac (simpset() addsimps [real_inverse_eq_divide,
-                                      pos_real_divide_less_eq]) 1); 
-qed "LIMSEQ_divide_realpow_zero";
-
-(*----------------------------------------------------------------
-               Limit of c^n for |c| < 1  
- ---------------------------------------------------------------*)
-Goal "abs(c) < #1 ==> (%n. abs(c) ^ n) ----> #0";
-by (blast_tac (claset() addSIs [LIMSEQ_realpow_zero,abs_ge_zero]) 1);
-qed "LIMSEQ_rabs_realpow_zero";
-
-Goal "abs(c) < #1 ==> (%n. abs(c) ^ n) ----NS> #0";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_rabs_realpow_zero,
-    LIMSEQ_NSLIMSEQ_iff RS sym]) 1);
-qed "NSLIMSEQ_rabs_realpow_zero";
-
-Goal "abs(c) < #1 ==> (%n. c ^ n) ----> #0";
-by (rtac (LIMSEQ_rabs_zero RS iffD1) 1);
-by (auto_tac (claset() addIs [LIMSEQ_rabs_realpow_zero],
-              simpset() addsimps [realpow_abs RS sym]));
-qed "LIMSEQ_rabs_realpow_zero2";
-
-Goal "abs(c) < #1 ==> (%n. c ^ n) ----NS> #0";
-by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_rabs_realpow_zero2,
-    LIMSEQ_NSLIMSEQ_iff RS sym]) 1);
-qed "NSLIMSEQ_rabs_realpow_zero2";
-
-(***---------------------------------------------------------------
-                 Hyperreals and Sequences
- ---------------------------------------------------------------***)
-(*** A bounded sequence is a finite hyperreal ***)
-Goal "NSBseq X ==> Abs_hypreal(hyprel^^{X}) : HFinite";
-by (auto_tac (claset() addSIs [bexI,lemma_hyprel_refl] addIs 
-       [FreeUltrafilterNat_all RS FreeUltrafilterNat_subset],
-       simpset() addsimps [HFinite_FreeUltrafilterNat_iff,
-        Bseq_NSBseq_iff RS sym, Bseq_iff1a]));
-qed "NSBseq_HFinite_hypreal";
-
-(*** A sequence converging to zero defines an infinitesimal ***)
-Goalw [NSLIMSEQ_def] 
-      "X ----NS> #0 ==> Abs_hypreal(hyprel^^{X}) : Infinitesimal";
-by (dres_inst_tac [("x","whn")] bspec 1);
-by (simp_tac (simpset() addsimps [HNatInfinite_whn]) 1);
-by (auto_tac (claset(),simpset() addsimps [hypnat_omega_def,
-    mem_infmal_iff RS sym,starfunNat,hypreal_of_real_zero]));
-qed "NSLIMSEQ_zero_Infinitesimal_hypreal";
-
-(***---------------------------------------------------------------
-    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 "EX 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 <= f(n)";
-Goal "subseq f ==> EX n. N1 <= n & N2 <= f(n)";
- ---------------------------------------------------------------***)
--- a/src/HOL/Real/Hyperreal/SEQ.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,63 +0,0 @@
-(*  Title       : SEQ.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Convergence of sequences and series
-*) 
-
-SEQ = NatStar + HyperPow + 
-
-constdefs
-
-  (* Standard definition of convergence of sequence *)           
-  LIMSEQ :: [nat=>real,real] => bool    ("((_)/ ----> (_))" [60, 60] 60)
-  "X ----> L == (ALL r. #0 < r --> (EX no. ALL n. no <= n --> abs (X n + -L) < r))"
-  
-  (* Nonstandard definition of convergence of sequence *)
-  NSLIMSEQ :: [nat=>real,real] => bool    ("((_)/ ----NS> (_))" [60, 60] 60)
-  "X ----NS> L == (ALL N: HNatInfinite. (*fNat* X) N @= hypreal_of_real L)"
-
-  (* define value of limit using choice operator*)
-  lim :: (nat => real) => real
-  "lim X == (@L. (X ----> L))"
-
-  nslim :: (nat => real) => real
-  "nslim X == (@L. (X ----NS> L))"
-
-  (* Standard definition of convergence *)
-  convergent :: (nat => real) => bool
-  "convergent X == (EX L. (X ----> L))"
-
-  (* Nonstandard definition of convergence *)
-  NSconvergent :: (nat => real) => bool
-  "NSconvergent X == (EX L. (X ----NS> L))"
-
-  (* Standard definition for bounded sequence *)
-  Bseq :: (nat => real) => bool
-  "Bseq X == (EX K. (#0 < K & (ALL n. abs(X n) <= K)))"
- 
-  (* Nonstandard definition for bounded sequence *)
-  NSBseq :: (nat=>real) => bool
-  "NSBseq X == (ALL N: HNatInfinite. (*fNat* X) N : HFinite)" 
-
-  (* Definition for monotonicity *)
-  monoseq :: (nat=>real)=>bool
-  "monoseq X == ((ALL (m::nat) n. m <= n --> X m <= X n) |
-                 (ALL m n. m <= n --> X n <= X m))"   
-
-  (* Definition of subsequence *)
-  subseq :: (nat => nat) => bool
-  "subseq f == (ALL m n. m < n --> (f m) < (f n))"
-
-  (** Cauchy condition **)
-
-  (* Standard definition *)
-  Cauchy :: (nat => real) => bool
-  "Cauchy X == (ALL e. (#0 < e -->
-                       (EX M. (ALL m n. M <= m & M <= n 
-                             --> abs((X m) + -(X n)) < e))))"
-
-  NSCauchy :: (nat => real) => bool
-  "NSCauchy X == (ALL M: HNatInfinite. ALL N: HNatInfinite.
-                      (*fNat* X) M @= (*fNat* X) N)"
-end
-
--- a/src/HOL/Real/Hyperreal/Series.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,609 +0,0 @@
-(*  Title       : Series.ML
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-                : 1999  University of Edinburgh
-    Description : Finite summation and infinite series
-*) 
-
-Goal "sumr (Suc n) n f = #0";
-by (induct_tac "n" 1);
-by (Auto_tac);
-qed "sumr_Suc_zero";
-Addsimps [sumr_Suc_zero];
-
-Goal "sumr m m f = #0";
-by (induct_tac "m" 1);
-by (Auto_tac);
-qed "sumr_eq_bounds";
-Addsimps [sumr_eq_bounds];
-
-Goal "sumr m (Suc m) f = f(m)";
-by (Auto_tac);
-qed "sumr_Suc_eq";
-Addsimps [sumr_Suc_eq];
-
-Goal "sumr (m+k) k f = #0";
-by (induct_tac "k" 1);
-by (Auto_tac);
-qed "sumr_add_lbound_zero";
-Addsimps [sumr_add_lbound_zero];
-
-Goal "sumr m n f + sumr m n g = sumr m n (%n. f n + g n)";
-by (induct_tac "n" 1);
-by (auto_tac (claset(),simpset() addsimps real_add_ac));
-qed "sumr_add";
-
-Goal "r * sumr m n f = sumr m n (%n. r * f n)";
-by (induct_tac "n" 1);
-by (Auto_tac);
-by (auto_tac (claset(),simpset() addsimps 
-    [real_add_mult_distrib2]));
-qed "sumr_mult";
-
-Goal "n < p --> sumr 0 n f + sumr n p f = sumr 0 p f";
-by (induct_tac "p" 1);
-by (auto_tac (claset() addSDs [CLAIM "n < Suc na ==> n <= na",
-    leI] addDs [le_anti_sym], simpset()));
-qed_spec_mp "sumr_split_add";
-
-Goal "n < p ==> sumr 0 p f + \
-\                - sumr 0 n f = sumr n p f";
-by (dres_inst_tac [("f1","f")] (sumr_split_add RS sym) 1);
-by (asm_simp_tac (simpset() addsimps real_add_ac) 1);
-qed "sumr_split_add_minus";
-
-Goal "abs(sumr m n f) <= sumr m n (%i. abs(f i))";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [(abs_triangle_ineq 
-    RS real_le_trans),real_add_le_mono1], simpset()));
-qed "sumr_rabs";
-
-Goal "!!f g. (ALL r. m <= r & r < n --> f r = g r) \
-\                --> sumr m n f = sumr m n g";
-by (induct_tac "n" 1);
-by (Auto_tac);
-qed_spec_mp "sumr_fun_eq";
-
-Goal "sumr 0 n (%i. r) = real_of_nat n*r";
-by (induct_tac "n" 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [real_add_mult_distrib,real_of_nat_zero]));
-qed "sumr_const";
-Addsimps [sumr_const];
-
-Goal "sumr 0 n f + -(real_of_nat n*r) = sumr 0 n (%i. f i + -r)";
-by (full_simp_tac (simpset() addsimps [sumr_add RS sym,
-    real_minus_mult_eq2] delsimps [real_minus_mult_eq2 RS sym]) 1);
-qed "sumr_add_mult_const";
-
-Goalw [real_diff_def] 
-     "sumr 0 n f - (real_of_nat n*r) = sumr 0 n (%i. f i - r)";
-by (full_simp_tac (simpset() addsimps [sumr_add_mult_const]) 1);
-qed "sumr_diff_mult_const";
-
-Goal "n < m --> sumr m n f = #0";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addDs [less_imp_le], simpset()));
-qed_spec_mp "sumr_less_bounds_zero";
-Addsimps [sumr_less_bounds_zero];
-
-Goal "sumr m n (%i. - f i) = - sumr m n f";
-by (induct_tac "n" 1);
-by (case_tac "Suc n <= m" 2);
-by (auto_tac (claset(),simpset() addsimps 
-    [real_minus_add_distrib]));
-qed "sumr_minus";
-
-Goal "sumr (m+k) (n+k) f = sumr m n (%i. f(i + k))";
-by (induct_tac "n" 1);
-by (Auto_tac);
-qed "sumr_shift_bounds";
-
-Goal "sumr 0 (#2*n) (%i. (#-1) ^ Suc i) = #0";
-by (induct_tac "n" 1);
-by (Auto_tac);
-qed "sumr_minus_one_realpow_zero";
-Addsimps [sumr_minus_one_realpow_zero];
-
-Goal "m < Suc n ==> Suc n - m = Suc (n - m)";
-by (dtac less_imp_Suc_add 1);
-by (Auto_tac);
-val Suc_diff_n = result();
-
-Goal "(ALL n. m <= Suc n --> f n = r) & m <= na \
-\                --> sumr m na f = (real_of_nat (na - m) * r)";
-by (induct_tac "na" 1);
-by (Step_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [le_Suc_eq])  2);
-by (Step_tac 1);
-by (dres_inst_tac [("x","n")] spec 3);
-by (auto_tac (claset() addSDs [le_imp_less_or_eq],
-    simpset()));
-by (asm_simp_tac (simpset() addsimps [real_add_mult_distrib,
-    Suc_diff_n]) 1);
-qed_spec_mp "sumr_interval_const";
-
-Goal "(ALL n. m <= n --> f n = r) & m <= na \
-\     --> sumr m na f = (real_of_nat (na - m) * r)";
-by (induct_tac "na" 1);
-by (Step_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [le_Suc_eq])  2);
-by (Step_tac 1);
-by (dres_inst_tac [("x","n")] spec 3);
-by (auto_tac (claset() addSDs [le_imp_less_or_eq],
-    simpset()));
-by (asm_simp_tac (simpset() addsimps [Suc_diff_n,
-    real_add_mult_distrib]) 1);
-qed_spec_mp "sumr_interval_const2";
-
-Goal "(m <= n)= (m < Suc n)";
-by (Auto_tac);
-qed "le_eq_less_Suc";
-
-Goal "(ALL n. m <= n --> #0 <= f n) & m < na \
-\                --> sumr 0 m f <= sumr 0 na f";
-by (induct_tac "na" 1);
-by (Step_tac 1);
-by (ALLGOALS(asm_full_simp_tac (simpset() addsimps 
-    [le_eq_less_Suc RS sym])));
-by (ALLGOALS(dres_inst_tac [("x","n")] spec));
-by (Step_tac 1);
-by (dtac le_imp_less_or_eq 1 THEN Step_tac 1);
-by (dtac real_add_le_mono 2);
-by (dres_inst_tac [("i","sumr 0 m f")] 
-    (real_le_refl RS real_add_le_mono) 1);
-by (Auto_tac);
-qed_spec_mp "sumr_le";
-
-Goal "!!f g. (ALL r. m <= r & r < n --> f r <= g r) \
-\                --> sumr m n f <= sumr m n g";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [real_add_le_mono],
-    simpset() addsimps [le_def]));
-qed_spec_mp "sumr_le2";
-
-Goal "(ALL n. #0 <= f n) --> #0 <= sumr m n f";
-by (induct_tac "n" 1);
-by Auto_tac;
-by (dres_inst_tac [("x","n")] spec 1);
-by (arith_tac 1);
-qed_spec_mp "sumr_ge_zero";
-
-Goal "(ALL n. m <= n --> #0 <= f n) --> #0 <= sumr m n f";
-by (induct_tac "n" 1);
-by Auto_tac;
-by (dres_inst_tac [("x","n")] spec 1);
-by (arith_tac 1);
-qed_spec_mp "sumr_ge_zero2";
-
-Goal "#0 <= sumr m n (%n. abs (f n))";
-by (induct_tac "n" 1);
-by Auto_tac;
-by (arith_tac 1);
-qed "sumr_rabs_ge_zero";
-Addsimps [sumr_rabs_ge_zero];
-AddSIs [sumr_rabs_ge_zero]; 
-
-Goal "abs (sumr m n (%n. abs (f n))) = (sumr m n (%n. abs (f n)))";
-by (induct_tac "n" 1);
-by (Auto_tac THEN arith_tac 1);
-qed "rabs_sumr_rabs_cancel";
-Addsimps [rabs_sumr_rabs_cancel];
-
-Goal "ALL n. N <= n --> f n = #0 \
-\     ==> ALL m n. N <= m --> sumr m n f = #0";   
-by (Step_tac 1);
-by (induct_tac "n" 1);
-by (Auto_tac);
-qed "sumr_zero";
-
-Goal "Suc N <= n --> N <= n - 1";
-by (induct_tac "n" 1);
-by (Auto_tac);
-qed_spec_mp "Suc_le_imp_diff_ge";
-
-Goal "ALL n. N <= n --> f (Suc n) = #0 \
-\     ==> ALL m n. Suc N <= m --> sumr m n f = #0";   
-by (rtac sumr_zero 1 THEN Step_tac 1);
-by (subgoal_tac "0 < n" 1);
-by (dtac Suc_le_imp_diff_ge 1);
-by (Auto_tac);
-qed "Suc_le_imp_diff_ge2";
-
-(* proved elsewhere? *)
-Goal "(0 < n) = (EX m. n = Suc m)";
-by (induct_tac "n" 1);
-by (Auto_tac);
-qed "gt_zero_eq_Ex";
-
-Goal "sumr 1 n (%n. f(n) * #0 ^ n) = #0";
-by (induct_tac "n" 1);
-by (auto_tac (claset(),simpset() addsimps [gt_zero_eq_Ex]));
-qed "sumr_one_lb_realpow_zero";
-Addsimps [sumr_one_lb_realpow_zero];
-
-Goalw [real_diff_def] "sumr m n f - sumr m n g = sumr m n (%n. f n - g n)";
-by (simp_tac (simpset() addsimps [sumr_add RS sym,sumr_minus]) 1);
-qed "sumr_diff";
-
-Goal "(ALL p. (m <= p & p < m + n --> (f p = g p))) --> sumr m n f = sumr m n g";
-by (induct_tac "n" 1);
-by (Auto_tac);
-qed_spec_mp "sumr_subst";
-
-Goal "(ALL p. m <= p & p < m + n --> (f(p) <= K)) \
-\     --> (sumr m (m + n) f <= (real_of_nat n * K))";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [real_add_le_mono],
-    simpset() addsimps [real_add_mult_distrib]));
-qed_spec_mp "sumr_bound";
-
-Goal "(ALL p. 0 <= p & p < n --> (f(p) <= K)) \
-\     --> (sumr 0 n f <= (real_of_nat n * K))";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [real_add_le_mono],
-    simpset() addsimps [real_add_mult_distrib]));
-qed_spec_mp "sumr_bound2";
-
-Goal "sumr 0 n (%m. sumr (m * k) (m*k + k) f) = sumr 0 (n * k) f";
-by (subgoal_tac "k = 0 | 0 < k" 1);
-by (Auto_tac);
-by (induct_tac "n" 1);
-by (auto_tac (claset(),simpset() addsimps [sumr_split_add,add_commute]));
-qed "sumr_group";
-Addsimps [sumr_group];
-
-(*-------------------------------------------------------------------
-                        Infinite sums
-    Theorems follow from properties of limits and sums
- -------------------------------------------------------------------*)
-(*----------------------
-   suminf is the sum   
- ---------------------*)
-Goalw [sums_def,summable_def] 
-      "f sums l ==> summable f";
-by (Blast_tac 1);
-qed "sums_summable";
-
-Goalw [summable_def,suminf_def]
-     "summable f ==> f sums (suminf f)";
-by (blast_tac (claset() addIs [someI2]) 1);
-qed "summable_sums";
-
-Goalw [summable_def,suminf_def,sums_def]
-     "summable f ==> (%n. sumr 0 n f) ----> (suminf f)";
-by (blast_tac (claset() addIs [someI2]) 1);
-qed "summable_sumr_LIMSEQ_suminf";
-
-(*-------------------
-    sum is unique                    
- ------------------*)
-Goal "f sums s ==> (s = suminf f)";
-by (forward_tac [sums_summable RS summable_sums] 1);
-by (auto_tac (claset() addSIs [LIMSEQ_unique],
-    simpset() addsimps [sums_def]));
-qed "sums_unique";
-
-(*
-Goalw [sums_def,LIMSEQ_def] 
-     "(ALL m. n <= Suc m --> f(m) = #0) ==> f sums (sumr 0 n f)";
-by (Step_tac 1);
-by (res_inst_tac [("x","n")] exI 1);
-by (Step_tac 1 THEN forward_tac [le_imp_less_or_eq] 1);
-by (Step_tac 1);
-by (dres_inst_tac [("f","f")] sumr_split_add_minus 1);
-by (ALLGOALS (Asm_simp_tac));
-by (dtac (conjI RS sumr_interval_const) 1);
-by (Auto_tac);
-qed "series_zero";
-next one was called series_zero2
-**********************)
-
-Goalw [sums_def,LIMSEQ_def] 
-     "(ALL m. n <= m --> f(m) = #0) ==> f sums (sumr 0 n f)";
-by (Step_tac 1);
-by (res_inst_tac [("x","n")] exI 1);
-by (Step_tac 1 THEN forward_tac [le_imp_less_or_eq] 1);
-by (Step_tac 1);
-by (dres_inst_tac [("f","f")] sumr_split_add_minus 1);
-by (ALLGOALS (Asm_simp_tac));
-by (dtac (conjI RS sumr_interval_const2) 1);
-by (Auto_tac);
-qed "series_zero";
-
-Goalw [sums_def] "x sums x0 ==> (%n. c * x(n)) sums (c * x0)";
-by (auto_tac (claset() addSIs [LIMSEQ_mult] addIs [LIMSEQ_const], 
-              simpset() addsimps [sumr_mult RS sym]));
-qed "sums_mult";
-
-Goal "x sums x' ==> (%n. x(n)/c) sums (x'/c)";
-by (asm_simp_tac (simpset() addsimps [real_divide_def, sums_mult,
-                              inst "w" "inverse c" real_mult_commute]) 1); 
-qed "sums_divide";
-
-Goalw [sums_def] "[| x sums x0; y sums y0 |] ==> (%n. x n - y n) sums (x0-y0)";
-by (auto_tac (claset() addIs [LIMSEQ_diff],
-              simpset() addsimps [sumr_diff RS sym]));
-qed "sums_diff";
-
-Goal "summable f ==> suminf f * c = suminf(%n. f n * c)";
-by (auto_tac (claset() addSIs [sums_unique,sums_mult,summable_sums],
-    simpset() addsimps [real_mult_commute]));
-qed "suminf_mult";
-
-Goal "summable f ==> c * suminf f  = suminf(%n. c * f n)";
-by (auto_tac (claset() addSIs [sums_unique,sums_mult,summable_sums],
-    simpset()));
-qed "suminf_mult2";
-
-Goal "[| summable f; summable g |]  \
-\     ==> suminf f - suminf g  = suminf(%n. f n - g n)";
-by (auto_tac (claset() addSIs [sums_diff,sums_unique,summable_sums], simpset()));
-qed "suminf_diff";
-
-Goalw [sums_def] 
-       "x sums x0 ==> (%n. - x n) sums - x0";
-by (auto_tac (claset() addSIs [LIMSEQ_minus], simpset() addsimps [sumr_minus]));
-qed "sums_minus";
-
-Goal "[|summable f; 0 < k |] \
-\     ==> (%n. sumr (n*k) (n*k + k) f) sums (suminf f)";
-by (dtac summable_sums 1);
-by (auto_tac (claset(),simpset() addsimps [sums_def,LIMSEQ_def]));
-by (dres_inst_tac [("x","r")] spec 1 THEN Step_tac 1); 
-by (res_inst_tac [("x","no")] exI 1 THEN Step_tac 1); 
-by (dres_inst_tac [("x","n*k")] spec 1); 
-by (auto_tac (claset() addSDs [not_leE], simpset()));
-by (dres_inst_tac [("j","no")] less_le_trans 1);
-by (Auto_tac);
-qed "sums_group";
-
-Goal "[|summable f; ALL d. #0 < (f(n + (2 * d))) + f(n + ((2 * d) + 1))|] \
-\     ==> sumr 0 n f < suminf f";
-by (dtac summable_sums 1);
-by (auto_tac (claset(),simpset() addsimps [sums_def,LIMSEQ_def]));
-by (dres_inst_tac [("x","f(n) + f(n + 1)")] spec 1);
-by (Auto_tac);
-by (rtac ccontr 2 THEN dtac real_leI 2);
-by (subgoal_tac "sumr 0 (n + 2) f <= sumr 0 (2 * (Suc no) + n) f" 2);
-by (induct_tac "no" 3 THEN Simp_tac 3);
-by (res_inst_tac [("j","sumr 0 (2*(Suc na)+n) f")] real_le_trans 3);
-by (assume_tac 3);
-by (dres_inst_tac [("x","Suc na")] spec 3);
-by (dres_inst_tac [("x","0")] spec 1);
-by (Asm_full_simp_tac 1);
-by (asm_full_simp_tac (simpset() addsimps add_ac) 2);
-by (rotate_tac 1 1 THEN dres_inst_tac [("x","2 * (Suc no) + n")] spec 1);
-by (Step_tac 1 THEN Asm_full_simp_tac 1);
-by (subgoal_tac "suminf f + (f(n) + f(n + 1)) <= \
-\                 sumr 0 (2 * (Suc no) + n) f" 1);
-by (res_inst_tac [("j","sumr 0 (n+2) f")] real_le_trans 2);
-by (assume_tac 3);
-by (res_inst_tac [("j","sumr 0 n f + (f(n) + f(n + 1))")] real_le_trans 2);
-by (REPEAT(Asm_simp_tac 2));
-by (subgoal_tac "suminf f <= sumr 0 (2 * (Suc no) + n) f" 1);
-by (res_inst_tac [("j","suminf f + (f(n) + f(n + 1))")] real_le_trans 2);
-by (assume_tac 3);
-by (dres_inst_tac [("x","0")] spec 2);
-by (Asm_full_simp_tac 2);
-by (subgoal_tac "#0 <= sumr 0 (2 * Suc no + n) f + - suminf f" 1);
-by (dtac (rename_numerals abs_eqI1) 1 );
-by (Asm_full_simp_tac 1);
-by (auto_tac (claset(),simpset() addsimps [real_le_def]));
-qed "sumr_pos_lt_pair";
-
-(*-----------------------------------------------------------------
-   Summable series of positive terms has limit >= any partial sum 
- ----------------------------------------------------------------*)
-Goal 
-     "[| summable f; ALL m. n <= m --> #0 <= f(m) |] \
-\          ==> sumr 0 n f <= suminf f";
-by (dtac summable_sums 1);
-by (rewtac sums_def);
-by (cut_inst_tac [("k","sumr 0 n f")] LIMSEQ_const 1);
-by (etac LIMSEQ_le 1 THEN Step_tac 1);
-by (res_inst_tac [("x","n")] exI 1 THEN Step_tac 1);
-by (dtac le_imp_less_or_eq 1 THEN Step_tac 1);
-by (auto_tac (claset() addIs [sumr_le], simpset()));
-qed "series_pos_le";
-
-Goal "[| summable f; ALL m. n <= m --> #0 < f(m) |] \
-\          ==> sumr 0 n f < suminf f";
-by (res_inst_tac [("j","sumr 0 (Suc n) f")] real_less_le_trans 1);
-by (rtac series_pos_le 2);
-by (Auto_tac);
-by (dres_inst_tac [("x","m")] spec 1);
-by (Auto_tac);
-qed "series_pos_less";
-
-(*-------------------------------------------------------------------
-                    sum of geometric progression                 
- -------------------------------------------------------------------*)
-
-Goal "x ~= #1 ==> sumr 0 n (%n. x ^ n) = (x ^ n - #1) / (x - #1)";
-by (induct_tac "n" 1);
-by (Auto_tac);
-by (res_inst_tac [("c1","x - #1")] (real_mult_right_cancel RS iffD1) 1);
-by (auto_tac (claset(),
-       simpset() addsimps [real_mult_assoc, real_add_mult_distrib]));
-by (auto_tac (claset(),
-       simpset() addsimps [real_add_mult_distrib2,
-                           real_diff_def, real_mult_commute]));
-qed "sumr_geometric";
-
-Goal "abs(x) < #1 ==> (%n. x ^ n) sums (#1/(#1 - x))";
-by (case_tac "x = #1" 1);
-by (auto_tac (claset() addSDs [LIMSEQ_rabs_realpow_zero2],
-             simpset() addsimps [sumr_geometric ,abs_one, sums_def,
-                                 real_diff_def, real_add_divide_distrib]));
-by (subgoal_tac "#1 / (#1 + - x) = #0/(x-#1) + - #1/(x-#1)" 1);
-by (asm_full_simp_tac (simpset() addsimps [real_divide_eq_cancel1,
-                 real_divide_minus_eq RS sym, real_diff_def]) 2); 
-by (etac ssubst 1); 
-by (rtac LIMSEQ_add 1 THEN rtac LIMSEQ_divide 1);
-by (auto_tac (claset() addIs [LIMSEQ_const], 
-              simpset() addsimps [real_diff_def, LIMSEQ_rabs_realpow_zero2]));
-qed "geometric_sums";
-
-(*-------------------------------------------------------------------
-    Cauchy-type criterion for convergence of series (c.f. Harrison)
- -------------------------------------------------------------------*)
-Goalw [summable_def,sums_def,convergent_def]
-      "summable f = convergent (%n. sumr 0 n f)";
-by (Auto_tac);
-qed "summable_convergent_sumr_iff";
-
-Goal "summable f = \
-\     (ALL e. #0 < e --> (EX N. ALL m n. N <= m --> abs(sumr m n f) < e))";
-by (auto_tac (claset(),simpset() addsimps [summable_convergent_sumr_iff,
-    Cauchy_convergent_iff RS sym,Cauchy_def]));
-by (ALLGOALS(dtac spec) THEN Auto_tac);
-by (res_inst_tac [("x","M")] exI 1);
-by (res_inst_tac [("x","N")] exI 2);
-by (Auto_tac);
-by (ALLGOALS(cut_inst_tac [("m","m"),("n","n")] less_linear));
-by (Auto_tac);
-by (ftac (le_less_trans RS less_imp_le) 1 THEN assume_tac 1);
-by (dres_inst_tac [("x","n")] spec 1);
-by (dres_inst_tac [("x","m")] spec 1);
-by (auto_tac (claset() addIs [(abs_minus_add_cancel RS subst)],
-    simpset() addsimps [sumr_split_add_minus,abs_minus_add_cancel]));
-qed "summable_Cauchy";
-
-(*-------------------------------------------------------------------
-     Terms of a convergent series tend to zero
-     > Goalw [LIMSEQ_def] "summable f ==> f ----> #0";
-     Proved easily in HSeries after proving nonstandard case.
- -------------------------------------------------------------------*)
-(*-------------------------------------------------------------------
-                    Comparison test
- -------------------------------------------------------------------*)
-Goal "[| EX N. ALL n. N <= n --> abs(f n) <= g n; \
-\        summable g \
-\     |] ==> summable f";
-by (auto_tac (claset(),simpset() addsimps [summable_Cauchy]));
-by (dtac spec 1 THEN Auto_tac);
-by (res_inst_tac [("x","N + Na")] exI 1 THEN Auto_tac);
-by (rotate_tac 2 1);
-by (dres_inst_tac [("x","m")] spec 1);
-by (Auto_tac THEN rotate_tac 2 1 THEN dres_inst_tac [("x","n")] spec 1);
-by (res_inst_tac [("j","sumr m n (%k. abs(f k))")] real_le_less_trans 1);
-by (rtac sumr_rabs 1);
-by (res_inst_tac [("j","sumr m n g")] real_le_less_trans 1);
-by (auto_tac (claset() addIs [sumr_le2],
-    simpset() addsimps [abs_interval_iff]));
-qed "summable_comparison_test";
-
-Goal "[| EX N. ALL n. N <= n --> abs(f n) <= g n; \
-\        summable g \
-\     |] ==> summable (%k. abs (f k))";
-by (rtac summable_comparison_test 1);
-by (auto_tac (claset(),simpset() addsimps [abs_idempotent]));
-qed "summable_rabs_comparison_test";
-
-(*------------------------------------------------------------------*)
-(*       Limit comparison property for series (c.f. jrh)            *)
-(*------------------------------------------------------------------*)
-Goal "[|ALL n. f n <= g n; summable f; summable g |] \
-\     ==> suminf f <= suminf g";
-by (REPEAT(dtac summable_sums 1));
-by (auto_tac (claset() addSIs [LIMSEQ_le], simpset() addsimps [sums_def]));
-by (rtac exI 1);
-by (auto_tac (claset() addSIs [sumr_le2], simpset()));
-qed "summable_le";
-
-Goal "[|ALL n. abs(f n) <= g n; summable g |] \
-\     ==> summable f & suminf f <= suminf g";
-by (auto_tac (claset() addIs [summable_comparison_test]
-    addSIs [summable_le], simpset()));
-by (auto_tac (claset(),simpset() addsimps [abs_le_interval_iff]));
-qed "summable_le2";
-
-(*-------------------------------------------------------------------
-         Absolute convergence imples normal convergence
- -------------------------------------------------------------------*)
-Goal "summable (%n. abs (f n)) ==> summable f";
-by (auto_tac (claset(),simpset() addsimps [sumr_rabs,summable_Cauchy]));
-by (dtac spec 1 THEN Auto_tac);
-by (res_inst_tac [("x","N")] exI 1 THEN Auto_tac);
-by (dtac spec 1 THEN Auto_tac);
-by (res_inst_tac [("j","sumr m n (%n. abs(f n))")] real_le_less_trans 1);
-by (auto_tac (claset() addIs [sumr_rabs], simpset()));
-qed "summable_rabs_cancel";
-
-(*-------------------------------------------------------------------
-         Absolute convergence of series
- -------------------------------------------------------------------*)
-Goal "summable (%n. abs (f n)) ==> abs(suminf f) <= suminf (%n. abs(f n))";
-by (auto_tac (claset() addIs [LIMSEQ_le,LIMSEQ_imp_rabs,summable_rabs_cancel,
-    summable_sumr_LIMSEQ_suminf,sumr_rabs], simpset()));
-qed "summable_rabs";
-
-(*-------------------------------------------------------------------
-                        The ratio test
- -------------------------------------------------------------------*)
-Goal "[| c <= #0; abs x <= c * abs y |] ==> x = (#0::real)";
-by (dtac real_le_imp_less_or_eq 1);
-by (Auto_tac);
-by (dres_inst_tac [("x1","y")] (abs_ge_zero RS 
-    rename_numerals real_mult_le_zero) 1);
-by (asm_full_simp_tac (simpset() addsimps [real_mult_commute]) 1);
-by (dtac real_le_trans 1 THEN assume_tac 1);
-by (TRYALL(arith_tac));
-qed "rabs_ratiotest_lemma";
-
-(* lemmas *)
-
-Goal "(k::nat) <= l ==> (EX n. l = k + n)";
-by (dtac le_imp_less_or_eq 1);
-by (auto_tac (claset() addDs [less_imp_Suc_add], simpset()));
-qed "le_Suc_ex";
-
-Goal "((k::nat) <= l) = (EX n. l = k + n)";
-by (auto_tac (claset(),simpset() addsimps [le_Suc_ex]));
-qed "le_Suc_ex_iff";
-
-(*All this trouble just to get #0<c *)
-Goal "[| ALL n. N <= n --> abs(f(Suc n)) <= c*abs(f n) |] \
-\     ==> #0 < c | summable f";
-by (simp_tac (simpset() addsimps [linorder_not_le RS sym]) 1); 
-by (asm_full_simp_tac (simpset() addsimps [summable_Cauchy]) 1);
-by (Step_tac 1 THEN subgoal_tac "ALL n. N <= n --> f (Suc n) = #0" 1);
-by (blast_tac (claset() addIs [rabs_ratiotest_lemma]) 2);
-by (res_inst_tac [("x","Suc N")] exI 1 THEN Step_tac 1);
-by (dtac Suc_le_imp_diff_ge2 1 THEN Auto_tac);
-qed "ratio_test_lemma2";
-
-
-Goal "[| c < #1; ALL n. N <= n --> abs(f(Suc n)) <= c*abs(f n) |] \
-\     ==> summable f";
-by (forward_tac [ratio_test_lemma2] 1);
-by Auto_tac;  
-by (res_inst_tac [("g","%n. (abs (f N)/(c ^ N))*c ^ n")] 
-    summable_comparison_test 1);
-by (res_inst_tac [("x","N")] exI 1 THEN Step_tac 1);
-by (dtac (le_Suc_ex_iff RS iffD1) 1);
-by (auto_tac (claset(),simpset() addsimps [realpow_add,
-    rename_numerals realpow_not_zero]));
-by (induct_tac "na" 1 THEN Auto_tac);
-by (res_inst_tac [("j","c*abs(f(N + n))")] real_le_trans 1);
-by (auto_tac (claset() addIs [real_mult_le_le_mono1],
-    simpset() addsimps [summable_def]));
-by (asm_full_simp_tac (simpset() addsimps real_mult_ac) 1);
-by (res_inst_tac [("x","abs(f N) * (#1/(#1 - c)) / (c ^ N)")] exI 1);
-by (rtac sums_divide 1); 
-by (rtac sums_mult 1); 
-by (auto_tac (claset() addSIs [sums_mult,geometric_sums], 
-              simpset() addsimps [real_abs_def]));
-qed "ratio_test";
-
-
-(*----------------------------------------------------------------------------*)
-(* Differentiation of finite sum                                              *)
-(*----------------------------------------------------------------------------*)
-
-Goal "(ALL r. m <= r & r < (m + n) --> DERIV (%x. f r x) x :> (f' r x)) \
-\     --> DERIV (%x. sumr m n (%n. f n x)) x :> sumr m n (%r. f' r x)";
-by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [DERIV_add], simpset()));
-qed "DERIV_sumr";
--- a/src/HOL/Real/Hyperreal/Series.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,27 +0,0 @@
-(*  Title       : Series.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Finite summation and infinite series
-*) 
-
-
-Series = SEQ + Lim +
-
-consts sumr :: "[nat,nat,(nat=>real)] => real"
-primrec
-   sumr_0   "sumr m 0 f = #0"
-   sumr_Suc "sumr m (Suc n) f = (if n < m then #0 
-                               else sumr m n f + f(n))"
-
-constdefs
-   sums  :: [nat=>real,real] => bool     (infixr 80)
-   "f sums s  == (%n. sumr 0 n f) ----> s"
-
-   summable :: (nat=>real) => bool
-   "summable f == (EX s. f sums s)"
-
-   suminf   :: (nat=>real) => real
-   "suminf f == (@s. f sums s)"
-end
-
-
--- a/src/HOL/Real/Hyperreal/Star.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,490 +0,0 @@
-(*  Title       : STAR.ML
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : *-transforms
-*) 
-
-(*--------------------------------------------------------
-   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) 
-*)
-Goal "ALL x. EX y. Q x y ==> EX (f :: nat => nat). ALL x. Q x (f x)";
-by (res_inst_tac [("x","%x. LEAST y. Q x y")] exI 1);
-by (blast_tac (claset() addIs [LeastI]) 1);
-qed "no_choice";
-
-(*------------------------------------------------------------
-    Properties of the *-transform applied to sets of reals
- ------------------------------------------------------------*)
-
-Goalw [starset_def] "*s*(UNIV::real set) = (UNIV::hypreal set)";
-by (Auto_tac);
-qed "STAR_real_set";
-Addsimps [STAR_real_set];
-
-Goalw [starset_def] "*s* {} = {}";
-by (Step_tac 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (dres_inst_tac [("x","%n. xa n")] bspec 1);
-by (Auto_tac);
-qed "STAR_empty_set";
-Addsimps [STAR_empty_set];
-
-Goalw [starset_def] "*s* (A Un B) = *s* A Un *s* B";
-by (Auto_tac);
-by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2));
-by (dtac FreeUltrafilterNat_Compl_mem 1);
-by (dtac bspec 1 THEN assume_tac 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (Auto_tac);
-by (Fuf_tac 1);
-qed "STAR_Un";
-
-Goalw [starset_def] "*s* (A Int B) = *s* A Int *s* B";
-by (Auto_tac);
-by (blast_tac (claset() addIs [FreeUltrafilterNat_Int,
-               FreeUltrafilterNat_subset]) 3);
-by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1));
-qed "STAR_Int";
-
-Goalw [starset_def] "*s* -A = -(*s* A)";
-by (Auto_tac);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 2);
-by (REPEAT(Step_tac 1) THEN Auto_tac);
-by (Fuf_empty_tac 1);
-by (dtac FreeUltrafilterNat_Compl_mem 1);
-by (Fuf_tac 1);
-qed "STAR_Compl";
-
-goal Set.thy "(A - B) = (A Int (- B))";
-by (Step_tac 1);
-qed "set_diff_iff2";
-
-Goal "x ~: *s* F ==> x : *s* (- F)";
-by (auto_tac (claset(),simpset() addsimps [STAR_Compl]));
-qed "STAR_mem_Compl";
-
-Goal "*s* (A - B) = *s* A - *s* B";
-by (auto_tac (claset(),simpset() addsimps 
-         [set_diff_iff2,STAR_Int,STAR_Compl]));
-qed "STAR_diff";
-
-Goalw [starset_def] "A <= B ==> *s* A <= *s* B";
-by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1));
-qed "STAR_subset";
-
-Goalw [starset_def,hypreal_of_real_def] 
-          "a : A ==> hypreal_of_real a : *s* A";
-by (auto_tac (claset() addIs [FreeUltrafilterNat_subset],simpset()));
-qed "STAR_mem";
-
-Goalw [starset_def] "hypreal_of_real `` A <= *s* A";
-by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_def]));
-by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1);
-qed "STAR_hypreal_of_real_image_subset";
-
-Goalw [starset_def] "*s* X Int SReal = hypreal_of_real `` X";
-by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_def,SReal_def]));
-by (fold_tac [hypreal_of_real_def]);
-by (rtac imageI 1 THEN rtac ccontr 1);
-by (dtac bspec 1);
-by (rtac lemma_hyprel_refl 1);
-by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2);
-by (Auto_tac);
-qed "STAR_hypreal_of_real_Int";
-
-Goal "x ~: hypreal_of_real `` A ==> ALL y: A. x ~= hypreal_of_real y";
-by (Auto_tac);
-qed "lemma_not_hyprealA";
-
-Goal "- {n. X n = xa} = {n. X n ~= xa}";
-by (Auto_tac);
-qed "lemma_Compl_eq";
-
-Goalw [starset_def]
-    "ALL n. (X n) ~: M \
-\         ==> Abs_hypreal(hyprel^^{X}) ~: *s* M";
-by (Auto_tac THEN rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
-by (Auto_tac);
-qed "STAR_real_seq_to_hypreal";
-
-Goalw [starset_def] "*s* {x} = {hypreal_of_real x}";
-by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_def]));
-by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addIs [FreeUltrafilterNat_subset],simpset()));
-qed "STAR_singleton";
-Addsimps [STAR_singleton];
-
-Goal "x ~: F ==> hypreal_of_real x ~: *s* F";
-by (auto_tac (claset(),simpset() addsimps
-    [starset_def,hypreal_of_real_def]));
-by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
-by (Auto_tac);
-qed "STAR_not_mem";
-
-Goal "[| x : *s* A; A <= B |] ==> x : *s* B";
-by (blast_tac (claset() addDs [STAR_subset]) 1);
-qed "STAR_subset_closed";
-
-(*------------------------------------------------------------------ 
-   Nonstandard extension of a set (defined using a constant 
-   sequence) as a special case of an internal set
- -----------------------------------------------------------------*)
-
-Goalw [starset_n_def,starset_def] 
-     "ALL n. (As n = A) ==> *sn* As = *s* A";
-by (Auto_tac);
-qed "starset_n_starset";
-
-
-(*----------------------------------------------------------------*)
-(* Theorems about nonstandard extensions of functions             *)
-(*----------------------------------------------------------------*)
-
-(*----------------------------------------------------------------*) 
-(* Nonstandard extension of a function (defined using a           *)
-(* constant sequence) as a special case of an internal function   *)
-(*----------------------------------------------------------------*)
-
-Goalw [starfun_n_def,starfun_def] 
-     "ALL n. (F n = f) ==> *fn* F = *f* f";
-by (Auto_tac);
-qed "starfun_n_starfun";
-
-
-(* 
-   Prove that hrabs is a nonstandard extension of rabs without
-   use of congruence property (proved after this for general
-   nonstandard extensions of real valued functions). This makes 
-   proof much longer- see comments at end of HREALABS.thy where
-   we proved a congruence theorem for hrabs. 
-
-   NEW!!! No need to prove all the lemmas anymore. Use the ultrafilter
-   tactic! 
-*)
-  
-Goalw [is_starext_def] "is_starext abs abs";
-by (Step_tac 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by Auto_tac;
-by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
-by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
-by (auto_tac (claset() addSDs [spec],simpset() addsimps [hypreal_minus,hrabs_def,
-    hypreal_zero_def,hypreal_le_def,hypreal_less_def]));
-by (TRYALL(Ultra_tac));
-by (TRYALL(arith_tac));
-qed "hrabs_is_starext_rabs";
-
-Goal "[| X: Rep_hypreal z; Y: Rep_hypreal z |] \
-\              ==> {n. X n = Y n} : FreeUltrafilterNat";
-by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (Auto_tac THEN Fuf_tac 1);
-qed "Rep_hypreal_FreeUltrafilterNat";
-
-(*-----------------------------------------------------------------------
-    Nonstandard extension of functions- congruence 
- -----------------------------------------------------------------------*) 
-
-Goalw [congruent_def] "congruent hyprel (%X. hyprel^^{%n. f (X n)})";
-by (safe_tac (claset()));
-by (ALLGOALS(Fuf_tac));
-qed "starfun_congruent";
-
-Goalw [starfun_def]
-      "(*f* f) (Abs_hypreal(hyprel^^{%n. X n})) = \
-\      Abs_hypreal(hyprel ^^ {%n. f (X n)})";
-by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
-by (simp_tac (simpset() addsimps 
-   [hyprel_in_hypreal RS Abs_hypreal_inverse,[equiv_hyprel,
-   starfun_congruent] MRS UN_equiv_class]) 1);
-qed "starfun";
-
-(*-------------------------------------------
-  multiplication: ( *f ) x ( *g ) = *(f x g)  
- ------------------------------------------*)
-Goal "(*f* f) xa * (*f* g) xa = (*f* (%x. f x * g x)) xa";
-by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [starfun,hypreal_mult]));
-qed "starfun_mult";
-Addsimps [starfun_mult RS sym];
-
-(*---------------------------------------
-  addition: ( *f ) + ( *g ) = *(f + g)  
- ---------------------------------------*)
-Goal "(*f* f) xa + (*f* g) xa = (*f* (%x. f x + g x)) xa";
-by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [starfun,hypreal_add]));
-qed "starfun_add";
-Addsimps [starfun_add RS sym];
-
-(*--------------------------------------------
-  subtraction: ( *f ) + -( *g ) = *(f + -g)  
- -------------------------------------------*)
-
-Goal "- (*f* f) x = (*f* (%x. - f x)) x";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [starfun, hypreal_minus]));
-qed "starfun_minus";
-Addsimps [starfun_minus RS sym];
-
-(*FIXME: delete*)
-Goal "(*f* f) xa + -(*f* g) xa = (*f* (%x. f x + -g x)) xa";
-by (Simp_tac 1);
-qed "starfun_add_minus";
-Addsimps [starfun_add_minus RS sym];
-
-Goalw [hypreal_diff_def,real_diff_def]
-  "(*f* f) xa  - (*f* g) xa = (*f* (%x. f x - g x)) xa";
-by (rtac starfun_add_minus 1);
-qed "starfun_diff";
-Addsimps [starfun_diff RS sym];
-
-(*--------------------------------------
-  composition: ( *f ) o ( *g ) = *(f o g)  
- ---------------------------------------*)
-
-Goal "(%x. (*f* f) ((*f* g) x)) = *f* (%x. f (g x))"; 
-by (rtac ext 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [starfun]));
-qed "starfun_o2";
-
-Goalw [o_def] "(*f* f) o (*f* g) = (*f* (f o g))";
-by (simp_tac (simpset() addsimps [starfun_o2]) 1);
-qed "starfun_o";
-
-(*--------------------------------------
-  NS extension of constant function
- --------------------------------------*)
-Goal "(*f* (%x. k)) xa = hypreal_of_real  k";
-by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [starfun,
-    hypreal_of_real_def]));
-qed "starfun_const_fun";
-
-Addsimps [starfun_const_fun];
-
-(*----------------------------------------------------
-   the NS extension of the identity function
- ----------------------------------------------------*)
-
-Goal "x @= hypreal_of_real a ==> (*f* (%x. x)) x @= hypreal_of_real  a";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [starfun]));
-qed "starfun_Idfun_inf_close";
-
-Goal "(*f* (%x. x)) x = x";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [starfun]));
-qed "starfun_Id";
-Addsimps [starfun_Id];  
-
-(*----------------------------------------------------------------------
-      the *-function is a (nonstandard) extension of the function
- ----------------------------------------------------------------------*)
-
-Goalw [is_starext_def] "is_starext (*f* f) f";
-by (Auto_tac);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSIs [bexI] ,simpset() addsimps [starfun]));
-qed "is_starext_starfun";
-
-(*----------------------------------------------------------------------
-     Any nonstandard extension is in fact the *-function
- ----------------------------------------------------------------------*)
-
-Goalw [is_starext_def] "is_starext F f ==> F = *f* f";
-by (rtac ext 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (dres_inst_tac [("x","x")] spec 1);
-by (dres_inst_tac [("x","(*f* f) x")] spec 1);
-by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem],
-    simpset() addsimps [starfun]));
-by (Fuf_empty_tac 1);
-qed "is_starfun_starext";
-
-Goal "(is_starext F f) = (F = *f* f)";
-by (blast_tac (claset() addIs [is_starfun_starext,is_starext_starfun]) 1);
-qed "is_starext_starfun_iff";
-
-(*--------------------------------------------------------
-   extented function has same solution as its standard
-   version for real arguments. i.e they are the same
-   for all real arguments
- -------------------------------------------------------*)
-Goal "(*f* f) (hypreal_of_real a) = hypreal_of_real (f a)";
-by (auto_tac (claset(),simpset() addsimps 
-     [starfun,hypreal_of_real_def]));
-qed "starfun_eq";
-
-Addsimps [starfun_eq];
-
-Goal "(*f* f) (hypreal_of_real a) @= hypreal_of_real (f a)";
-by (Auto_tac);
-qed "starfun_inf_close";
-
-(* useful for NS definition of derivatives *)
-Goal "(*f* (%h. f (x + h))) xa  = (*f* f) (hypreal_of_real  x + xa)";
-by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [starfun,
-    hypreal_of_real_def,hypreal_add]));
-qed "starfun_lambda_cancel";
-
-Goal "(*f* (%h. f(g(x + h)))) xa = (*f* (f o g)) (hypreal_of_real x + xa)";
-by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [starfun,
-    hypreal_of_real_def,hypreal_add]));
-qed "starfun_lambda_cancel2";
-
-Goal "[| (*f* f) xa @= l; (*f* g) xa @= m; \
-\                 l: HFinite; m: HFinite  \
-\              |] ==>  (*f* (%x. f x * g x)) xa @= l * m";
-by (dtac inf_close_mult_HFinite 1);
-by (REPEAT(assume_tac 1));
-by (auto_tac (claset() addIs [inf_close_sym RSN (2,inf_close_HFinite)],
-              simpset()));
-qed "starfun_mult_HFinite_inf_close";
-
-Goal "[| (*f* f) xa @= l; (*f* g) xa @= m \
-\              |] ==>  (*f* (%x. f x + g x)) xa @= l + m";
-by (auto_tac (claset() addIs [inf_close_add], simpset()));
-qed "starfun_add_inf_close";
-
-(*----------------------------------------------------
-    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          *)
-
-Goal "*f* abs = abs";
-by (rtac (hrabs_is_starext_rabs RS 
-          (is_starext_starfun_iff RS iffD1) RS sym) 1);
-qed "starfun_rabs_hrabs";
-
-Goal "(*f* inverse) x = inverse(x)";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-            simpset() addsimps [starfun, hypreal_inverse, hypreal_zero_def]));
-qed "starfun_inverse_inverse";
-Addsimps [starfun_inverse_inverse];
-
-Goal "inverse ((*f* f) x) = (*f* (%x. inverse (f x))) x";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
-              simpset() addsimps [starfun, hypreal_inverse]));
-qed "starfun_inverse";
-Addsimps [starfun_inverse RS sym];
-
-Goalw [hypreal_divide_def,real_divide_def]
-  "(*f* f) xa  / (*f* g) xa = (*f* (%x. f x / g x)) xa";
-by Auto_tac;
-qed "starfun_divide";
-Addsimps [starfun_divide RS sym];
-
-Goal "inverse ((*f* f) x) = (*f* (%x. inverse (f x))) x";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addIs [FreeUltrafilterNat_subset]
-                       addSDs [FreeUltrafilterNat_Compl_mem],
-    simpset() addsimps [starfun, hypreal_inverse, hypreal_zero_def]));
-qed "starfun_inverse2";
-
-(*-------------------------------------------------------------
-    General lemma/theorem needed for proofs in elementary
-    topology of the reals
- ------------------------------------------------------------*)
-Goalw [starset_def] 
-      "(*f* f) x : *s* A ==> x : *s* {x. f x : A}";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [starfun]));
-by (dres_inst_tac [("x","%n. f (Xa n)")] bspec 1);
-by (Auto_tac THEN Fuf_tac 1);
-qed "starfun_mem_starset";
-
-(*------------------------------------------------------------
-   Alternative definition for hrabs with rabs function
-   applied entrywise to equivalence class representative.
-   This is easily proved using starfun and ns extension thm
- ------------------------------------------------------------*)
-Goal "abs (Abs_hypreal (hyprel ^^ {X})) = \
-\                 Abs_hypreal(hyprel ^^ {%n. abs (X n)})";
-by (simp_tac (simpset() addsimps [starfun_rabs_hrabs RS sym,starfun]) 1);
-qed "hypreal_hrabs";
-
-(*----------------------------------------------------------------
-   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 ( *f* f). See second NS set extension below.
- ----------------------------------------------------------------*)
-Goalw [starset_def]
-   "*s* {x. abs (x + - y) < r} = \
-\    {x. abs(x + -hypreal_of_real y) < hypreal_of_real r}";
-by (Step_tac 1);
-by (ALLGOALS(res_inst_tac [("z","x")] eq_Abs_hypreal));
-by (auto_tac (claset() addSIs [exI] addSDs [bspec],
-          simpset() addsimps [hypreal_minus, hypreal_of_real_def,hypreal_add,
-                              hypreal_hrabs,hypreal_less_def]));
-by (Fuf_tac 1);
-qed "STAR_rabs_add_minus";
-
-Goalw [starset_def]
-  "*s* {x. abs (f x + - y) < r} = \
-\      {x. abs((*f* f) x + -hypreal_of_real y) < hypreal_of_real r}";
-by (Step_tac 1);
-by (ALLGOALS(res_inst_tac [("z","x")] eq_Abs_hypreal));
-by (auto_tac (claset() addSIs [exI] addSDs [bspec],
-    simpset() addsimps [hypreal_minus, hypreal_of_real_def,hypreal_add,
-    hypreal_hrabs,hypreal_less_def,starfun]));
-by (Fuf_tac 1);
-qed "STAR_starfun_rabs_add_minus";
-
-(*-------------------------------------------------------------------
-   Another charaterization of Infinitesimal and one of @= relation. 
-   In this theory since hypreal_hrabs proved here. (To Check:) Maybe 
-   move both if possible? 
- -------------------------------------------------------------------*)
-Goal "(x:Infinitesimal) = \
-\     (EX X:Rep_hypreal(x). \
-\       ALL m. {n. abs(X n) < inverse(real_of_posnat m)}:FreeUltrafilterNat)";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSIs [bexI,lemma_hyprel_refl],
-	simpset() addsimps [Infinitesimal_hypreal_of_posnat_iff,
-	    hypreal_of_real_of_posnat,hypreal_of_real_def,hypreal_inverse,
-	    hypreal_hrabs,hypreal_less])); 
-by (dres_inst_tac [("x","n")] spec 1);
-by (Fuf_tac 1);
-qed  "Infinitesimal_FreeUltrafilterNat_iff2";
-
-Goal "(Abs_hypreal(hyprel^^{X}) @= Abs_hypreal(hyprel^^{Y})) = \
-\     (ALL m. {n. abs (X n + - Y n) < \
-\                 inverse(real_of_posnat m)} : FreeUltrafilterNat)";
-by (rtac (inf_close_minus_iff RS ssubst) 1);
-by (rtac (mem_infmal_iff RS subst) 1);
-by (auto_tac (claset(), 
-              simpset() addsimps [hypreal_minus, hypreal_add,
-                                  Infinitesimal_FreeUltrafilterNat_iff2]));
-by (dres_inst_tac [("x","m")] spec 1);
-by (Fuf_tac 1);
-qed "inf_close_FreeUltrafilterNat_iff";
-
-Goal "inj starfun";
-by (rtac injI 1);
-by (rtac ext 1 THEN rtac ccontr 1);
-by (dres_inst_tac [("x","Abs_hypreal(hyprel ^^{%n. xa})")] fun_cong 1);
-by (auto_tac (claset(),simpset() addsimps [starfun]));
-qed "inj_starfun";
-
-
-
-
-
--- a/src/HOL/Real/Hyperreal/Star.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,39 +0,0 @@
-(*  Title       : Star.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : defining *-transforms in NSA which extends sets of reals, 
-                  and real=>real functions
-*) 
-
-Star = NSA +
-
-constdefs
-    (* nonstandard extension of sets *)
-    starset :: real set => hypreal set          ("*s* _" [80] 80)
-    "*s* A  == {x. ALL X: Rep_hypreal(x). {n::nat. X n : A}: FreeUltrafilterNat}"
-
-    (* internal sets *)
-    starset_n :: (nat => real set) => hypreal set        ("*sn* _" [80] 80)
-    "*sn* As  == {x. ALL X: Rep_hypreal(x). {n::nat. X n : (As n)}: FreeUltrafilterNat}"   
-    
-    InternalSets :: "hypreal set set"
-    "InternalSets == {X. EX As. X = *sn* As}"
-
-    (* nonstandard extension of function *)
-    is_starext  ::  [hypreal => hypreal, real => real] => bool
-    "is_starext F f == (ALL x y. EX X: Rep_hypreal(x). EX Y: Rep_hypreal(y).
-                        ((y = (F x)) = ({n. Y n = f(X n)} : FreeUltrafilterNat)))"
-    
-    starfun :: (real => real) => hypreal => hypreal        ("*f* _" [80] 80)
-    "*f* f  == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel^^{%n. f (X n)}))" 
-
-    (* internal functions *)
-    starfun_n :: (nat => (real => real)) => hypreal => hypreal        ("*fn* _" [80] 80)
-    "*fn* F  == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel^^{%n. (F n)(X n)}))" 
-
-    InternalFuns :: (hypreal => hypreal) set
-    "InternalFuns == {X. EX F. X = *fn* F}"
-end
-
-
-
--- a/src/HOL/Real/Hyperreal/Zorn.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,285 +0,0 @@
-(*  Title       : Zorn.ML
-    ID          : $Id$
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Zorn's Lemma -- adapted proofs from lcp's ZF/Zorn.ML
-*) 
-
-(*---------------------------------------------------------------
-      Section 1.  Mathematical Preamble 
- ---------------------------------------------------------------*)
-
-Goal "(ALL x:C. x<=A | B<=x) ==> Union(C)<=A | B<=Union(C)";
-by (Blast_tac 1);
-qed "Union_lemma0";
-
-(*-- similar to subset_cs in ZF/subset.thy --*)
-bind_thms ("thissubset_SIs",
-    [subset_refl,Union_least, UN_least, Un_least, 
-     Inter_greatest, Int_greatest,
-     Un_upper1, Un_upper2, Int_lower1, Int_lower2]);
-
-
-(*A claset for subset reasoning*)
-val thissubset_cs = claset() 
-    delrules [subsetI, subsetCE]
-    addSIs thissubset_SIs
-    addIs  [Union_upper, Inter_lower];
-
-(* increasingD2 of ZF/Zorn.ML *) 
-Goalw [succ_def] "x <= succ S x";
-by (rtac (split_if RS iffD2) 1);
-by (auto_tac (claset(),simpset() addsimps [super_def,
-               maxchain_def,psubset_def]));
-by (rtac swap 1 THEN assume_tac 1);
-by (rtac someI2 1);
-by (ALLGOALS(Blast_tac));
-qed "Abrial_axiom1";
-
-val [TFin_succI, Pow_TFin_UnionI] = TFin.intrs;
-val TFin_UnionI = PowI RS Pow_TFin_UnionI;
-bind_thm ("TFin_succI", TFin_succI);
-bind_thm ("Pow_TFin_UnionI", Pow_TFin_UnionI);
-bind_thm ("TFin_UnionI", TFin_UnionI);
-
-val major::prems = Goal  
-          "[| n : TFin S; \
-\            !!x. [| x: TFin S; P(x) |] ==> P(succ S x); \
-\            !!Y. [| Y <= TFin S; Ball Y P |] ==> P(Union Y) |] \
-\         ==> P(n)";
-by (rtac (major RS TFin.induct) 1);
-by (ALLGOALS (fast_tac (claset() addIs prems)));
-qed "TFin_induct";
-
-(*Perform induction on n, then prove the major premise using prems. *)
-fun TFin_ind_tac a prems i = 
-    EVERY [induct_thm_tac TFin_induct a i,
-           rename_last_tac a ["1"] (i+1),
-           rename_last_tac a ["2"] (i+2),
-           ares_tac prems i];
-
-Goal "x <= y ==> x <= succ S y";
-by (etac (Abrial_axiom1 RSN (2,subset_trans)) 1);
-qed "succ_trans";
-
-(*Lemma 1 of section 3.1*)
-Goal "[| n: TFin S;  m: TFin S;  \
-\        ALL x: TFin S. x <= m --> x = m | succ S x <= m \
-\     |] ==> n <= m | succ S m <= n";
-by (etac TFin_induct 1);
-by (etac Union_lemma0 2);               (*or just Blast_tac*)
-by (blast_tac (thissubset_cs addIs [succ_trans]) 1);
-qed "TFin_linear_lemma1";
-
-(* Lemma 2 of section 3.2 *)
-Goal "m: TFin S ==> ALL n: TFin S. n<=m --> n=m | succ S n<=m";
-by (etac TFin_induct 1);
-by (rtac (impI RS ballI) 1);
-(*case split using TFin_linear_lemma1*)
-by (res_inst_tac [("n1","n"), ("m1","x")] 
-    (TFin_linear_lemma1 RS disjE) 1  THEN  REPEAT (assume_tac 1));
-by (dres_inst_tac [("x","n")] bspec 1 THEN assume_tac 1);
-by (blast_tac (thissubset_cs addIs [succ_trans]) 1);
-by (REPEAT (ares_tac [disjI1,equalityI] 1));
-(*second induction step*)
-by (rtac (impI RS ballI) 1);
-by (rtac (Union_lemma0 RS disjE) 1);
-by (rtac disjI2 3);
-by (REPEAT (ares_tac [disjI1,equalityI] 2));
-by (rtac ballI 1);
-by (ball_tac 1);
-by (set_mp_tac 1);
-by (res_inst_tac [("n1","n"), ("m1","x")] 
-    (TFin_linear_lemma1 RS disjE) 1  THEN  REPEAT (assume_tac 1));
-by (blast_tac thissubset_cs 1);
-by (rtac (Abrial_axiom1 RS subset_trans RS disjI1) 1);
-by (assume_tac 1);
-qed "TFin_linear_lemma2";
-
-(*a more convenient form for Lemma 2*)
-Goal "[| n<=m;  m: TFin S;  n: TFin S |] ==> n=m | succ S n<=m";
-by (rtac (TFin_linear_lemma2 RS bspec RS mp) 1);
-by (REPEAT (assume_tac 1));
-qed "TFin_subsetD";
-
-(*Consequences from section 3.3 -- Property 3.2, the ordering is total*)
-Goal "[| m: TFin S;  n: TFin S|] ==> n<=m | m<=n";
-by (rtac (TFin_linear_lemma2 RSN (3,TFin_linear_lemma1) RS disjE) 1);
-by (REPEAT (assume_tac 1) THEN etac disjI2 1);
-by (blast_tac (thissubset_cs addIs [Abrial_axiom1 RS subset_trans]) 1);
-qed "TFin_subset_linear";
-
-(*Lemma 3 of section 3.3*)
-Goal "[| n: TFin S;  m: TFin S;  m = succ S m |] ==> n<=m";
-by (etac TFin_induct 1);
-by (dtac TFin_subsetD 1);
-by (REPEAT (assume_tac 1));
-by (fast_tac (claset() addEs [ssubst]) 1);
-by (blast_tac (thissubset_cs) 1);
-qed "eq_succ_upper";
-
-(*Property 3.3 of section 3.3*)
-Goal "m: TFin S ==> (m = succ S m) = (m = Union(TFin S))";
-by (rtac iffI 1);
-by (rtac (Union_upper RS equalityI) 1);
-by (rtac (eq_succ_upper RS Union_least) 2);
-by (REPEAT (assume_tac 1));
-by (etac ssubst 1);
-by (rtac (Abrial_axiom1 RS equalityI) 1);
-by (blast_tac (thissubset_cs addIs [TFin_UnionI, TFin_succI]) 1);
-qed "equal_succ_Union";
-
-(*-------------------------------------------------------------------------
-    Section 4.  Hausdorff's Theorem: every set contains a maximal chain 
-    NB: We assume the partial ordering is <=, the subset relation! 
- -------------------------------------------------------------------------*)
-
-Goalw [chain_def] "({} :: 'a set set) : chain S";
-by (Auto_tac);
-qed "empty_set_mem_chain";
-
-Goalw [super_def] "super S c <= chain S";
-by (Fast_tac 1);
-qed "super_subset_chain";
-
-Goalw [maxchain_def] "maxchain S <= chain S";
-by (Fast_tac 1);
-qed "maxchain_subset_chain";
-
-Goalw [succ_def] "c ~: chain S ==> succ S c = c";
-by (fast_tac (claset() addSIs [if_P]) 1);
-qed "succI1";
-
-Goalw [succ_def] "c: maxchain S ==> succ S c = c";
-by (fast_tac (claset() addSIs [if_P]) 1);
-qed "succI2";
-
-Goalw [succ_def] "c: chain S - maxchain S ==> \
-\                         succ S c = (@c'. c': super S c)";
-by (fast_tac (claset() addSIs [if_not_P]) 1);
-qed "succI3";
-
-Goal "c: chain S - maxchain S ==> ? d. d: super S c";
-by (rewrite_goals_tac [super_def,maxchain_def]);
-by (Auto_tac);
-qed "mem_super_Ex";
-
-Goal "c: chain S - maxchain S ==> \
-\                         (@c'. c': super S c): super S c";
-by (etac (mem_super_Ex RS exE) 1);
-by (rtac someI2 1);
-by (Auto_tac);
-qed "select_super";
-
-Goal "c: chain S - maxchain S ==> \
-\                         (@c'. c': super S c) ~= c";
-by (rtac notI 1);
-by (dtac select_super 1);
-by (asm_full_simp_tac (simpset() addsimps [super_def,psubset_def]) 1);
-qed "select_not_equals";
-
-Goal "c: chain S - maxchain S ==> \
-\                         succ S c ~= c";
-by (ftac succI3 1);
-by (Asm_simp_tac 1);
-by (rtac select_not_equals 1);
-by (assume_tac 1);
-qed "succ_not_equals";
-
-Goal "c: TFin S ==> (c :: 'a set set): chain S";
-by (etac TFin_induct 1);
-by (asm_simp_tac (simpset() addsimps [succ_def,
-    select_super RS (super_subset_chain RS subsetD)]
-                   addsplits [split_if]) 1);
-by (rewtac chain_def);
-by (rtac CollectI 1);
-by Safe_tac;
-by (dtac bspec 1 THEN assume_tac 1);
-by (res_inst_tac  [("m1","Xa"), ("n1","X")] (TFin_subset_linear RS disjE) 2);
-by (ALLGOALS(Blast_tac));
-qed "TFin_chain_lemm4";
- 
-Goal "EX c. (c :: 'a set set): maxchain S";
-by (res_inst_tac [("x", "Union(TFin S)")] exI 1);
-by (rtac classical 1);
-by (subgoal_tac "succ S (Union(TFin S)) = Union(TFin S)" 1);
-by (resolve_tac [equal_succ_Union RS iffD2 RS sym] 2);
-by (resolve_tac [subset_refl RS TFin_UnionI] 2);
-by (rtac refl 2);
-by (cut_facts_tac [subset_refl RS TFin_UnionI RS TFin_chain_lemm4] 1);
-by (dtac (DiffI RS succ_not_equals) 1);
-by (ALLGOALS(Blast_tac));
-qed "Hausdorff";
-
-
-(*---------------------------------------------------------------
-  Section 5.  Zorn's Lemma: if all chains have upper bounds 
-                               there is  a maximal element 
- ----------------------------------------------------------------*)
-Goalw [chain_def]
-    "[| c: chain S; z: S; \
-\             ALL x:c. x<=(z:: 'a set) |] ==> {z} Un c : chain S";
-by (Blast_tac 1);
-qed "chain_extend";
-
-Goalw [chain_def] "[| c: chain S; x: c |] ==> x <= Union(c)";
-by (Auto_tac);
-qed "chain_Union_upper";
-
-Goalw [chain_def] "c: chain S ==> ! x: c. x <= Union(c)";
-by (Auto_tac);
-qed "chain_ball_Union_upper";
-
-Goal "[| c: maxchain S; u: S; Union(c) <= u |] ==> Union(c) = u";
-by (rtac ccontr 1);
-by (asm_full_simp_tac (simpset() addsimps [maxchain_def]) 1);
-by (etac conjE 1);
-by (subgoal_tac "({u} Un c): super S c" 1);
-by (Asm_full_simp_tac 1);
-by (rewrite_tac [super_def,psubset_def]);
-by (blast_tac (claset() addIs [chain_extend] addDs [chain_Union_upper]) 1);
-qed "maxchain_Zorn";
-
-Goal "ALL c: chain S. Union(c): S ==> \
-\     EX y: S. ALL z: S. y <= z --> y = z";
-by (cut_facts_tac [Hausdorff,maxchain_subset_chain] 1);
-by (etac exE 1);
-by (dtac subsetD 1 THEN assume_tac 1);
-by (dtac bspec 1 THEN assume_tac 1);
-by (res_inst_tac [("x","Union(c)")] bexI 1);
-by (rtac ballI 1 THEN rtac impI 1);
-by (blast_tac (claset() addSDs [maxchain_Zorn]) 1);
-by (assume_tac 1);
-qed "Zorn_Lemma";
-
-(*-------------------------------------------------------------
-             Alternative version of Zorn's Lemma
- --------------------------------------------------------------*)
-Goal "ALL (c:: 'a set set): chain S. EX y : S. ALL x : c. x <= y ==> \
-\     EX y : S. ALL x : S. (y :: 'a set) <= x --> y = x";
-by (cut_facts_tac [Hausdorff,maxchain_subset_chain] 1);
-by (EVERY1[etac exE, dtac subsetD, assume_tac]);
-by (EVERY1[dtac bspec, assume_tac, etac bexE]);
-by (res_inst_tac [("x","y")] bexI 1);
-by (assume_tac 2);
-by (EVERY1[rtac ballI, rtac impI, rtac ccontr]);
-by (forw_inst_tac [("z","x")]  chain_extend 1);
-by (assume_tac 1 THEN Blast_tac 1);
-by (rewrite_tac [maxchain_def,super_def,psubset_def]);
-by (blast_tac (claset() addSEs [equalityCE]) 1);
-qed "Zorn_Lemma2";
-
-(** misc. lemmas **)
-
-Goalw [chain_def] "[| c : chain S; x: c; y: c |] ==> x <= y | y <= x";
-by (Blast_tac 1);
-qed "chainD";
-
-Goalw [chain_def] "!!(c :: 'a set set). c: chain S ==> c <= S";
-by (Blast_tac 1);
-qed "chainD2";
-
-(* proved elsewhere? *) 
-Goal "x : Union(c) ==> EX m:c. x:m";
-by (Blast_tac 1);
-qed "mem_UnionD";
--- a/src/HOL/Real/Hyperreal/Zorn.thy	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,34 +0,0 @@
-(*  Title       : Zorn.thy
-    ID          : $Id$
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Zorn's Lemma -- See lcp's Zorn.thy in ZF
-*) 
-
-Zorn = Main +
-
-constdefs
-  chain     ::  'a::ord set => 'a set set
-    "chain S  == {F. F <= S & (ALL x: F. ALL y: F. x <= y | y <= x)}" 
-
-  super     ::  ['a::ord set,'a set] => 'a set set
-    "super S c == {d. d: chain(S) & c < d}"
-
-  maxchain  ::  'a::ord set => 'a set set
-    "maxchain S == {c. c: chain S & super S c = {}}"
-
-  succ      ::  ['a::ord set,'a set] => 'a set
-    "succ S c == if (c ~: chain S| c: maxchain S) 
-                 then c else (@c'. c': (super S c))" 
-
-consts 
-  "TFin" :: 'a::ord set => 'a set set
-
-inductive "TFin(S)"
-  intrs
-    succI        "x : TFin S ==> succ S x : TFin S"
-    Pow_UnionI   "Y : Pow(TFin S) ==> Union(Y) : TFin S"
-           
-  monos          Pow_mono
-end
-
--- a/src/HOL/Real/Hyperreal/fuf.ML	Sat Dec 30 22:03:46 2000 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,78 +0,0 @@
-(*  Title       : HOL/Real/Hyperreal/fuf.ML
-    ID          : $Id$
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-                  1999  University of Edinburgh
-
-Simple tactics to help proofs involving our free ultrafilter
-(FreeUltrafilterNat). We rely on the fact that filters satisfy the
-finite intersection property.
-*)
-
-local
-
-exception FUFempty;
-
-fun get_fuf_hyps [] zs = zs
-|   get_fuf_hyps (x::xs) zs =
-        case (concl_of x) of
-        (_ $ (Const ("Not",_) $ (Const ("op :",_) $ _ $
-             Const ("HyperDef.FreeUltrafilterNat",_)))) =>  get_fuf_hyps xs
-                                              ((x RS FreeUltrafilterNat_Compl_mem)::zs)
-       |(_ $ (Const ("op :",_) $ _ $
-             Const ("HyperDef.FreeUltrafilterNat",_)))  =>  get_fuf_hyps xs (x::zs)
-       | _ => get_fuf_hyps xs zs;
-
-fun inter_prems [] = raise FUFempty
-|   inter_prems [x] = x
-|   inter_prems (x::y::ys) =
-      inter_prems (([x,y] MRS FreeUltrafilterNat_Int) :: ys);
-
-in
-
-(*---------------------------------------------------------------
-   solves goals of the form
-    [| A1: FUF; A2: FUF; ...; An: FUF |] ==> B : FUF
-   where A1 Int A2 Int ... Int An <= B
- ---------------------------------------------------------------*)
-
-fun fuf_tac css i = METAHYPS(fn prems =>
-                    (rtac ((inter_prems (get_fuf_hyps prems [])) RS
-                           FreeUltrafilterNat_subset) 1) THEN
-                    auto_tac css) i;
-
-fun Fuf_tac i = fuf_tac (clasimpset ()) i;
-
-
-(*---------------------------------------------------------------
-   solves goals of the form
-    [| A1: FUF; A2: FUF; ...; An: FUF |] ==> P
-   where A1 Int A2 Int ... Int An <= {} since {} ~: FUF
-   (i.e. uses fact that FUF is a proper filter)
- ---------------------------------------------------------------*)
-
-fun fuf_empty_tac css i = METAHYPS (fn prems =>
-  rtac ((inter_prems (get_fuf_hyps prems [])) RS
-    (FreeUltrafilterNat_subset RS (FreeUltrafilterNat_empty RS notE))) 1
-                     THEN auto_tac css) i;
-
-fun Fuf_empty_tac i = fuf_empty_tac (clasimpset ()) i;
-
-
-(*---------------------------------------------------------------
-  All in one -- not really needed.
- ---------------------------------------------------------------*)
-
-fun fuf_auto_tac css i = SOLVE (fuf_empty_tac css i) ORELSE TRY (fuf_tac css i);
-fun Fuf_auto_tac i = fuf_auto_tac (clasimpset ()) i;
-
-
-(*---------------------------------------------------------------
-   In fact could make this the only tactic: just need to
-   use contraposition and then look for empty set.
- ---------------------------------------------------------------*)
-
-fun ultra_tac css i = rtac ccontr i THEN fuf_empty_tac css i;
-fun Ultra_tac i = ultra_tac (clasimpset ()) i;
-
-end;