starfun, starset, and other functions on NS types are now polymorphic;
authorhuffman
Fri, 09 Sep 2005 19:34:22 +0200
changeset 17318 bc1c75855f3d
parent 17317 3f12de2e2e6e
child 17319 14c6e073252f
starfun, starset, and other functions on NS types are now polymorphic; many similar theorems have been generalized and merged; (star_n X) replaces (Abs_star(starrel `` {X})); many proofs have been simplified with the transfer tactic.
src/HOL/Complex/CLim.thy
src/HOL/Complex/CStar.thy
src/HOL/Complex/NSCA.thy
src/HOL/Complex/NSComplex.thy
src/HOL/Hyperreal/HLog.thy
src/HOL/Hyperreal/HSeries.thy
src/HOL/Hyperreal/HTranscendental.thy
src/HOL/Hyperreal/HyperArith.thy
src/HOL/Hyperreal/HyperDef.thy
src/HOL/Hyperreal/HyperNat.thy
src/HOL/Hyperreal/HyperPow.thy
src/HOL/Hyperreal/Integration.thy
src/HOL/Hyperreal/Lim.thy
src/HOL/Hyperreal/NSA.thy
src/HOL/Hyperreal/NatStar.thy
src/HOL/Hyperreal/SEQ.thy
src/HOL/Hyperreal/Star.thy
src/HOL/Hyperreal/StarType.thy
src/HOL/Hyperreal/Transcendental.thy
src/HOL/Hyperreal/hypreal_arith.ML
--- a/src/HOL/Complex/CLim.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Complex/CLim.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -46,7 +46,7 @@
 			      ("((_)/ -- (_)/ --NSC> (_))" [60, 0, 60] 60)
   "f -- a --NSC> L == (\<forall>x. (x \<noteq> hcomplex_of_complex a &
            		         x @c= hcomplex_of_complex a
-                                   --> ( *fc* f) x @c= hcomplex_of_complex L))"
+                                   --> ( *f* f) x @c= hcomplex_of_complex L))"
 
   (* f: C --> R *)
   CRLIM :: "[complex=>real,complex,real] => bool"
@@ -60,7 +60,7 @@
 			      ("((_)/ -- (_)/ --NSCR> (_))" [60, 0, 60] 60)
   "f -- a --NSCR> L == (\<forall>x. (x \<noteq> hcomplex_of_complex a &
            		         x @c= hcomplex_of_complex a
-                                   --> ( *fcR* f) x @= hypreal_of_real L))"
+                                   --> ( *f* f) x @= hypreal_of_real L))"
 
 
   isContc :: "[complex=>complex,complex] => bool"
@@ -69,7 +69,7 @@
   (* NS definition dispenses with limit notions *)
   isNSContc :: "[complex=>complex,complex] => bool"
   "isNSContc f a == (\<forall>y. y @c= hcomplex_of_complex a -->
-			   ( *fc* f) y @c= hcomplex_of_complex (f a))"
+			   ( *f* f) y @c= hcomplex_of_complex (f a))"
 
   isContCR :: "[complex=>real,complex] => bool"
   "isContCR f a == (f -- a --CR> (f a))"
@@ -77,7 +77,7 @@
   (* NS definition dispenses with limit notions *)
   isNSContCR :: "[complex=>real,complex] => bool"
   "isNSContCR f a == (\<forall>y. y @c= hcomplex_of_complex a -->
-			   ( *fcR* f) y @= hypreal_of_real (f a))"
+			   ( *f* f) y @= hypreal_of_real (f a))"
 
   (* differentiation: D is derivative of function f at x *)
   cderiv:: "[complex=>complex,complex,complex] => bool"
@@ -87,7 +87,7 @@
   nscderiv :: "[complex=>complex,complex,complex] => bool"
 			    ("(NSCDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60)
   "NSCDERIV f x :> D == (\<forall>h \<in> CInfinitesimal - {0}.
-			      (( *fc* f)(hcomplex_of_complex x + h)
+			      (( *f* f)(hcomplex_of_complex x + h)
         			 - hcomplex_of_complex (f x))/h @c= hcomplex_of_complex D)"
 
   cdifferentiable :: "[complex=>complex,complex] => bool"
@@ -105,7 +105,7 @@
 			    --> cmod(f x - f y) < r)))"
 
   isNSUContc :: "(complex=>complex) => bool"
-  "isNSUContc f == (\<forall>x y. x @c= y --> ( *fc* f) x @c= ( *fc* f) y)"
+  "isNSUContc f == (\<forall>x y. x @c= y --> ( *f* f) x @c= ( *f* f) y)"
 
 
 
@@ -123,20 +123,18 @@
 lemma CLIM_NSCLIM:
       "f -- x --C> L ==> f -- x --NSC> L"
 apply (simp add: CLIM_def NSCLIM_def capprox_def, auto)
-apply (rule_tac z = xa in eq_Abs_star)
-apply (auto simp add: hcomplex_of_complex_def starfunC hcomplex_diff 
-         star_of_def star_n_def
+apply (rule_tac x = xa in star_cases)
+apply (auto simp add: starfun star_n_diff star_of_def star_n_eq_iff
          CInfinitesimal_hcmod_iff hcmod Infinitesimal_FreeUltrafilterNat_iff)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl, safe)
+apply (rule bexI [OF _ Rep_star_star_n], safe)
 apply (drule_tac x = u in spec, auto)
 apply (drule_tac x = s in spec, auto, ultra)
 apply (drule sym, auto)
 done
 
-lemma eq_Abs_star_ALL:
-     "(\<forall>t. P t) = (\<forall>X. P (Abs_star(starrel `` {X})))"
+lemma eq_Abs_star_ALL: "(\<forall>t. P t) = (\<forall>X. P (star_n X))"
 apply auto
-apply (rule_tac z = t in eq_Abs_star, auto)
+apply (rule_tac x = t in star_cases, auto)
 done
 
 lemma lemma_CLIM:
@@ -169,18 +167,16 @@
      "f -- x --NSC> L ==> f -- x --C> L"
 apply (simp add: CLIM_def NSCLIM_def)
 apply (rule ccontr) 
-apply (auto simp add: eq_Abs_star_ALL starfunC 
-            CInfinitesimal_capprox_minus [symmetric] hcomplex_diff
-            CInfinitesimal_hcmod_iff hcomplex_of_complex_def 
-            star_of_def star_n_def
+apply (auto simp add: eq_Abs_star_ALL starfun
+            CInfinitesimal_capprox_minus [symmetric] star_n_diff
+            CInfinitesimal_hcmod_iff star_of_def star_n_eq_iff
             Infinitesimal_FreeUltrafilterNat_iff hcmod)
 apply (simp add: linorder_not_less)
 apply (drule lemma_skolemize_CLIM2, safe)
 apply (drule_tac x = X in spec, auto)
 apply (drule lemma_csimp [THEN complex_seq_to_hcomplex_CInfinitesimal])
-apply (simp add: CInfinitesimal_hcmod_iff hcomplex_of_complex_def 
-            star_of_def star_n_def
-            Infinitesimal_FreeUltrafilterNat_iff hcomplex_diff hcmod,  blast)
+apply (simp add: CInfinitesimal_hcmod_iff star_of_def
+            Infinitesimal_FreeUltrafilterNat_iff star_n_diff hcmod,  blast)
 apply (drule_tac x = r in spec, clarify)
 apply (drule FreeUltrafilterNat_all, ultra, arith)
 done
@@ -195,13 +191,13 @@
 
 lemma CRLIM_NSCRLIM: "f -- x --CR> L ==> f -- x --NSCR> L"
 apply (simp add: CRLIM_def NSCRLIM_def capprox_def, auto)
-apply (rule_tac z = xa in eq_Abs_star)
-apply (auto simp add: hcomplex_of_complex_def starfunCR hcomplex_diff
-              CInfinitesimal_hcmod_iff hcmod hypreal_diff
+apply (rule_tac x = xa in star_cases)
+apply (auto simp add: starfun star_n_diff
+              CInfinitesimal_hcmod_iff hcmod
               Infinitesimal_FreeUltrafilterNat_iff
-              star_of_def star_n_def
-              Infinitesimal_approx_minus [symmetric] hypreal_of_real_def)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl, safe)
+              star_of_def star_n_eq_iff
+              Infinitesimal_approx_minus [symmetric])
+apply (rule bexI [OF _ Rep_star_star_n], safe)
 apply (drule_tac x = u in spec, auto)
 apply (drule_tac x = s in spec, auto, ultra)
 apply (drule sym, auto)
@@ -235,19 +231,18 @@
 lemma NSCRLIM_CRLIM: "f -- x --NSCR> L ==> f -- x --CR> L"
 apply (simp add: CRLIM_def NSCRLIM_def capprox_def)
 apply (rule ccontr) 
-apply (auto simp add: eq_Abs_star_ALL starfunCR hcomplex_diff 
-             hcomplex_of_complex_def hypreal_diff CInfinitesimal_hcmod_iff 
+apply (auto simp add: eq_Abs_star_ALL starfun star_n_diff 
+             CInfinitesimal_hcmod_iff 
              hcmod Infinitesimal_approx_minus [symmetric]
-             star_of_def star_n_def 
+             star_of_def star_n_eq_iff
              Infinitesimal_FreeUltrafilterNat_iff)
 apply (simp add: linorder_not_less)
 apply (drule lemma_skolemize_CRLIM2, safe)
 apply (drule_tac x = X in spec, auto)
 apply (drule lemma_crsimp [THEN complex_seq_to_hcomplex_CInfinitesimal])
-apply (simp add: CInfinitesimal_hcmod_iff hcomplex_of_complex_def 
-             star_of_def star_n_def
-             Infinitesimal_FreeUltrafilterNat_iff hcomplex_diff hcmod)
-apply (auto simp add: hypreal_of_real_def star_of_def star_n_def hypreal_diff)
+apply (simp add: CInfinitesimal_hcmod_iff star_of_def
+             Infinitesimal_FreeUltrafilterNat_iff star_n_diff hcmod)
+apply (auto simp add: star_of_def star_n_diff)
 apply (drule_tac x = r in spec, clarify)
 apply (drule FreeUltrafilterNat_all, ultra)
 done
@@ -413,8 +408,8 @@
    "(f -- x --NSC> L) = ((%y. cmod(f y - L)) -- x --NSCR> 0)"
 apply (auto simp add: NSCLIM_def NSCRLIM_def CInfinitesimal_capprox_minus [symmetric] CInfinitesimal_hcmod_iff)
 apply (auto dest!: spec) 
-apply (rule_tac [!] z = xa in eq_Abs_star)
-apply (auto simp add: hcomplex_diff starfunC starfunCR hcomplex_of_complex_def hcmod mem_infmal_iff star_of_def star_n_def)
+apply (rule_tac [!] x = xa in star_cases)
+apply (auto simp add: star_n_diff starfun hcmod mem_infmal_iff star_of_def)
 done
 
 (** much, much easier standard proof **)
@@ -432,8 +427,8 @@
 apply (auto intro: NSCLIM_NSCRLIM_Re NSCLIM_NSCRLIM_Im)
 apply (auto simp add: NSCLIM_def NSCRLIM_def)
 apply (auto dest!: spec) 
-apply (rule_tac z = x in eq_Abs_star)
-apply (simp add: capprox_approx_iff starfunC hcomplex_of_complex_def starfunCR hypreal_of_real_def star_of_def star_n_def)
+apply (rule_tac x = x in star_cases)
+apply (simp add: capprox_approx_iff starfun star_of_def)
 done
 
 lemma CLIM_CRLIM_Re_Im_iff:
@@ -446,7 +441,7 @@
 
 lemma isNSContcD:
       "[| isNSContc f a; y @c= hcomplex_of_complex a |]
-       ==> ( *fc* f) y @c= hcomplex_of_complex (f a)"
+       ==> ( *f* f) y @c= hcomplex_of_complex (f a)"
 by (simp add: isNSContc_def)
 
 lemma isNSContc_NSCLIM: "isNSContc f a ==> f -- a --NSC> (f a) "
@@ -485,10 +480,10 @@
 apply (drule_tac [2] x = "- hcomplex_of_complex a + x" in spec, safe, simp)
 apply (rule mem_cinfmal_iff [THEN iffD2, THEN CInfinitesimal_add_capprox_self [THEN capprox_sym]])
 apply (rule_tac [4] capprox_minus_iff2 [THEN iffD1])
- prefer 3 apply (simp add: compare_rls hcomplex_add_commute)
-apply (rule_tac [2] z = x in eq_Abs_star)
-apply (rule_tac [4] z = x in eq_Abs_star)
-apply (auto simp add: starfunC hcomplex_of_complex_def hcomplex_minus hcomplex_add star_of_def star_n_def)
+ prefer 3 apply (simp add: compare_rls add_commute)
+apply (rule_tac [2] x = x in star_cases)
+apply (rule_tac [4] x = x in star_cases)
+apply (auto simp add: starfun star_n_minus star_n_add star_of_def)
 done
 
 lemma NSCLIM_isContc_iff:
@@ -507,13 +502,13 @@
 
 lemma isContc_mult:
      "[| isContc f a; isContc g a |] ==> isContc (%x. f(x) * g(x)) a"
-by (auto intro!: starfunC_mult_CFinite_capprox 
-            simp del: starfunC_mult [symmetric]
+by (auto intro!: starfun_mult_CFinite_capprox
+            [simplified starfun_mult [symmetric]]
             simp add: isNSContc_isContc_iff [symmetric] isNSContc_def)
 
 
 lemma isContc_o: "[| isContc f a; isContc g (f a) |] ==> isContc (g o f) a"
-by (simp add: isNSContc_isContc_iff [symmetric] isNSContc_def starfunC_o [symmetric])
+by (simp add: isNSContc_isContc_iff [symmetric] isNSContc_def starfun_o [symmetric])
 
 lemma isContc_o2:
      "[| isContc f a; isContc g (f a) |] ==> isContc (%x. g (f x)) a"
@@ -550,7 +545,7 @@
 
 lemma isNSContCRD:
       "[| isNSContCR f a; y @c= hcomplex_of_complex a |]
-       ==> ( *fcR* f) y @= hypreal_of_real (f a)"
+       ==> ( *f* f) y @= hypreal_of_real (f a)"
 by (simp add: isNSContCR_def)
 
 lemma isNSContCR_NSCRLIM: "isNSContCR f a ==> f -- a --NSCR> (f a) "
@@ -651,7 +646,7 @@
 apply (drule_tac x = xa in bspec)
 apply (rule_tac [3] ccontr)
 apply (drule_tac [3] x = h in spec)
-apply (auto simp add: mem_cinfmal_iff starfunC_lambda_cancel)
+apply (auto simp add: mem_cinfmal_iff starfun_lambda_cancel)
 done
 
 (*** 2nd equivalence ***)
@@ -662,7 +657,7 @@
 lemma NSCDERIV_iff2:
      "(NSCDERIV f x :> D) =
       (\<forall>xa. xa \<noteq> hcomplex_of_complex x & xa @c= hcomplex_of_complex x -->
-        ( *fc* (%z. (f z - f x) / (z - x))) xa @c= hcomplex_of_complex D)"
+        ( *f* (%z. (f z - f x) / (z - x))) xa @c= hcomplex_of_complex D)"
 by (simp add: NSCDERIV_NSCLIM_iff2 NSCLIM_def)
 
 lemma NSCDERIV_CDERIV_iff: "(NSCDERIV f x :> D) = (CDERIV f x :> D)"
@@ -734,7 +729,7 @@
       ==> NSCDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))"
 apply (simp add: NSCDERIV_NSCLIM_iff NSCLIM_def, clarify) 
 apply (auto dest!: spec
-            simp add: starfunC_lambda_cancel lemma_nscderiv1)
+            simp add: starfun_lambda_cancel lemma_nscderiv1)
 apply (simp (no_asm) add: add_divide_distrib)
 apply (drule bex_CInfinitesimal_iff2 [THEN iffD2])+
 apply (auto simp del: times_divide_eq_right simp add: times_divide_eq_right [symmetric])
@@ -803,7 +798,7 @@
 (* lemmas *)
 lemma NSCDERIV_zero:
       "[| NSCDERIV g x :> D;
-          ( *fc* g) (hcomplex_of_complex(x) + xa) = hcomplex_of_complex(g x);
+          ( *f* g) (hcomplex_of_complex(x) + xa) = hcomplex_of_complex(g x);
           xa : CInfinitesimal; xa \<noteq> 0
        |] ==> D = 0"
 apply (simp add: nscderiv_def)
@@ -812,7 +807,7 @@
 
 lemma NSCDERIV_capprox:
   "[| NSCDERIV f x :> D;  h: CInfinitesimal;  h \<noteq> 0 |]
-   ==> ( *fc* f) (hcomplex_of_complex(x) + h) - hcomplex_of_complex(f x) @c= 0"
+   ==> ( *f* f) (hcomplex_of_complex(x) + h) - hcomplex_of_complex(f x) @c= 0"
 apply (simp add: nscderiv_def mem_cinfmal_iff [symmetric])
 apply (rule CInfinitesimal_ratio)
 apply (rule_tac [3] capprox_hcomplex_of_complex_CFinite, auto)
@@ -829,11 +824,11 @@
 
 lemma NSCDERIVD1:
    "[| NSCDERIV f (g x) :> Da;
-       ( *fc* g) (hcomplex_of_complex(x) + xa) \<noteq> hcomplex_of_complex (g x);
-       ( *fc* g) (hcomplex_of_complex(x) + xa) @c= hcomplex_of_complex (g x)|]
-    ==> (( *fc* f) (( *fc* g) (hcomplex_of_complex(x) + xa))
+       ( *f* g) (hcomplex_of_complex(x) + xa) \<noteq> hcomplex_of_complex (g x);
+       ( *f* g) (hcomplex_of_complex(x) + xa) @c= hcomplex_of_complex (g x)|]
+    ==> (( *f* f) (( *f* g) (hcomplex_of_complex(x) + xa))
 	 - hcomplex_of_complex (f (g x))) /
-	(( *fc* g) (hcomplex_of_complex(x) + xa) - hcomplex_of_complex (g x))
+	(( *f* g) (hcomplex_of_complex(x) + xa) - hcomplex_of_complex (g x))
 	   @c= hcomplex_of_complex (Da)"
 by (simp add: NSCDERIV_NSCLIM_iff2 NSCLIM_def)
 
@@ -847,9 +842,9 @@
 
 lemma NSCDERIVD2:
   "[| NSCDERIV g x :> Db; xa: CInfinitesimal; xa \<noteq> 0 |]
-   ==> (( *fc* g) (hcomplex_of_complex x + xa) - hcomplex_of_complex(g x)) / xa
+   ==> (( *f* g) (hcomplex_of_complex x + xa) - hcomplex_of_complex(g x)) / xa
        @c= hcomplex_of_complex (Db)"
-by (simp add: NSCDERIV_NSCLIM_iff NSCLIM_def mem_cinfmal_iff starfunC_lambda_cancel)
+by (simp add: NSCDERIV_NSCLIM_iff NSCLIM_def mem_cinfmal_iff starfun_lambda_cancel)
 
 lemma lemma_complex_chain: "(z::hcomplex) \<noteq> 0 ==> x*y = (x*inverse(z))*(z*y)"
 by auto
@@ -862,11 +857,11 @@
 apply (simp (no_asm_simp) add: NSCDERIV_NSCLIM_iff NSCLIM_def mem_cinfmal_iff [symmetric])
 apply safe
 apply (frule_tac f = g in NSCDERIV_capprox)
-apply (auto simp add: starfunC_lambda_cancel2 starfunC_o [symmetric])
-apply (case_tac "( *fc* g) (hcomplex_of_complex (x) + xa) = hcomplex_of_complex (g x) ")
+apply (auto simp add: starfun_lambda_cancel2 starfun_o [symmetric])
+apply (case_tac "( *f* g) (hcomplex_of_complex (x) + xa) = hcomplex_of_complex (g x) ")
 apply (drule_tac g = g in NSCDERIV_zero)
 apply (auto simp add: divide_inverse)
-apply (rule_tac z1 = "( *fc* g) (hcomplex_of_complex (x) + xa) - hcomplex_of_complex (g x) " and y1 = "inverse xa" in lemma_complex_chain [THEN ssubst])
+apply (rule_tac z1 = "( *f* g) (hcomplex_of_complex (x) + xa) - hcomplex_of_complex (g x) " and y1 = "inverse xa" in lemma_complex_chain [THEN ssubst])
 apply (simp (no_asm_simp))
 apply (rule capprox_mult_hcomplex_of_complex)
 apply (auto intro!: NSCDERIVD1 intro: capprox_minus_iff [THEN iffD2] 
@@ -933,9 +928,9 @@
      "x \<noteq> 0 ==> NSCDERIV (%x. inverse(x)) x :> (- (inverse x ^ 2))"
 apply (simp add: nscderiv_def Ball_def, clarify) 
 apply (frule CInfinitesimal_add_not_zero [where x=x])
-apply (auto simp add: starfunC_inverse_inverse diff_minus 
+apply (auto simp add: starfun_inverse_inverse diff_minus 
            simp del: minus_mult_left [symmetric] minus_mult_right [symmetric])
-apply (simp add: hcomplex_add_commute numeral_2_eq_2 inverse_add
+apply (simp add: add_commute numeral_2_eq_2 inverse_add
                  inverse_mult_distrib [symmetric] inverse_minus_eq [symmetric]
                  add_ac mult_ac 
             del: inverse_minus_eq inverse_mult_distrib
--- a/src/HOL/Complex/CStar.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Complex/CStar.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -10,706 +10,104 @@
 imports NSCA
 begin
 
-constdefs
-
-    (* nonstandard extension of sets *)
-    starsetC :: "complex set => hcomplex set"          ("*sc* _" [80] 80)
-    "*sc* A  == {x. \<forall>X \<in> Rep_star(x). {n. X n \<in> A} \<in> FreeUltrafilterNat}"
-
-    (* internal sets *)
-    starsetC_n :: "(nat => complex set) => hcomplex set"  ("*scn* _" [80] 80)
-    "*scn* As  == {x. \<forall>X \<in> Rep_star(x). 
-                      {n. X n \<in> (As n)} \<in> FreeUltrafilterNat}"
-
-    InternalCSets :: "hcomplex set set"
-    "InternalCSets == {X. \<exists>As. X = *scn* As}"
-
-    (* star transform of functions f: Complex --> Complex *)
-
-    starfunC :: "(complex => complex) => hcomplex => hcomplex"
-                ("*fc* _" [80] 80)
-    "*fc* f  == 
-        (%x. Abs_star(\<Union>X \<in> Rep_star(x). starrel``{%n. f (X n)}))"
-
-    starfunC_n :: "(nat => (complex => complex)) => hcomplex => hcomplex"
-                  ("*fcn* _" [80] 80)
-    "*fcn* F  == 
-      (%x. Abs_star(\<Union>X \<in> Rep_star(x). starrel``{%n. (F n)(X n)}))"
-
-    InternalCFuns :: "(hcomplex => hcomplex) set"
-    "InternalCFuns == {X. \<exists>F. X = *fcn* F}"
-
-
-    (* star transform of functions f: Real --> Complex *)
-
-    starfunRC :: "(real => complex) => hypreal => hcomplex"
-                 ("*fRc* _" [80] 80)
-    "*fRc* f  == (%x. Abs_star(\<Union>X \<in> Rep_star(x). starrel``{%n. f (X n)}))"
-
-    starfunRC_n :: "(nat => (real => complex)) => hypreal => hcomplex"
-                   ("*fRcn* _" [80] 80)
-    "*fRcn* F  == (%x. Abs_star(\<Union>X \<in> Rep_star(x). starrel``{%n. (F n)(X n)}))"
-
-    InternalRCFuns :: "(hypreal => hcomplex) set"
-    "InternalRCFuns == {X. \<exists>F. X = *fRcn* F}"
-
-    (* star transform of functions f: Complex --> Real; needed for Re and Im parts *)
-
-    starfunCR :: "(complex => real) => hcomplex => hypreal"
-                 ("*fcR* _" [80] 80)
-    "*fcR* f  == (%x. Abs_star(\<Union>X \<in> Rep_star(x). starrel``{%n. f (X n)}))"
-
-    starfunCR_n :: "(nat => (complex => real)) => hcomplex => hypreal"
-                   ("*fcRn* _" [80] 80)
-    "*fcRn* F  == (%x. Abs_star(\<Union>X \<in> Rep_star(x). starrel``{%n. (F n)(X n)}))"
-
-    InternalCRFuns :: "(hcomplex => hypreal) set"
-    "InternalCRFuns == {X. \<exists>F. X = *fcRn* F}"
-
-
 subsection{*Properties of the *-Transform Applied to Sets of Reals*}
 
-lemma STARC_complex_set [simp]: "*sc*(UNIV::complex set) = (UNIV)"
-by (simp add: starsetC_def)
-declare STARC_complex_set
-
-lemma STARC_empty_set: "*sc* {} = {}"
-by (simp add: starsetC_def)
-declare STARC_empty_set [simp]
-
-lemma STARC_Un: "*sc* (A Un B) = *sc* A Un *sc* B"
-apply (auto simp add: starsetC_def)
-apply (drule bspec, assumption)
-apply (rule_tac z = x in eq_Abs_star, simp, ultra)
-apply (blast intro: FreeUltrafilterNat_subset)+
-done
-
-lemma starsetC_n_Un: "*scn* (%n. (A n) Un (B n)) = *scn* A Un *scn* B"
-apply (auto simp add: starsetC_n_def)
-apply (drule_tac x = Xa in bspec)
-apply (rule_tac [2] z = x in eq_Abs_star)
-apply (auto dest!: bspec, ultra+)
-done
-
-lemma InternalCSets_Un:
-     "[| X \<in> InternalCSets; Y \<in> InternalCSets |] ==> (X Un Y) \<in> InternalCSets"
-by (auto simp add:  InternalCSets_def starsetC_n_Un [symmetric])
-
-lemma STARC_Int: "*sc* (A Int B) = *sc* A Int *sc* B"
-apply (auto simp add: starsetC_def)
-prefer 3 apply (blast intro: FreeUltrafilterNat_Int FreeUltrafilterNat_subset)
-apply (blast intro: FreeUltrafilterNat_subset)+
-done
-
-lemma starsetC_n_Int: "*scn* (%n. (A n) Int (B n)) = *scn* A Int *scn* B"
-apply (auto simp add: starsetC_n_def)
-apply (auto dest!: bspec, ultra+)
-done
-
-lemma InternalCSets_Int:
-    "[| X \<in> InternalCSets; Y \<in> InternalCSets |] ==> (X Int Y) \<in> InternalCSets"
-by (auto simp add: InternalCSets_def starsetC_n_Int [symmetric])
-
-lemma STARC_Compl: "*sc* -A = -( *sc* A)"
-apply (auto simp add: starsetC_def)
-apply (rule_tac z = x in eq_Abs_star)
-apply (rule_tac [2] z = x in eq_Abs_star)
-apply (auto dest!: bspec, ultra+)
-done
-
-lemma starsetC_n_Compl: "*scn* ((%n. - A n)) = -( *scn* A)"
-apply (auto simp add: starsetC_n_def)
-apply (rule_tac z = x in eq_Abs_star)
-apply (rule_tac [2] z = x in eq_Abs_star)
-apply (auto dest!: bspec, ultra+)
-done
-
-lemma InternalCSets_Compl: "X :InternalCSets ==> -X \<in> InternalCSets"
-by (auto simp add: InternalCSets_def starsetC_n_Compl [symmetric])
-
-lemma STARC_mem_Compl: "x \<notin> *sc* F ==> x \<in> *sc* (- F)"
-by (simp add: STARC_Compl)
-
-lemma STARC_diff: "*sc* (A - B) = *sc* A - *sc* B"
-by (simp add: Diff_eq STARC_Int STARC_Compl)
-
-lemma starsetC_n_diff:
-      "*scn* (%n. (A n) - (B n)) = *scn* A - *scn* B"
-apply (auto simp add: starsetC_n_def)
-apply (rule_tac [2] z = x in eq_Abs_star)
-apply (rule_tac [3] z = x in eq_Abs_star)
-apply (auto dest!: bspec, ultra+)
-done
-
-lemma InternalCSets_diff:
-     "[| X \<in> InternalCSets; Y \<in> InternalCSets |] ==> (X - Y) \<in> InternalCSets"
-by (auto simp add: InternalCSets_def starsetC_n_diff [symmetric])
-
-lemma STARC_subset: "A \<le> B ==> *sc* A \<le> *sc* B"
-apply (simp add: starsetC_def)
-apply (blast intro: FreeUltrafilterNat_subset)+
-done
-
-lemma STARC_mem: "a \<in> A ==> hcomplex_of_complex a \<in> *sc* A"
-apply (simp add: starsetC_def hcomplex_of_complex_def star_of_def star_n_def)
-apply (auto intro: FreeUltrafilterNat_subset)
-done
-
-lemma STARC_hcomplex_of_complex_image_subset:
-    "hcomplex_of_complex ` A \<le> *sc* A"
-apply (auto simp add: starsetC_def hcomplex_of_complex_def star_of_def star_n_def)
-apply (blast intro: FreeUltrafilterNat_subset)
-done
-
-lemma STARC_SComplex_subset: "SComplex \<le> *sc* (UNIV:: complex set)"
-by auto
+lemma STARC_SComplex_subset: "SComplex \<subseteq> *s* (UNIV:: complex set)"
+by simp
 
 lemma STARC_hcomplex_of_complex_Int:
-     "*sc* X Int SComplex = hcomplex_of_complex ` X"
-apply (auto simp add: starsetC_def hcomplex_of_complex_def SComplex_def star_of_def star_n_def)
-apply (fold star_n_def star_of_def hcomplex_of_complex_def)
-apply (rule imageI, rule ccontr)
-apply (drule bspec)
-apply (rule lemma_starrel_refl)
-prefer 2 apply (blast intro: FreeUltrafilterNat_subset, auto)
-done
+     "*s* X Int SComplex = hcomplex_of_complex ` X"
+by (auto simp add: SComplex_def STAR_mem_iff)
 
 lemma lemma_not_hcomplexA:
      "x \<notin> hcomplex_of_complex ` A ==> \<forall>y \<in> A. x \<noteq> hcomplex_of_complex y"
 by auto
 
-lemma starsetC_starsetC_n_eq: "*sc* X = *scn* (%n. X)"
-by (simp add: starsetC_n_def starsetC_def)
-
-lemma InternalCSets_starsetC_n [simp]: "( *sc* X) \<in> InternalCSets"
-by (auto simp add: InternalCSets_def starsetC_starsetC_n_eq)
-
-lemma InternalCSets_UNIV_diff:
-    "X \<in> InternalCSets ==> UNIV - X \<in> InternalCSets"
-apply (subgoal_tac "UNIV - X = - X")
-by (auto intro: InternalCSets_Compl)
-
-text{*Nonstandard extension of a set (defined using a constant sequence) as a special case of an internal set*}
-
-lemma starsetC_n_starsetC: "\<forall>n. (As n = A) ==> *scn* As = *sc* A"
-by (simp add:starsetC_n_def starsetC_def)
-
-
 subsection{*Theorems about Nonstandard Extensions of Functions*}
 
-lemma starfunC_n_starfunC: "\<forall>n. (F n = f) ==> *fcn* F = *fc* f"
-by (simp add: starfunC_n_def starfunC_def)
-
-lemma starfunRC_n_starfunRC: "\<forall>n. (F n = f) ==> *fRcn* F = *fRc* f"
-by (simp add: starfunRC_n_def starfunRC_def)
-
-lemma starfunCR_n_starfunCR: "\<forall>n. (F n = f) ==> *fcRn* F = *fcR* f"
-by (simp add: starfunCR_n_def starfunCR_def)
-
-lemma starfunC_congruent:
-      "(%X. starrel``{%n. f (X n)}) respects starrel"
-by (auto simp add: starrel_iff congruent_def, ultra)
-
-(* f::complex => complex *)
-lemma starfunC:
-      "( *fc* f) (Abs_star(starrel``{%n. X n})) =
-       Abs_star(starrel `` {%n. f (X n)})"
-apply (simp add: starfunC_def)
-apply (rule arg_cong [where f = Abs_star])
-apply (auto iff add: starrel_iff, ultra)
-done
-
 lemma cstarfun_if_eq:
      "w \<noteq> hcomplex_of_complex x
-       ==> ( *fc* (\<lambda>z. if z = x then a else g z)) w = ( *fc* g) w" 
-apply (rule_tac z=w in eq_Abs_star)
-apply (simp add: hcomplex_of_complex_def star_of_def star_n_def starfunC, ultra)
-done
-
-lemma starfunRC:
-      "( *fRc* f) (Abs_star(starrel``{%n. X n})) =
-       Abs_star(starrel `` {%n. f (X n)})"
-apply (simp add: starfunRC_def)
-apply (rule arg_cong [where f = Abs_star], auto, ultra)
-done
-
-lemma starfunCR:
-      "( *fcR* f) (Abs_star(starrel``{%n. X n})) =
-       Abs_star(starrel `` {%n. f (X n)})"
-apply (simp add: starfunCR_def)
-apply (rule arg_cong [where f = Abs_star])
-apply (auto iff add: starrel_iff, ultra)
-done
-
-(**  multiplication: ( *f) x ( *g) = *(f x g) **)
-
-lemma starfunC_mult: "( *fc* f) z * ( *fc* g) z = ( *fc* (%x. f x * g x)) z"
-apply (rule_tac z = z in eq_Abs_star)
-apply (auto simp add: starfunC hcomplex_mult)
-done
-declare starfunC_mult [symmetric, simp]
-
-lemma starfunRC_mult:
-    "( *fRc* f) z * ( *fRc* g) z = ( *fRc* (%x. f x * g x)) z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunRC hcomplex_mult)
-done
-declare starfunRC_mult [symmetric, simp]
-
-lemma starfunCR_mult:
-    "( *fcR* f) z * ( *fcR* g) z = ( *fcR* (%x. f x * g x)) z"
-apply (rule_tac z = z in eq_Abs_star)
-apply (simp add: starfunCR hypreal_mult)
-done
-declare starfunCR_mult [symmetric, simp]
-
-(**  addition: ( *f) + ( *g) = *(f + g)  **)
-
-lemma starfunC_add: "( *fc* f) z + ( *fc* g) z = ( *fc* (%x. f x + g x)) z"
-apply (rule_tac z = z in eq_Abs_star)
-apply (simp add: starfunC hcomplex_add)
-done
-declare starfunC_add [symmetric, simp]
-
-lemma starfunRC_add: "( *fRc* f) z + ( *fRc* g) z = ( *fRc* (%x. f x + g x)) z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunRC hcomplex_add)
-done
-declare starfunRC_add [symmetric, simp]
-
-lemma starfunCR_add: "( *fcR* f) z + ( *fcR* g) z = ( *fcR* (%x. f x + g x)) z"
-apply (rule_tac z = z in eq_Abs_star)
-apply (simp add: starfunCR hypreal_add)
-done
-declare starfunCR_add [symmetric, simp]
-
-(**  uminus **)
-lemma starfunC_minus [simp]: "( *fc* (%x. - f x)) x = - ( *fc* f) x"
-apply (rule_tac z = x in eq_Abs_star)
-apply (simp add: starfunC hcomplex_minus)
-done
-
-lemma starfunRC_minus [simp]: "( *fRc* (%x. - f x)) x = - ( *fRc* f) x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfunRC hcomplex_minus)
-done
-
-lemma starfunCR_minus [simp]: "( *fcR* (%x. - f x)) x = - ( *fcR* f) x"
-apply (rule_tac z = x in eq_Abs_star)
-apply (simp add: starfunCR hypreal_minus)
+       ==> ( *f* (\<lambda>z. if z = x then a else g z)) w = ( *f* g) w"
+apply (cases w)
+apply (simp add: star_of_def starfun star_n_eq_iff, ultra)
 done
 
-(**  addition: ( *f) - ( *g) = *(f - g)  **)
-
-lemma starfunC_diff: "( *fc* f) y  - ( *fc* g) y = ( *fc* (%x. f x - g x)) y"
-by (simp add: diff_minus)
-declare starfunC_diff [symmetric, simp]
-
-lemma starfunRC_diff:
-    "( *fRc* f) y  - ( *fRc* g) y = ( *fRc* (%x. f x - g x)) y"
-by (simp add: diff_minus)
-declare starfunRC_diff [symmetric, simp]
-
-lemma starfunCR_diff:
-  "( *fcR* f) y  - ( *fcR* g) y = ( *fcR* (%x. f x - g x)) y"
-by (simp add: diff_minus)
-declare starfunCR_diff [symmetric, simp]
-
-(**  composition: ( *f) o ( *g) = *(f o g) **)
-
-lemma starfunC_o2: "(%x. ( *fc* f) (( *fc* g) x)) = *fc* (%x. f (g x))"
-apply (rule ext)
-apply (rule_tac z = x in eq_Abs_star)
-apply (simp add: starfunC)
-done
-
-lemma starfunC_o: "( *fc* f) o ( *fc* g) = ( *fc* (f o g))"
-by (simp add: o_def starfunC_o2)
-
-lemma starfunC_starfunRC_o2:
-    "(%x. ( *fc* f) (( *fRc* g) x)) = *fRc* (%x. f (g x))"
-apply (rule ext)
-apply (rule_tac z = x in eq_Abs_star)
-apply (simp add: starfunRC starfunC)
-done
-
-lemma starfun_starfunCR_o2:
-    "(%x. ( *f* f) (( *fcR* g) x)) = *fcR* (%x. f (g x))"
-apply (rule ext)
-apply (rule_tac z = x in eq_Abs_star)
-apply (simp add: starfunCR starfun)
-done
-
-lemma starfunC_starfunRC_o: "( *fc* f) o ( *fRc* g) = ( *fRc* (f o g))"
-by (simp add: o_def starfunC_starfunRC_o2)
-
-lemma starfun_starfunCR_o: "( *f* f) o ( *fcR* g) = ( *fcR* (f o g))"
-by (simp add: o_def starfun_starfunCR_o2)
-
-lemma starfunC_const_fun [simp]: "( *fc* (%x. k)) z = hcomplex_of_complex k"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunC hcomplex_of_complex_def star_of_def star_n_def)
-done
-
-lemma starfunRC_const_fun [simp]: "( *fRc* (%x. k)) z = hcomplex_of_complex k"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunRC hcomplex_of_complex_def star_of_def star_n_def)
-done
-
-lemma starfunCR_const_fun [simp]: "( *fcR* (%x. k)) z = hypreal_of_real k"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunCR hypreal_of_real_def star_of_def star_n_def)
-done
-
-lemma starfunC_inverse: "inverse (( *fc* f) x) = ( *fc* (%x. inverse (f x))) x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfunC hcomplex_inverse)
-done
-declare starfunC_inverse [symmetric, simp]
-
-lemma starfunRC_inverse:
-    "inverse (( *fRc* f) x) = ( *fRc* (%x. inverse (f x))) x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfunRC hcomplex_inverse)
-done
-declare starfunRC_inverse [symmetric, simp]
-
-lemma starfunCR_inverse:
-    "inverse (( *fcR* f) x) = ( *fcR* (%x. inverse (f x))) x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfunCR hypreal_inverse)
-done
-declare starfunCR_inverse [symmetric, simp]
-
-lemma starfunC_eq [simp]:
-    "( *fc* f) (hcomplex_of_complex a) = hcomplex_of_complex (f a)"
-by (simp add: starfunC hcomplex_of_complex_def star_of_def star_n_def)
-
-lemma starfunRC_eq [simp]:
-    "( *fRc* f) (hypreal_of_real a) = hcomplex_of_complex (f a)"
-by (simp add: starfunRC hcomplex_of_complex_def hypreal_of_real_def star_of_def star_n_def)
-
-lemma starfunCR_eq [simp]:
-    "( *fcR* f) (hcomplex_of_complex a) = hypreal_of_real (f a)"
-by (simp add: starfunCR hcomplex_of_complex_def hypreal_of_real_def star_of_def star_n_def)
-
-lemma starfunC_capprox:
-    "( *fc* f) (hcomplex_of_complex a) @c= hcomplex_of_complex (f a)"
-by auto
-
-lemma starfunRC_capprox:
-    "( *fRc* f) (hypreal_of_real a) @c= hcomplex_of_complex (f a)"
-by auto
-
-lemma starfunCR_approx:
-    "( *fcR* f) (hcomplex_of_complex a) @= hypreal_of_real (f a)"
+lemma starfun_capprox:
+    "( *f* f) (hcomplex_of_complex a) @c= hcomplex_of_complex (f a)"
 by auto
 
 (*
-Goal "( *fcNat* (%n. z ^ n)) N = (hcomplex_of_complex z) hcpow N"
+Goal "( *fNat* (%n. z ^ n)) N = (hcomplex_of_complex z) hcpow N"
 *)
 
-lemma starfunC_hcpow: "( *fc* (%z. z ^ n)) Z = Z hcpow hypnat_of_nat n"
-apply (rule_tac z=Z in eq_Abs_star)
-apply (simp add: hcpow starfunC hypnat_of_nat_eq)
-done
-
-lemma starfunC_lambda_cancel:
-    "( *fc* (%h. f (x + h))) y  = ( *fc* f) (hcomplex_of_complex  x + y)"
-apply (rule_tac z=y in eq_Abs_star)
-apply (simp add: starfunC hcomplex_of_complex_def hcomplex_add star_of_def star_n_def)
-done
-
-lemma starfunCR_lambda_cancel:
-    "( *fcR* (%h. f (x + h))) y  = ( *fcR* f) (hcomplex_of_complex  x + y)"
-apply (rule_tac z=y in eq_Abs_star)
-apply (simp add: starfunCR hcomplex_of_complex_def hcomplex_add star_of_def star_n_def)
+lemma starfunC_hcpow: "( *f* (%z. z ^ n)) Z = Z hcpow hypnat_of_nat n"
+apply (cases Z)
+apply (simp add: hcpow starfun hypnat_of_nat_eq)
 done
 
-lemma starfunRC_lambda_cancel:
-    "( *fRc* (%h. f (x + h))) y  = ( *fRc* f) (hypreal_of_real x + y)"
-apply (rule_tac z=y in eq_Abs_star)
-apply (simp add: starfunRC hypreal_of_real_def star_of_def star_n_def hypreal_add)
-done
-
-lemma starfunC_lambda_cancel2:
-    "( *fc* (%h. f(g(x + h)))) y = ( *fc* (f o g)) (hcomplex_of_complex x + y)"
-apply (rule_tac z=y in eq_Abs_star)
-apply (simp add: starfunC hcomplex_of_complex_def hcomplex_add star_of_def star_n_def)
-done
-
-lemma starfunCR_lambda_cancel2:
-    "( *fcR* (%h. f(g(x + h)))) y = ( *fcR* (f o g)) (hcomplex_of_complex x + y)"
-apply (rule_tac z=y in eq_Abs_star)
-apply (simp add: starfunCR hcomplex_of_complex_def hcomplex_add star_of_def star_n_def)
-done
-
-lemma starfunRC_lambda_cancel2:
-    "( *fRc* (%h. f(g(x + h)))) y = ( *fRc* (f o g)) (hypreal_of_real x + y)"
-apply (rule_tac z=y in eq_Abs_star)
-apply (simp add: starfunRC hypreal_of_real_def star_of_def star_n_def hypreal_add)
-done
-
-lemma starfunC_mult_CFinite_capprox:
-    "[| ( *fc* f) y @c= l; ( *fc* g) y @c= m; l: CFinite; m: CFinite |]
-     ==>  ( *fc* (%x. f x * g x)) y @c= l * m"
+lemma starfun_mult_CFinite_capprox:
+    "[| ( *f* f) y @c= l; ( *f* g) y @c= m; l: CFinite; m: CFinite |]
+     ==>  ( *f* (%x. f x * g x)) y @c= l * m"
 apply (drule capprox_mult_CFinite, assumption+)
 apply (auto intro: capprox_sym [THEN [2] capprox_CFinite])
 done
 
-lemma starfunCR_mult_HFinite_capprox:
-    "[| ( *fcR* f) y @= l; ( *fcR* g) y @= m; l: HFinite; m: HFinite |]
-     ==>  ( *fcR* (%x. f x * g x)) y @= l * m"
-apply (drule approx_mult_HFinite, assumption+)
-apply (auto intro: approx_sym [THEN [2] approx_HFinite])
-done
-
-lemma starfunRC_mult_CFinite_capprox:
-    "[| ( *fRc* f) y @c= l; ( *fRc* g) y @c= m; l: CFinite; m: CFinite |]
-     ==>  ( *fRc* (%x. f x * g x)) y @c= l * m"
-apply (drule capprox_mult_CFinite, assumption+)
-apply (auto intro: capprox_sym [THEN [2] capprox_CFinite])
-done
-
-lemma starfunC_add_capprox:
-    "[| ( *fc* f) y @c= l; ( *fc* g) y @c= m |]
-     ==>  ( *fc* (%x. f x + g x)) y @c= l + m"
+lemma starfun_add_capprox:
+    "[| ( *f* f) y @c= l; ( *f* g) y @c= m |]
+     ==>  ( *f* (%x. f x + g x)) y @c= l + m"
 by (auto intro: capprox_add)
 
-lemma starfunRC_add_capprox:
-    "[| ( *fRc* f) y @c= l; ( *fRc* g) y @c= m |]
-     ==>  ( *fRc* (%x. f x + g x)) y @c= l + m"
-by (auto intro: capprox_add)
-
-lemma starfunCR_add_approx:
-    "[| ( *fcR* f) y @= l; ( *fcR* g) y @= m
-               |] ==>  ( *fcR* (%x. f x + g x)) y @= l + m"
-by (auto intro: approx_add)
-
-lemma starfunCR_cmod: "*fcR* cmod = hcmod"
+lemma starfunCR_cmod: "*f* cmod = hcmod"
 apply (rule ext)
-apply (rule_tac z = x in eq_Abs_star)
-apply (simp add: starfunCR hcmod)
-done
-
-lemma starfunC_inverse_inverse: "( *fc* inverse) x = inverse(x)"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfunC hcomplex_inverse)
+apply (rule_tac x = x in star_cases)
+apply (simp add: starfun hcmod)
 done
 
-lemma starfunC_divide: "( *fc* f) y  / ( *fc* g) y = ( *fc* (%x. f x / g x)) y"
-by (simp add: divide_inverse)
-declare starfunC_divide [symmetric, simp]
-
-lemma starfunCR_divide:
-  "( *fcR* f) y  / ( *fcR* g) y = ( *fcR* (%x. f x / g x)) y"
-by (simp add: divide_inverse)
-declare starfunCR_divide [symmetric, simp]
-
-lemma starfunRC_divide:
-  "( *fRc* f) y  / ( *fRc* g) y = ( *fRc* (%x. f x / g x)) y"
-by (simp add: divide_inverse)
-declare starfunRC_divide [symmetric, simp]
-
-
-subsection{*Internal Functions - Some Redundancy With *Fc* Now*}
-
-lemma starfunC_n_congruent:
-      "(%X. starrel``{%n. f n (X n)}) respects starrel"
-by (auto simp add: congruent_def starrel_iff, ultra)
-
-lemma starfunC_n:
-     "( *fcn* f) (Abs_star(starrel``{%n. X n})) =
-      Abs_star(starrel `` {%n. f n (X n)})"
-apply (simp add: starfunC_n_def)
-apply (rule arg_cong [where f = Abs_star])
-apply (auto iff add: starrel_iff, ultra)
-done
-
-(**  multiplication: ( *fn) x ( *gn) = *(fn x gn) **)
-
-lemma starfunC_n_mult:
-    "( *fcn* f) z * ( *fcn* g) z = ( *fcn* (% i x. f i x * g i x)) z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunC_n hcomplex_mult)
-done
-
-(**  addition: ( *fn) + ( *gn) = *(fn + gn) **)
-
-lemma starfunC_n_add:
-    "( *fcn* f) z + ( *fcn* g) z = ( *fcn* (%i x. f i x + g i x)) z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunC_n hcomplex_add)
-done
-
-(** uminus **)
-
-lemma starfunC_n_minus: "- ( *fcn* g) z = ( *fcn* (%i x. - g i x)) z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunC_n hcomplex_minus)
-done
+subsection{*Internal Functions - Some Redundancy With *f* Now*}
 
 (** subtraction: ( *fn) - ( *gn) = *(fn - gn) **)
 
-lemma starfunNat_n_diff:
-   "( *fcn* f) z - ( *fcn* g) z = ( *fcn* (%i x. f i x - g i x)) z"
-by (simp add: diff_minus  starfunC_n_add starfunC_n_minus)
+lemma starfun_n_diff:
+   "( *fn* f) z - ( *fn* g) z = ( *fn* (%i x. f i x - g i x)) z"
+apply (cases z)
+apply (simp add: starfun_n star_n_diff)
+done
 
 (** composition: ( *fn) o ( *gn) = *(fn o gn) **)
 
-lemma starfunC_n_const_fun [simp]:
-     "( *fcn* (%i x. k)) z = hcomplex_of_complex  k"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunC_n hcomplex_of_complex_def star_of_def star_n_def)
-done
-
-lemma starfunC_n_eq [simp]:
-    "( *fcn* f) (hcomplex_of_complex n) = Abs_star(starrel `` {%i. f i n})"
-by (simp add: starfunC_n hcomplex_of_complex_def star_of_def star_n_def)
-
-lemma starfunC_eq_iff: "(( *fc* f) = ( *fc* g)) = (f = g)"
-apply auto
-apply (rule ext, rule ccontr)
-apply (drule_tac x = "hcomplex_of_complex (x) " in fun_cong)
-apply (simp add: starfunC hcomplex_of_complex_def star_of_def star_n_def)
-done
-
-lemma starfunRC_eq_iff: "(( *fRc* f) = ( *fRc* g)) = (f = g)"
-apply auto
-apply (rule ext, rule ccontr)
-apply (drule_tac x = "hypreal_of_real (x) " in fun_cong)
-apply auto
-done
-
-lemma starfunCR_eq_iff: "(( *fcR* f) = ( *fcR* g)) = (f = g)"
-apply auto
-apply (rule ext, rule ccontr)
-apply (drule_tac x = "hcomplex_of_complex (x) " in fun_cong)
-apply auto
-done
-
 lemma starfunC_eq_Re_Im_iff:
-    "(( *fc* f) x = z) = ((( *fcR* (%x. Re(f x))) x = hRe (z)) &
-                          (( *fcR* (%x. Im(f x))) x = hIm (z)))"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=z in eq_Abs_star)
-apply (auto simp add: starfunCR starfunC hIm hRe complex_Re_Im_cancel_iff, ultra+)
+    "(( *f* f) x = z) = ((( *f* (%x. Re(f x))) x = hRe (z)) &
+                          (( *f* (%x. Im(f x))) x = hIm (z)))"
+apply (cases x, cases z)
+apply (auto simp add: starfun hIm hRe complex_Re_Im_cancel_iff star_n_eq_iff, ultra+)
 done
 
 lemma starfunC_approx_Re_Im_iff:
-    "(( *fc* f) x @c= z) = ((( *fcR* (%x. Re(f x))) x @= hRe (z)) &
-                            (( *fcR* (%x. Im(f x))) x @= hIm (z)))"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunCR starfunC hIm hRe capprox_approx_iff)
+    "(( *f* f) x @c= z) = ((( *f* (%x. Re(f x))) x @= hRe (z)) &
+                            (( *f* (%x. Im(f x))) x @= hIm (z)))"
+apply (cases x, cases z)
+apply (simp add: starfun hIm hRe capprox_approx_iff)
 done
 
 lemma starfunC_Idfun_capprox:
-    "x @c= hcomplex_of_complex a ==> ( *fc* (%x. x)) x @c= hcomplex_of_complex  a"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfunC)
-done
-
-lemma starfunC_Id [simp]: "( *fc* (%x. x)) x = x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfunC)
-done
+    "x @c= hcomplex_of_complex a ==> ( *f* (%x. x)) x @c= hcomplex_of_complex  a"
+by simp
 
 ML
 {*
-val STARC_complex_set = thm "STARC_complex_set";
-val STARC_empty_set = thm "STARC_empty_set";
-val STARC_Un = thm "STARC_Un";
-val starsetC_n_Un = thm "starsetC_n_Un";
-val InternalCSets_Un = thm "InternalCSets_Un";
-val STARC_Int = thm "STARC_Int";
-val starsetC_n_Int = thm "starsetC_n_Int";
-val InternalCSets_Int = thm "InternalCSets_Int";
-val STARC_Compl = thm "STARC_Compl";
-val starsetC_n_Compl = thm "starsetC_n_Compl";
-val InternalCSets_Compl = thm "InternalCSets_Compl";
-val STARC_mem_Compl = thm "STARC_mem_Compl";
-val STARC_diff = thm "STARC_diff";
-val starsetC_n_diff = thm "starsetC_n_diff";
-val InternalCSets_diff = thm "InternalCSets_diff";
-val STARC_subset = thm "STARC_subset";
-val STARC_mem = thm "STARC_mem";
-val STARC_hcomplex_of_complex_image_subset = thm "STARC_hcomplex_of_complex_image_subset";
 val STARC_SComplex_subset = thm "STARC_SComplex_subset";
 val STARC_hcomplex_of_complex_Int = thm "STARC_hcomplex_of_complex_Int";
 val lemma_not_hcomplexA = thm "lemma_not_hcomplexA";
-val starsetC_starsetC_n_eq = thm "starsetC_starsetC_n_eq";
-val InternalCSets_starsetC_n = thm "InternalCSets_starsetC_n";
-val InternalCSets_UNIV_diff = thm "InternalCSets_UNIV_diff";
-val starsetC_n_starsetC = thm "starsetC_n_starsetC";
-val starfunC_n_starfunC = thm "starfunC_n_starfunC";
-val starfunRC_n_starfunRC = thm "starfunRC_n_starfunRC";
-val starfunCR_n_starfunCR = thm "starfunCR_n_starfunCR";
-val starfunC_congruent = thm "starfunC_congruent";
-val starfunC = thm "starfunC";
-val starfunRC = thm "starfunRC";
-val starfunCR = thm "starfunCR";
-val starfunC_mult = thm "starfunC_mult";
-val starfunRC_mult = thm "starfunRC_mult";
-val starfunCR_mult = thm "starfunCR_mult";
-val starfunC_add = thm "starfunC_add";
-val starfunRC_add = thm "starfunRC_add";
-val starfunCR_add = thm "starfunCR_add";
-val starfunC_minus = thm "starfunC_minus";
-val starfunRC_minus = thm "starfunRC_minus";
-val starfunCR_minus = thm "starfunCR_minus";
-val starfunC_diff = thm "starfunC_diff";
-val starfunRC_diff = thm "starfunRC_diff";
-val starfunCR_diff = thm "starfunCR_diff";
-val starfunC_o2 = thm "starfunC_o2";
-val starfunC_o = thm "starfunC_o";
-val starfunC_starfunRC_o2 = thm "starfunC_starfunRC_o2";
-val starfun_starfunCR_o2 = thm "starfun_starfunCR_o2";
-val starfunC_starfunRC_o = thm "starfunC_starfunRC_o";
-val starfun_starfunCR_o = thm "starfun_starfunCR_o";
-val starfunC_const_fun = thm "starfunC_const_fun";
-val starfunRC_const_fun = thm "starfunRC_const_fun";
-val starfunCR_const_fun = thm "starfunCR_const_fun";
-val starfunC_inverse = thm "starfunC_inverse";
-val starfunRC_inverse = thm "starfunRC_inverse";
-val starfunCR_inverse = thm "starfunCR_inverse";
-val starfunC_eq = thm "starfunC_eq";
-val starfunRC_eq = thm "starfunRC_eq";
-val starfunCR_eq = thm "starfunCR_eq";
-val starfunC_capprox = thm "starfunC_capprox";
-val starfunRC_capprox = thm "starfunRC_capprox";
-val starfunCR_approx = thm "starfunCR_approx";
+val starfun_capprox = thm "starfun_capprox";
 val starfunC_hcpow = thm "starfunC_hcpow";
-val starfunC_lambda_cancel = thm "starfunC_lambda_cancel";
-val starfunCR_lambda_cancel = thm "starfunCR_lambda_cancel";
-val starfunRC_lambda_cancel = thm "starfunRC_lambda_cancel";
-val starfunC_lambda_cancel2 = thm "starfunC_lambda_cancel2";
-val starfunCR_lambda_cancel2 = thm "starfunCR_lambda_cancel2";
-val starfunRC_lambda_cancel2 = thm "starfunRC_lambda_cancel2";
-val starfunC_mult_CFinite_capprox = thm "starfunC_mult_CFinite_capprox";
-val starfunCR_mult_HFinite_capprox = thm "starfunCR_mult_HFinite_capprox";
-val starfunRC_mult_CFinite_capprox = thm "starfunRC_mult_CFinite_capprox";
-val starfunC_add_capprox = thm "starfunC_add_capprox";
-val starfunRC_add_capprox = thm "starfunRC_add_capprox";
-val starfunCR_add_approx = thm "starfunCR_add_approx";
+val starfun_mult_CFinite_capprox = thm "starfun_mult_CFinite_capprox";
+val starfun_add_capprox = thm "starfun_add_capprox";
 val starfunCR_cmod = thm "starfunCR_cmod";
-val starfunC_inverse_inverse = thm "starfunC_inverse_inverse";
-val starfunC_divide = thm "starfunC_divide";
-val starfunCR_divide = thm "starfunCR_divide";
-val starfunRC_divide = thm "starfunRC_divide";
-val starfunC_n_congruent = thm "starfunC_n_congruent";
-val starfunC_n = thm "starfunC_n";
-val starfunC_n_mult = thm "starfunC_n_mult";
-val starfunC_n_add = thm "starfunC_n_add";
-val starfunC_n_minus = thm "starfunC_n_minus";
-val starfunNat_n_diff = thm "starfunNat_n_diff";
-val starfunC_n_const_fun = thm "starfunC_n_const_fun";
-val starfunC_n_eq = thm "starfunC_n_eq";
-val starfunC_eq_iff = thm "starfunC_eq_iff";
-val starfunRC_eq_iff = thm "starfunRC_eq_iff";
-val starfunCR_eq_iff = thm "starfunCR_eq_iff";
+val starfun_inverse_inverse = thm "starfun_inverse_inverse";
+val starfun_n_diff = thm "starfun_n_diff";
 val starfunC_eq_Re_Im_iff = thm "starfunC_eq_Re_Im_iff";
 val starfunC_approx_Re_Im_iff = thm "starfunC_approx_Re_Im_iff";
 val starfunC_Idfun_capprox = thm "starfunC_Idfun_capprox";
-val starfunC_Id = thm "starfunC_Id";
 *}
 
 end
--- a/src/HOL/Complex/NSCA.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Complex/NSCA.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -80,7 +80,7 @@
 
 lemma SReal_hcmod_hcomplex_of_complex [simp]:
      "hcmod (hcomplex_of_complex r) \<in> Reals"
-by (simp add: hcomplex_of_complex_def hcmod SReal_def hypreal_of_real_def star_of_def star_n_def)
+by (auto simp add: hcmod SReal_def star_of_def)
 
 lemma SReal_hcmod_number_of [simp]: "hcmod (number_of w ::hcomplex) \<in> Reals"
 apply (subst hcomplex_number_of [symmetric])
@@ -134,8 +134,8 @@
 lemma SComplex_hcmod_SReal: 
       "z \<in> SComplex ==> hcmod z \<in> Reals"
 apply (simp add: SComplex_def SReal_def)
-apply (rule_tac z = z in eq_Abs_star)
-apply (auto simp add: hcmod hypreal_of_real_def star_of_def star_n_def hcomplex_of_complex_def cmod_def)
+apply (cases z)
+apply (auto simp add: hcmod star_of_def cmod_def star_n_eq_iff)
 apply (rule_tac x = "cmod r" in exI)
 apply (simp add: cmod_def, ultra)
 done
@@ -144,7 +144,7 @@
 by (simp add: SComplex_def hcomplex_of_complex_zero_iff)
 
 lemma SComplex_one [simp]: "1 \<in> SComplex"
-by (simp add: SComplex_def hcomplex_of_complex_def star_of_def star_n_def hypreal_one_def)
+by (simp add: SComplex_def star_of_def star_n_one_num star_n_eq_iff)
 
 (*
 Goalw [SComplex_def,SReal_def] "hcmod z \<in> Reals ==> z \<in> SComplex"
@@ -235,7 +235,7 @@
 
 lemma CInfinitesimal_CFinite_mult2:
      "[| x \<in> CInfinitesimal; y \<in> CFinite |] ==> (y * x) \<in> CInfinitesimal"
-by (auto dest: CInfinitesimal_CFinite_mult simp add: hcomplex_mult_commute)
+by (auto dest: CInfinitesimal_CFinite_mult simp add: mult_commute)
 
 lemma CInfinite_hcmod_iff: "(z \<in> CInfinite) = (hcmod z \<in> HInfinite)"
 by (simp add: CInfinite_def HInfinite_def)
@@ -392,11 +392,11 @@
 lemma capprox_mult1: 
       "[| a @c= b; c \<in> CFinite|] ==> a*c @c= b*c"
 apply (simp add: capprox_def diff_minus)
-apply (simp only: CInfinitesimal_CFinite_mult minus_mult_left hcomplex_add_mult_distrib [symmetric])
+apply (simp only: CInfinitesimal_CFinite_mult minus_mult_left left_distrib [symmetric])
 done
 
 lemma capprox_mult2: "[|a @c= b; c \<in> CFinite|] ==> c*a @c= c*b"
-by (simp add: capprox_mult1 hcomplex_mult_commute)
+by (simp add: capprox_mult1 mult_commute)
 
 lemma capprox_mult_subst:
      "[|u @c= v*x; x @c= y; v \<in> CFinite|] ==> u @c= v*y"
@@ -549,7 +549,7 @@
      "u @c= 0 ==> hcmod(x + u) - hcmod x \<in> Infinitesimal"
 apply (rule_tac e = "hcmod u" and e' = "- hcmod u" in Infinitesimal_interval2)
 apply (auto dest: capprox_approx_zero_iff [THEN iffD1]
-             simp add: mem_infmal_iff [symmetric] hypreal_diff_def)
+             simp add: mem_infmal_iff [symmetric] diff_def)
 apply (rule_tac c1 = "hcmod x" in add_le_cancel_left [THEN iffD1])
 apply (auto simp add: diff_minus [symmetric])
 done
@@ -659,12 +659,11 @@
 by (blast intro: SComplex_capprox_iff [THEN iffD1] capprox_trans2)
 
 lemma hcomplex_capproxD1:
-     "Abs_star(starrel ``{%n. X n}) @c= Abs_star(starrel``{%n. Y n})  
-      ==> Abs_star(starrel `` {%n. Re(X n)}) @=  
-          Abs_star(starrel `` {%n. Re(Y n)})"
+     "star_n X @c= star_n Y
+      ==> star_n (%n. Re(X n)) @= star_n (%n. Re(Y n))"
 apply (auto simp add: approx_FreeUltrafilterNat_iff)
 apply (drule capprox_minus_iff [THEN iffD1])
-apply (auto simp add: hcomplex_minus hcomplex_add mem_cinfmal_iff [symmetric] CInfinitesimal_hcmod_iff hcmod Infinitesimal_FreeUltrafilterNat_iff2)
+apply (auto simp add: star_n_minus star_n_add mem_cinfmal_iff [symmetric] CInfinitesimal_hcmod_iff hcmod Infinitesimal_FreeUltrafilterNat_iff2)
 apply (drule_tac x = m in spec, ultra)
 apply (rename_tac Z x)
 apply (case_tac "X x")
@@ -678,12 +677,11 @@
 
 (* same proof *)
 lemma hcomplex_capproxD2:
-     "Abs_star(starrel ``{%n. X n}) @c= Abs_star(starrel``{%n. Y n})  
-      ==> Abs_star(starrel `` {%n. Im(X n)}) @=  
-          Abs_star(starrel `` {%n. Im(Y n)})"
+     "star_n X @c= star_n Y
+      ==> star_n (%n. Im(X n)) @= star_n (%n. Im(Y n))"
 apply (auto simp add: approx_FreeUltrafilterNat_iff)
 apply (drule capprox_minus_iff [THEN iffD1])
-apply (auto simp add: hcomplex_minus hcomplex_add mem_cinfmal_iff [symmetric] CInfinitesimal_hcmod_iff hcmod Infinitesimal_FreeUltrafilterNat_iff2)
+apply (auto simp add: star_n_minus star_n_add mem_cinfmal_iff [symmetric] CInfinitesimal_hcmod_iff hcmod Infinitesimal_FreeUltrafilterNat_iff2)
 apply (drule_tac x = m in spec, ultra)
 apply (rename_tac Z x)
 apply (case_tac "X x")
@@ -695,16 +693,14 @@
 done
 
 lemma hcomplex_capproxI:
-     "[| Abs_star(starrel `` {%n. Re(X n)}) @=  
-         Abs_star(starrel `` {%n. Re(Y n)});  
-         Abs_star(starrel `` {%n. Im(X n)}) @=  
-         Abs_star(starrel `` {%n. Im(Y n)})  
-      |] ==> Abs_star(starrel ``{%n. X n}) @c= Abs_star(starrel``{%n. Y n})"
+     "[| star_n (%n. Re(X n)) @= star_n (%n. Re(Y n));  
+         star_n (%n. Im(X n)) @= star_n (%n. Im(Y n))  
+      |] ==> star_n X @c= star_n Y"
 apply (drule approx_minus_iff [THEN iffD1])
 apply (drule approx_minus_iff [THEN iffD1])
 apply (rule capprox_minus_iff [THEN iffD2])
-apply (auto simp add: mem_cinfmal_iff [symmetric] mem_infmal_iff [symmetric] hypreal_minus hypreal_add hcomplex_minus hcomplex_add CInfinitesimal_hcmod_iff hcmod Infinitesimal_FreeUltrafilterNat_iff)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl, auto)
+apply (auto simp add: mem_cinfmal_iff [symmetric] mem_infmal_iff [symmetric] star_n_add star_n_minus CInfinitesimal_hcmod_iff hcmod Infinitesimal_FreeUltrafilterNat_iff)
+apply (rule bexI [OF _ Rep_star_star_n], auto)
 apply (drule_tac x = "u/2" in spec)
 apply (drule_tac x = "u/2" in spec, auto, ultra)
 apply (drule sym, drule sym)
@@ -718,23 +714,23 @@
 done
 
 lemma capprox_approx_iff:
-     "(Abs_star(starrel ``{%n. X n}) @c= Abs_star(starrel``{%n. Y n})) = 
-       (Abs_star(starrel `` {%n. Re(X n)}) @= Abs_star(starrel `` {%n. Re(Y n)}) &  
-        Abs_star(starrel `` {%n. Im(X n)}) @= Abs_star(starrel `` {%n. Im(Y n)}))"
+     "(star_n X @c= star_n Y) = 
+       (star_n (%n. Re(X n)) @= star_n (%n. Re(Y n)) &  
+        star_n (%n. Im(X n)) @= star_n (%n. Im(Y n)))"
 apply (blast intro: hcomplex_capproxI hcomplex_capproxD1 hcomplex_capproxD2)
 done
 
 lemma hcomplex_of_hypreal_capprox_iff [simp]:
      "(hcomplex_of_hypreal x @c= hcomplex_of_hypreal z) = (x @= z)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=z in eq_Abs_star)
+apply (cases x, cases z)
 apply (simp add: hcomplex_of_hypreal capprox_approx_iff)
 done
 
 lemma CFinite_HFinite_Re:
-     "Abs_star(starrel ``{%n. X n}) \<in> CFinite  
-      ==> Abs_star(starrel `` {%n. Re(X n)}) \<in> HFinite"
+     "star_n X \<in> CFinite  
+      ==> star_n (%n. Re(X n)) \<in> HFinite"
 apply (auto simp add: CFinite_hcmod_iff hcmod HFinite_FreeUltrafilterNat_iff)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl)
+apply (rule bexI [OF _ Rep_star_star_n])
 apply (rule_tac x = u in exI, ultra)
 apply (drule sym, case_tac "X x")
 apply (auto simp add: complex_mod numeral_2_eq_2 simp del: realpow_Suc)
@@ -745,10 +741,10 @@
 done
 
 lemma CFinite_HFinite_Im:
-     "Abs_star(starrel ``{%n. X n}) \<in> CFinite  
-      ==> Abs_star(starrel `` {%n. Im(X n)}) \<in> HFinite"
+     "star_n X \<in> CFinite  
+      ==> star_n (%n. Im(X n)) \<in> HFinite"
 apply (auto simp add: CFinite_hcmod_iff hcmod HFinite_FreeUltrafilterNat_iff)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl)
+apply (rule bexI [OF _ Rep_star_star_n])
 apply (rule_tac x = u in exI, ultra)
 apply (drule sym, case_tac "X x")
 apply (auto simp add: complex_mod simp del: realpow_Suc)
@@ -758,12 +754,12 @@
 done
 
 lemma HFinite_Re_Im_CFinite:
-     "[| Abs_star(starrel `` {%n. Re(X n)}) \<in> HFinite;  
-         Abs_star(starrel `` {%n. Im(X n)}) \<in> HFinite  
-      |] ==> Abs_star(starrel ``{%n. X n}) \<in> CFinite"
+     "[| star_n (%n. Re(X n)) \<in> HFinite;  
+         star_n (%n. Im(X n)) \<in> HFinite  
+      |] ==> star_n X \<in> CFinite"
 apply (auto simp add: CFinite_hcmod_iff hcmod HFinite_FreeUltrafilterNat_iff)
 apply (rename_tac Y Z u v)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl)
+apply (rule bexI [OF _ Rep_star_star_n])
 apply (rule_tac x = "2* (u + v) " in exI)
 apply ultra
 apply (drule sym, case_tac "X x")
@@ -778,66 +774,65 @@
 done
 
 lemma CFinite_HFinite_iff:
-     "(Abs_star(starrel ``{%n. X n}) \<in> CFinite) =  
-      (Abs_star(starrel `` {%n. Re(X n)}) \<in> HFinite &  
-       Abs_star(starrel `` {%n. Im(X n)}) \<in> HFinite)"
+     "(star_n X \<in> CFinite) =  
+      (star_n (%n. Re(X n)) \<in> HFinite &  
+       star_n (%n. Im(X n)) \<in> HFinite)"
 by (blast intro: HFinite_Re_Im_CFinite CFinite_HFinite_Im CFinite_HFinite_Re)
 
 lemma SComplex_Re_SReal:
-     "Abs_star(starrel ``{%n. X n}) \<in> SComplex  
-      ==> Abs_star(starrel `` {%n. Re(X n)}) \<in> Reals"
-apply (auto simp add: SComplex_def hcomplex_of_complex_def SReal_def hypreal_of_real_def star_of_def star_n_def)
+     "star_n X \<in> SComplex  
+      ==> star_n (%n. Re(X n)) \<in> Reals"
+apply (auto simp add: SComplex_def SReal_def star_of_def star_n_eq_iff)
 apply (rule_tac x = "Re r" in exI, ultra)
 done
 
 lemma SComplex_Im_SReal:
-     "Abs_star(starrel ``{%n. X n}) \<in> SComplex  
-      ==> Abs_star(starrel `` {%n. Im(X n)}) \<in> Reals"
-apply (auto simp add: SComplex_def hcomplex_of_complex_def SReal_def hypreal_of_real_def star_of_def star_n_def)
+     "star_n X \<in> SComplex  
+      ==> star_n (%n. Im(X n)) \<in> Reals"
+apply (auto simp add: SComplex_def SReal_def star_of_def star_n_eq_iff)
 apply (rule_tac x = "Im r" in exI, ultra)
 done
 
 lemma Reals_Re_Im_SComplex:
-     "[| Abs_star(starrel `` {%n. Re(X n)}) \<in> Reals;  
-         Abs_star(starrel `` {%n. Im(X n)}) \<in> Reals  
-      |] ==> Abs_star(starrel ``{%n. X n}) \<in> SComplex"
-apply (auto simp add: SComplex_def hcomplex_of_complex_def SReal_def hypreal_of_real_def star_of_def star_n_def)
+     "[| star_n (%n. Re(X n)) \<in> Reals;  
+         star_n (%n. Im(X n)) \<in> Reals  
+      |] ==> star_n X \<in> SComplex"
+apply (auto simp add: SComplex_def SReal_def star_of_def star_n_eq_iff)
 apply (rule_tac x = "Complex r ra" in exI, ultra)
 done
 
 lemma SComplex_SReal_iff:
-     "(Abs_star(starrel ``{%n. X n}) \<in> SComplex) =  
-      (Abs_star(starrel `` {%n. Re(X n)}) \<in> Reals &  
-       Abs_star(starrel `` {%n. Im(X n)}) \<in> Reals)"
+     "(star_n X \<in> SComplex) =  
+      (star_n (%n. Re(X n)) \<in> Reals &  
+       star_n (%n. Im(X n)) \<in> Reals)"
 by (blast intro: SComplex_Re_SReal SComplex_Im_SReal Reals_Re_Im_SComplex)
 
 lemma CInfinitesimal_Infinitesimal_iff:
-     "(Abs_star(starrel ``{%n. X n}) \<in> CInfinitesimal) =  
-      (Abs_star(starrel `` {%n. Re(X n)}) \<in> Infinitesimal &  
-       Abs_star(starrel `` {%n. Im(X n)}) \<in> Infinitesimal)"
-by (simp add: mem_cinfmal_iff mem_infmal_iff hcomplex_zero_num hypreal_zero_num capprox_approx_iff)
+     "(star_n X \<in> CInfinitesimal) =  
+      (star_n (%n. Re(X n)) \<in> Infinitesimal &  
+       star_n (%n. Im(X n)) \<in> Infinitesimal)"
+by (simp add: mem_cinfmal_iff mem_infmal_iff star_n_zero_num capprox_approx_iff)
 
 lemma eq_Abs_star_EX:
-     "(\<exists>t. P t) = (\<exists>X. P (Abs_star(starrel `` {X})))"
+     "(\<exists>t. P t) = (\<exists>X. P (star_n X))"
 apply auto
-apply (rule_tac z = t in eq_Abs_star, auto)
+apply (rule_tac x = t in star_cases, auto)
 done
 
 lemma eq_Abs_star_Bex:
-     "(\<exists>t \<in> A. P t) = (\<exists>X. (Abs_star(starrel `` {X})) \<in> A &  
-                         P (Abs_star(starrel `` {X})))"
+     "(\<exists>t \<in> A. P t) = (\<exists>X. star_n X \<in> A & P (star_n X))"
 apply auto
-apply (rule_tac z = t in eq_Abs_star, auto)
+apply (rule_tac x = t in star_cases, auto)
 done
 
 (* Here we go - easy proof now!! *)
 lemma stc_part_Ex: "x:CFinite ==> \<exists>t \<in> SComplex. x @c= t"
-apply (rule_tac z = x in eq_Abs_star)
+apply (cases x)
 apply (auto simp add: CFinite_HFinite_iff eq_Abs_star_Bex SComplex_SReal_iff capprox_approx_iff)
 apply (drule st_part_Ex, safe)+
-apply (rule_tac z = t in eq_Abs_star)
-apply (rule_tac z = ta in eq_Abs_star, auto)
-apply (rule_tac x = "%n. Complex (xa n) (xb n) " in exI)
+apply (rule_tac x = t in star_cases)
+apply (rule_tac x = ta in star_cases, auto)
+apply (rule_tac x = "%n. Complex (Xa n) (Xb n) " in exI)
 apply auto
 done
 
@@ -1097,7 +1092,7 @@
  prefer 2 apply simp 
 apply (erule_tac V = "x = stc x + e" in thin_rl)
 apply (erule_tac V = "y = stc y + ea" in thin_rl)
-apply (simp add: hcomplex_add_mult_distrib right_distrib)
+apply (simp add: left_distrib right_distrib)
 apply (drule stc_SComplex)+
 apply (simp (no_asm_use) add: add_assoc)
 apply (rule stc_CInfinitesimal_add_SComplex)
@@ -1135,14 +1130,14 @@
 
 lemma CFinite_HFinite_hcomplex_of_hypreal:
      "z \<in> HFinite ==> hcomplex_of_hypreal z \<in> CFinite"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hcomplex_of_hypreal CFinite_HFinite_iff hypreal_zero_def [symmetric])
+apply (cases z)
+apply (simp add: hcomplex_of_hypreal CFinite_HFinite_iff star_n_zero_num [symmetric])
 done
 
 lemma SComplex_SReal_hcomplex_of_hypreal:
      "x \<in> Reals ==>  hcomplex_of_hypreal x \<in> SComplex"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: hcomplex_of_hypreal SComplex_SReal_iff hypreal_zero_def [symmetric])
+apply (cases x)
+apply (simp add: hcomplex_of_hypreal SComplex_SReal_iff star_n_zero_num [symmetric])
 done
 
 lemma stc_hcomplex_of_hypreal: 
@@ -1171,53 +1166,52 @@
 by (simp add: CInfinitesimal_hcmod_iff)
 
 lemma CInfinite_HInfinite_iff:
-     "(Abs_star(starrel ``{%n. X n}) \<in> CInfinite) =  
-      (Abs_star(starrel `` {%n. Re(X n)}) \<in> HInfinite |  
-       Abs_star(starrel `` {%n. Im(X n)}) \<in> HInfinite)"
+     "(star_n X \<in> CInfinite) =  
+      (star_n (%n. Re(X n)) \<in> HInfinite |  
+       star_n (%n. Im(X n)) \<in> HInfinite)"
 by (simp add: CInfinite_CFinite_iff HInfinite_HFinite_iff CFinite_HFinite_iff)
 
 text{*These theorems should probably be deleted*}
 lemma hcomplex_split_CInfinitesimal_iff:
      "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y \<in> CInfinitesimal) =  
       (x \<in> Infinitesimal & y \<in> Infinitesimal)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: iii_def hcomplex_add hcomplex_mult hcomplex_of_hypreal CInfinitesimal_Infinitesimal_iff)
+apply (cases x, cases y)
+apply (simp add: iii_def star_of_def star_n_add star_n_mult hcomplex_of_hypreal CInfinitesimal_Infinitesimal_iff)
 done
 
 lemma hcomplex_split_CFinite_iff:
      "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y \<in> CFinite) =  
       (x \<in> HFinite & y \<in> HFinite)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: iii_def hcomplex_add hcomplex_mult hcomplex_of_hypreal CFinite_HFinite_iff)
+apply (cases x, cases y)
+apply (simp add: iii_def star_of_def star_n_add star_n_mult hcomplex_of_hypreal CFinite_HFinite_iff)
 done
 
 lemma hcomplex_split_SComplex_iff:
      "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y \<in> SComplex) =  
       (x \<in> Reals & y \<in> Reals)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: iii_def hcomplex_add hcomplex_mult hcomplex_of_hypreal SComplex_SReal_iff)
+apply (cases x, cases y)
+apply (simp add: iii_def star_of_def star_n_add star_n_mult hcomplex_of_hypreal SComplex_SReal_iff)
 done
 
 lemma hcomplex_split_CInfinite_iff:
      "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y \<in> CInfinite) =  
       (x \<in> HInfinite | y \<in> HInfinite)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: iii_def hcomplex_add hcomplex_mult hcomplex_of_hypreal CInfinite_HInfinite_iff)
+apply (cases x, cases y)
+apply (simp add: iii_def star_of_def star_n_add star_n_mult hcomplex_of_hypreal CInfinite_HInfinite_iff)
 done
 
 lemma hcomplex_split_capprox_iff:
      "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y @c=  
        hcomplex_of_hypreal x' + iii * hcomplex_of_hypreal y') =  
       (x @= x' & y @= y')"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (rule_tac z=x' in eq_Abs_star, rule_tac z=y' in eq_Abs_star)
-apply (simp add: iii_def hcomplex_add hcomplex_mult hcomplex_of_hypreal capprox_approx_iff)
+apply (cases x, cases y, cases x', cases y')
+apply (simp add: iii_def star_of_def star_n_add star_n_mult hcomplex_of_hypreal capprox_approx_iff)
 done
 
 lemma complex_seq_to_hcomplex_CInfinitesimal:
      "\<forall>n. cmod (X n - x) < inverse (real (Suc n)) ==>  
-      Abs_star(starrel``{X}) - hcomplex_of_complex x \<in> CInfinitesimal"
-apply (simp add: hcomplex_diff CInfinitesimal_hcmod_iff hcomplex_of_complex_def star_of_def star_n_def Infinitesimal_FreeUltrafilterNat_iff hcmod)
+      star_n X - hcomplex_of_complex x \<in> CInfinitesimal"
+apply (simp add: star_n_diff CInfinitesimal_hcmod_iff star_of_def Infinitesimal_FreeUltrafilterNat_iff hcmod)
 apply (rule bexI, auto)
 apply (auto dest: FreeUltrafilterNat_inverse_real_of_posnat FreeUltrafilterNat_all FreeUltrafilterNat_Int intro: order_less_trans FreeUltrafilterNat_subset)
 done
@@ -1228,13 +1222,11 @@
 
 lemma hcomplex_of_complex_approx_zero_iff [simp]:
      "(hcomplex_of_complex z @c= 0) = (z = 0)"
-by (simp add: hcomplex_of_complex_zero [symmetric]
-         del: hcomplex_of_complex_zero)
+by (simp add: star_of_zero [symmetric] del: star_of_zero)
 
 lemma hcomplex_of_complex_approx_zero_iff2 [simp]:
      "(0 @c= hcomplex_of_complex z) = (z = 0)"
-by (simp add: hcomplex_of_complex_zero [symmetric]
-         del: hcomplex_of_complex_zero)
+by (simp add: star_of_zero [symmetric] del: star_of_zero)
 
 
 ML
--- a/src/HOL/Complex/NSComplex.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Complex/NSComplex.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -12,85 +12,52 @@
 begin
 
 types hcomplex = "complex star"
-(*
-constdefs
-    hcomplexrel :: "((nat=>complex)*(nat=>complex)) set"
-    "hcomplexrel == {p. \<exists>X Y. p = ((X::nat=>complex),Y) &
-                        {n::nat. X(n) = Y(n)}: FreeUltrafilterNat}"
 
-typedef hcomplex = "{x::nat=>complex. True}//hcomplexrel"
-  by (auto simp add: quotient_def)
-
-instance hcomplex :: "{zero, one, plus, times, minus, inverse, power}" ..
-
-defs (overloaded)
-  hcomplex_zero_def:
-  "0 == Abs_hcomplex(hcomplexrel `` {%n. (0::complex)})"
-
-  hcomplex_one_def:
-  "1 == Abs_hcomplex(hcomplexrel `` {%n. (1::complex)})"
-
-
-  hcomplex_minus_def:
-  "- z == Abs_hcomplex(UN X: Rep_hcomplex(z).
-                       hcomplexrel `` {%n::nat. - (X n)})"
-
-  hcomplex_diff_def:
-  "w - z == w + -(z::hcomplex)"
-
-  hcinv_def:
-  "inverse(P) == Abs_hcomplex(UN X: Rep_hcomplex(P).
-                    hcomplexrel `` {%n. inverse(X n)})"
-*)
+syntax hcomplex_of_complex :: "real => real star"
+translations "hcomplex_of_complex" => "star_of :: complex => complex star"
 
 constdefs
 
-  hcomplex_of_complex :: "complex => hcomplex"
-(*  "hcomplex_of_complex z == Abs_star(starrel `` {%n. z})"*)
-  "hcomplex_of_complex z == star_of z"
-
   (*--- real and Imaginary parts ---*)
 
   hRe :: "hcomplex => hypreal"
-  "hRe(z) == Abs_star(UN X:Rep_star(z). starrel `` {%n. Re (X n)})"
+  "hRe(z) == ( *f* Re) z"
 
   hIm :: "hcomplex => hypreal"
-  "hIm(z) == Abs_star(UN X:Rep_star(z). starrel `` {%n. Im (X n)})"
+  "hIm(z) == ( *f* Im) z"
 
 
   (*----------- modulus ------------*)
 
   hcmod :: "hcomplex => hypreal"
-  "hcmod z == Abs_star(UN X: Rep_star(z).
-			  starrel `` {%n. cmod (X n)})"
+  "hcmod z == ( *f* cmod) z"
 
   (*------ imaginary unit ----------*)
 
   iii :: hcomplex
-  "iii == Abs_star(starrel `` {%n. ii})"
+  "iii == star_of ii"
 
   (*------- complex conjugate ------*)
 
   hcnj :: "hcomplex => hcomplex"
-  "hcnj z == Abs_star(UN X:Rep_star(z). starrel `` {%n. cnj (X n)})"
+  "hcnj z == ( *f* cnj) z"
 
   (*------------ Argand -------------*)
 
   hsgn :: "hcomplex => hcomplex"
-  "hsgn z == Abs_star(UN X:Rep_star(z). starrel `` {%n. sgn(X n)})"
+  "hsgn z == ( *f* sgn) z"
 
   harg :: "hcomplex => hypreal"
-  "harg z == Abs_star(UN X:Rep_star(z). starrel `` {%n. arg(X n)})"
+  "harg z == ( *f* arg) z"
 
   (* abbreviation for (cos a + i sin a) *)
   hcis :: "hypreal => hcomplex"
-  "hcis a == Abs_star(UN X:Rep_star(a). starrel `` {%n. cis (X n)})"
+  "hcis a == ( *f* cis) a"
 
   (*----- injection from hyperreals -----*)
 
   hcomplex_of_hypreal :: "hypreal => hcomplex"
-  "hcomplex_of_hypreal r == Abs_star(UN X:Rep_star(r).
-			       starrel `` {%n. complex_of_real (X n)})"
+  "hcomplex_of_hypreal r == ( *f* complex_of_real) r"
 
   (* abbreviation for r*(cos a + i sin a) *)
   hrcis :: "[hypreal, hypreal] => hcomplex"
@@ -104,225 +71,63 @@
 
 constdefs
   HComplex :: "[hypreal,hypreal] => hcomplex"
-   "HComplex x y == hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y"
-
-(*
-defs (overloaded)
-*)
-  (*----------- division ----------*)
-(*
-  hcomplex_divide_def:
-  "w / (z::hcomplex) == w * inverse z"
-
-  hcomplex_add_def:
-  "w + z == Abs_star(UN X:Rep_star(w). UN Y:Rep_star(z).
-		      starrel `` {%n. X n + Y n})"
-
-  hcomplex_mult_def:
-  "w * z == Abs_star(UN X:Rep_star(w). UN Y:Rep_star(z).
-		      starrel `` {%n. X n * Y n})"
-*)
+(*   "HComplex x y == hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y"*)
+   "HComplex == Ifun2_of Complex"
 
 
 consts
-  "hcpow"  :: "[hcomplex,hypnat] => hcomplex"     (infixr 80)
+  "hcpow"  :: "[hcomplex,hypnat] => hcomplex"     (infixr "hcpow" 80)
 
 defs
   (* hypernatural powers of nonstandard complex numbers *)
   hcpow_def:
   "(z::hcomplex) hcpow (n::hypnat)
-      == Abs_star(UN X:Rep_star(z). UN Y: Rep_star(n).
-             starrel `` {%n. (X n) ^ (Y n)})"
+      == Ifun2_of (op ^) z n"
 
 
 subsection{*Properties of Nonstandard Real and Imaginary Parts*}
 
-lemma hRe:
-     "hRe(Abs_star (starrel `` {X})) =
-      Abs_star(starrel `` {%n. Re(X n)})"
-apply (simp add: hRe_def)
-apply (rule_tac f = Abs_star in arg_cong)
-apply (auto iff: starrel_iff, ultra)
-done
+lemma hRe: "hRe (star_n X) = star_n (%n. Re(X n))"
+by (simp add: hRe_def starfun)
 
-lemma hIm:
-     "hIm(Abs_star (starrel `` {X})) =
-      Abs_star(starrel `` {%n. Im(X n)})"
-apply (simp add: hIm_def)
-apply (rule_tac f = Abs_star in arg_cong)
-apply (auto iff: starrel_iff, ultra)
-done
+lemma hIm: "hIm (star_n X) = star_n (%n. Im(X n))"
+by (simp add: hIm_def starfun)
 
 lemma hcomplex_hRe_hIm_cancel_iff:
-     "(w=z) = (hRe(w) = hRe(z) & hIm(w) = hIm(z))"
-apply (rule_tac z=z in eq_Abs_star, rule_tac z=w in eq_Abs_star)
-apply (auto simp add: hRe hIm complex_Re_Im_cancel_iff iff: starrel_iff)
-apply (ultra+)
-done
+     "!!w z. (w=z) = (hRe(w) = hRe(z) & hIm(w) = hIm(z))"
+by (unfold hRe_def hIm_def, transfer, rule complex_Re_Im_cancel_iff)
 
 lemma hcomplex_equality [intro?]: "hRe z = hRe w ==> hIm z = hIm w ==> z = w"
-by (simp add: hcomplex_hRe_hIm_cancel_iff) 
+by (simp add: hcomplex_hRe_hIm_cancel_iff)
 
 lemma hcomplex_hRe_zero [simp]: "hRe 0 = 0"
-by (simp add: hRe hypreal_zero_num)
+by (simp add: hRe star_n_zero_num)
 
 lemma hcomplex_hIm_zero [simp]: "hIm 0 = 0"
-by (simp add: hIm hypreal_zero_num)
+by (simp add: hIm star_n_zero_num)
 
 lemma hcomplex_hRe_one [simp]: "hRe 1 = 1"
-by (simp add: hRe hypreal_one_num)
+by (simp add: hRe star_n_one_num)
 
 lemma hcomplex_hIm_one [simp]: "hIm 1 = 0"
-by (simp add: hIm hypreal_one_def hypreal_zero_num)
+by (simp add: hIm star_n_one_num star_n_zero_num)
 
 
 subsection{*Addition for Nonstandard Complex Numbers*}
-(*
-lemma hcomplex_add_congruent2:
-    "congruent2 starrel starrel (%X Y. starrel `` {%n. X n + Y n})"
-by (auto simp add: congruent2_def iff: starrel_iff, ultra) 
-*)
-lemma hcomplex_add:
-  "Abs_star(starrel``{%n. X n}) + 
-   Abs_star(starrel``{%n. Y n}) =
-     Abs_star(starrel``{%n. X n + Y n})"
-by (rule hypreal_add)
-
-lemma hcomplex_add_commute: "(z::hcomplex) + w = w + z"
-by (rule add_commute)
-
-lemma hcomplex_add_assoc: "((z1::hcomplex) + z2) + z3 = z1 + (z2 + z3)"
-by (rule add_assoc)
-
-lemma hcomplex_add_zero_left: "(0::hcomplex) + z = z"
-by simp
-
-lemma hcomplex_add_zero_right: "z + (0::hcomplex) = z"
-by simp
-
-lemma hRe_add: "hRe(x + y) = hRe(x) + hRe(y)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hRe hcomplex_add hypreal_add complex_Re_add)
-done
-
-lemma hIm_add: "hIm(x + y) = hIm(x) + hIm(y)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hIm hcomplex_add hypreal_add complex_Im_add)
-done
-
-
-subsection{*Additive Inverse on Nonstandard Complex Numbers*}
-(*
-lemma hcomplex_minus_congruent:
-     "(%X. starrel `` {%n. - (X n)}) respects starrel"
-by (simp add: congruent_def)
-*)
-lemma hcomplex_minus:
-  "- (Abs_star(starrel `` {%n. X n})) =
-      Abs_star(starrel `` {%n. -(X n)})"
-by (rule hypreal_minus)
-
-lemma hcomplex_add_minus_left: "-z + z = (0::hcomplex)"
-by simp
-
-
-subsection{*Multiplication for Nonstandard Complex Numbers*}
-
-lemma hcomplex_mult:
-  "Abs_star(starrel``{%n. X n}) *
-     Abs_star(starrel``{%n. Y n}) =
-     Abs_star(starrel``{%n. X n * Y n})"
-by (rule hypreal_mult)
-
-lemma hcomplex_mult_commute: "(w::hcomplex) * z = z * w"
-by (rule mult_commute)
-
-lemma hcomplex_mult_assoc: "((u::hcomplex) * v) * w = u * (v * w)"
-by (rule mult_assoc)
-
-lemma hcomplex_mult_one_left: "(1::hcomplex) * z = z"
-by (rule mult_1_left)
 
-lemma hcomplex_mult_zero_left: "(0::hcomplex) * z = 0"
-by (rule mult_zero_left)
-
-lemma hcomplex_add_mult_distrib:
-     "((z1::hcomplex) + z2) * w = (z1 * w) + (z2 * w)"
-by (rule left_distrib)
-
-lemma hcomplex_zero_not_eq_one: "(0::hcomplex) \<noteq> (1::hcomplex)"
-by (rule zero_neq_one)
-
-declare hcomplex_zero_not_eq_one [THEN not_sym, simp]
-
-
-subsection{*Inverse of Nonstandard Complex Number*}
-
-lemma hcomplex_inverse:
-  "inverse (Abs_star(starrel `` {%n. X n})) =
-      Abs_star(starrel `` {%n. inverse (X n)})"
-apply (fold star_n_def)
-apply (simp add: star_inverse_def Ifun_of_def star_of_def Ifun_star_n)
-done
-
-lemma hcomplex_mult_inv_left:
-      "z \<noteq> (0::hcomplex) ==> inverse(z) * z = (1::hcomplex)"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hypreal_zero_def hypreal_one_def hcomplex_inverse hcomplex_mult, ultra)
-apply (rule ccontr)
-apply (drule left_inverse, auto)
-done
+lemma hRe_add: "!!x y. hRe(x + y) = hRe(x) + hRe(y)"
+by (unfold hRe_def, transfer, rule complex_Re_add)
 
-subsection {* The Field of Nonstandard Complex Numbers *}
-(*
-instance hcomplex :: field
-proof
-  fix z u v w :: hcomplex
-  show "(u + v) + w = u + (v + w)"
-    by (simp add: hcomplex_add_assoc)
-  show "z + w = w + z"
-    by (simp add: hcomplex_add_commute)
-  show "0 + z = z"
-    by (simp add: hcomplex_add_zero_left)
-  show "-z + z = 0"
-    by (simp add: hcomplex_add_minus_left)
-  show "z - w = z + -w"
-    by (simp add: hcomplex_diff_def)
-  show "(u * v) * w = u * (v * w)"
-    by (simp add: hcomplex_mult_assoc)
-  show "z * w = w * z"
-    by (simp add: hcomplex_mult_commute)
-  show "1 * z = z"
-    by (simp add: hcomplex_mult_one_left)
-  show "0 \<noteq> (1::hcomplex)"
-    by (rule hcomplex_zero_not_eq_one)
-  show "(u + v) * w = u * w + v * w"
-    by (simp add: hcomplex_add_mult_distrib)
-  show "z / w = z * inverse w"
-    by (simp add: hcomplex_divide_def)
-  assume "w \<noteq> 0"
-  thus "inverse w * w = 1"
-    by (rule hcomplex_mult_inv_left)
-qed
-
-instance hcomplex :: division_by_zero
-proof
-  show "inverse 0 = (0::hcomplex)"
-    by (simp add: hcomplex_inverse hcomplex_zero_def)
-qed
-*)
+lemma hIm_add: "!!x y. hIm(x + y) = hIm(x) + hIm(y)"
+by (unfold hIm_def, transfer, rule complex_Im_add)
 
 subsection{*More Minus Laws*}
 
-lemma hRe_minus: "hRe(-z) = - hRe(z)"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hRe hcomplex_minus hypreal_minus complex_Re_minus)
-done
+lemma hRe_minus: "!!z. hRe(-z) = - hRe(z)"
+by (unfold hRe_def, transfer, rule complex_Re_minus)
 
-lemma hIm_minus: "hIm(-z) = - hIm(z)"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hIm hcomplex_minus hypreal_minus complex_Im_minus)
-done
+lemma hIm_minus: "!!z. hIm(-z) = - hIm(z)"
+by (unfold hIm_def, transfer, rule complex_Im_minus)
 
 lemma hcomplex_add_minus_eq_minus:
       "x + y = (0::hcomplex) ==> x = -y"
@@ -331,25 +136,22 @@
 done
 
 lemma hcomplex_i_mult_eq [simp]: "iii * iii = - 1"
-by (simp add: iii_def hcomplex_mult hypreal_one_def hcomplex_minus)
+by (simp add: iii_def star_of_def star_n_mult star_n_one_num star_n_minus)
 
 lemma hcomplex_i_mult_left [simp]: "iii * (iii * z) = -z"
 by (simp add: mult_assoc [symmetric])
 
 lemma hcomplex_i_not_zero [simp]: "iii \<noteq> 0"
-by (simp add: iii_def hypreal_zero_def)
+by (simp add: iii_def star_of_def star_n_zero_num star_n_eq_iff)
 
 
 subsection{*More Multiplication Laws*}
 
-lemma hcomplex_mult_one_right: "z * (1::hcomplex) = z"
-by (rule OrderedGroup.mult_1_right)
-
 lemma hcomplex_mult_minus_one [simp]: "- 1 * (z::hcomplex) = -z"
 by simp
 
 lemma hcomplex_mult_minus_one_right [simp]: "(z::hcomplex) * - 1 = -z"
-by (subst hcomplex_mult_commute, simp)
+by simp
 
 lemma hcomplex_mult_left_cancel:
      "(c::hcomplex) \<noteq> (0::hcomplex) ==> (c*a=c*b) = (a=b)"
@@ -362,11 +164,6 @@
 
 subsection{*Subraction and Division*}
 
-lemma hcomplex_diff:
- "Abs_star(starrel``{%n. X n}) - Abs_star(starrel``{%n. Y n}) =
-  Abs_star(starrel``{%n. X n - Y n})"
-by (rule hypreal_diff)
-
 lemma hcomplex_diff_eq_eq [simp]: "((x::hcomplex) - y = z) = (x = z + y)"
 by (rule OrderedGroup.diff_eq_eq)
 
@@ -377,92 +174,78 @@
 subsection{*Embedding Properties for @{term hcomplex_of_hypreal} Map*}
 
 lemma hcomplex_of_hypreal:
-  "hcomplex_of_hypreal (Abs_star(starrel `` {%n. X n})) =
-      Abs_star(starrel `` {%n. complex_of_real (X n)})"
-apply (simp add: hcomplex_of_hypreal_def)
-apply (rule_tac f = Abs_star in arg_cong, auto iff: starrel_iff, ultra)
-done
+  "hcomplex_of_hypreal (star_n X) = star_n (%n. complex_of_real (X n))"
+by (simp add: hcomplex_of_hypreal_def starfun)
 
 lemma hcomplex_of_hypreal_cancel_iff [iff]:
-     "(hcomplex_of_hypreal x = hcomplex_of_hypreal y) = (x = y)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hcomplex_of_hypreal)
-done
+     "!!x y. (hcomplex_of_hypreal x = hcomplex_of_hypreal y) = (x = y)"
+by (unfold hcomplex_of_hypreal_def, transfer, simp)
 
 lemma hcomplex_of_hypreal_one [simp]: "hcomplex_of_hypreal 1 = 1"
-by (simp add: hypreal_one_def hcomplex_of_hypreal hypreal_one_num)
+by (simp add: hcomplex_of_hypreal star_n_one_num)
 
 lemma hcomplex_of_hypreal_zero [simp]: "hcomplex_of_hypreal 0 = 0"
-by (simp add: hypreal_zero_def hypreal_zero_def hcomplex_of_hypreal)
+by (simp add: star_n_zero_num hcomplex_of_hypreal)
 
 lemma hcomplex_of_hypreal_minus [simp]:
-     "hcomplex_of_hypreal(-x) = - hcomplex_of_hypreal x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: hcomplex_of_hypreal hcomplex_minus hypreal_minus)
-done
+     "!!x. hcomplex_of_hypreal(-x) = - hcomplex_of_hypreal x"
+by (unfold hcomplex_of_hypreal_def, transfer, simp)
 
 lemma hcomplex_of_hypreal_inverse [simp]:
-     "hcomplex_of_hypreal(inverse x) = inverse(hcomplex_of_hypreal x)"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: hcomplex_of_hypreal hypreal_inverse hcomplex_inverse)
-done
+     "!!x. hcomplex_of_hypreal(inverse x) = inverse(hcomplex_of_hypreal x)"
+by (unfold hcomplex_of_hypreal_def, transfer, simp)
 
 lemma hcomplex_of_hypreal_add [simp]:
-  "hcomplex_of_hypreal (x + y) = hcomplex_of_hypreal x + hcomplex_of_hypreal y"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hcomplex_of_hypreal hypreal_add hcomplex_add)
-done
+     "!!x y. hcomplex_of_hypreal (x + y) =
+      hcomplex_of_hypreal x + hcomplex_of_hypreal y"
+by (unfold hcomplex_of_hypreal_def, transfer, simp)
 
 lemma hcomplex_of_hypreal_diff [simp]:
-     "hcomplex_of_hypreal (x - y) =
+     "!!x y. hcomplex_of_hypreal (x - y) =
       hcomplex_of_hypreal x - hcomplex_of_hypreal y "
-by (simp add: hypreal_diff_def)
+by (unfold hcomplex_of_hypreal_def, transfer, simp)
 
 lemma hcomplex_of_hypreal_mult [simp]:
-  "hcomplex_of_hypreal (x * y) = hcomplex_of_hypreal x * hcomplex_of_hypreal y"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hcomplex_of_hypreal hypreal_mult hcomplex_mult)
-done
+     "!!x y. hcomplex_of_hypreal (x * y) =
+      hcomplex_of_hypreal x * hcomplex_of_hypreal y"
+by (unfold hcomplex_of_hypreal_def, transfer, simp)
 
 lemma hcomplex_of_hypreal_divide [simp]:
-  "hcomplex_of_hypreal(x/y) = hcomplex_of_hypreal x / hcomplex_of_hypreal y"
-by (simp add: divide_inverse)
+     "!!x y. hcomplex_of_hypreal(x/y) =
+      hcomplex_of_hypreal x / hcomplex_of_hypreal y"
+by (unfold hcomplex_of_hypreal_def, transfer, simp)
 
 lemma hRe_hcomplex_of_hypreal [simp]: "hRe(hcomplex_of_hypreal z) = z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (auto simp add: hcomplex_of_hypreal hRe)
+apply (cases z)
+apply (simp add: hcomplex_of_hypreal hRe)
 done
 
 lemma hIm_hcomplex_of_hypreal [simp]: "hIm(hcomplex_of_hypreal z) = 0"
-apply (rule_tac z=z in eq_Abs_star)
-apply (auto simp add: hcomplex_of_hypreal hIm hypreal_zero_num)
+apply (cases z)
+apply (simp add: hcomplex_of_hypreal hIm star_n_zero_num)
 done
 
 lemma hcomplex_of_hypreal_epsilon_not_zero [simp]:
      "hcomplex_of_hypreal epsilon \<noteq> 0"
-by (auto simp add: hcomplex_of_hypreal epsilon_def star_n_def hypreal_zero_def)
+by (simp add: hcomplex_of_hypreal epsilon_def star_n_zero_num star_n_eq_iff)
 
 
 subsection{*HComplex theorems*}
 
-lemma hRe_HComplex [simp]: "hRe (HComplex x y) = x"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: HComplex_def hRe iii_def hcomplex_add hcomplex_mult hcomplex_of_hypreal)
-done
+lemma hRe_HComplex [simp]: "!!x y. hRe (HComplex x y) = x"
+by (unfold hRe_def HComplex_def, transfer, simp)
 
-lemma hIm_HComplex [simp]: "hIm (HComplex x y) = y"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: HComplex_def hIm iii_def hcomplex_add hcomplex_mult hcomplex_of_hypreal)
-done
+lemma hIm_HComplex [simp]: "!!x y. hIm (HComplex x y) = y"
+by (unfold hIm_def HComplex_def, transfer, simp)
 
 text{*Relates the two nonstandard constructions*}
 lemma HComplex_eq_Abs_star_Complex:
-     "HComplex (Abs_star (starrel `` {X})) (Abs_star (starrel `` {Y})) =
-      Abs_star(starrel `` {%n::nat. Complex (X n) (Y n)})";
-by (simp add: hcomplex_hRe_hIm_cancel_iff hRe hIm) 
+     "HComplex (star_n X) (star_n Y) =
+      star_n (%n::nat. Complex (X n) (Y n))"
+by (simp add: hcomplex_hRe_hIm_cancel_iff hRe hIm)
 
 lemma hcomplex_surj [simp]: "HComplex (hRe z) (hIm z) = z"
-by (simp add: hcomplex_equality) 
+by (simp add: hcomplex_equality)
 
 lemma hcomplex_induct [case_names rect(*, induct type: hcomplex*)]:
      "(\<And>x y. P (HComplex x y)) ==> P z"
@@ -471,24 +254,18 @@
 
 subsection{*Modulus (Absolute Value) of Nonstandard Complex Number*}
 
-lemma hcmod:
-  "hcmod (Abs_star(starrel `` {%n. X n})) =
-      Abs_star(starrel `` {%n. cmod (X n)})"
-
-apply (simp add: hcmod_def)
-apply (rule_tac f = Abs_star in arg_cong)
-apply (auto iff: starrel_iff, ultra)
-done
+lemma hcmod: "hcmod (star_n X) = star_n (%n. cmod (X n))"
+by (simp add: hcmod_def starfun)
 
 lemma hcmod_zero [simp]: "hcmod(0) = 0"
-by (simp add: hypreal_zero_def hypreal_zero_def hcmod)
+by (simp add: star_n_zero_num hcmod)
 
 lemma hcmod_one [simp]: "hcmod(1) = 1"
-by (simp add: hypreal_one_def hcmod hypreal_one_num)
+by (simp add: hcmod star_n_one_num)
 
 lemma hcmod_hcomplex_of_hypreal [simp]: "hcmod(hcomplex_of_hypreal x) = abs x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (auto simp add: hcmod hcomplex_of_hypreal hypreal_hrabs)
+apply (cases x)
+apply (auto simp add: hcmod hcomplex_of_hypreal star_n_abs)
 done
 
 lemma hcomplex_of_hypreal_abs:
@@ -496,36 +273,32 @@
       hcomplex_of_hypreal(hcmod(hcomplex_of_hypreal x))"
 by simp
 
-lemma HComplex_inject [simp]: "HComplex x y = HComplex x' y' = (x=x' & y=y')"
-apply (rule iffI) 
- prefer 2 apply simp 
-apply (simp add: HComplex_def iii_def) 
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star,
-       rule_tac z=x' in eq_Abs_star, rule_tac z=y' in eq_Abs_star)
-apply (auto simp add: iii_def hcomplex_mult hcomplex_add hcomplex_of_hypreal)
-apply (ultra+) 
-done
+lemma HComplex_inject [simp]:
+  "!!x y x' y'. HComplex x y = HComplex x' y' = (x=x' & y=y')"
+by (unfold HComplex_def, transfer, simp)
 
 lemma HComplex_add [simp]:
-     "HComplex x1 y1 + HComplex x2 y2 = HComplex (x1+x2) (y1+y2)"
-by (simp add: HComplex_def add_ac right_distrib) 
+  "!!x1 y1 x2 y2. HComplex x1 y1 + HComplex x2 y2 = HComplex (x1+x2) (y1+y2)"
+by (unfold HComplex_def, transfer, simp)
 
-lemma HComplex_minus [simp]: "- HComplex x y = HComplex (-x) (-y)"
-by (simp add: HComplex_def hcomplex_of_hypreal_minus) 
+lemma HComplex_minus [simp]: "!!x y. - HComplex x y = HComplex (-x) (-y)"
+by (unfold HComplex_def, transfer, simp)
 
 lemma HComplex_diff [simp]:
-     "HComplex x1 y1 - HComplex x2 y2 = HComplex (x1-x2) (y1-y2)"
-by (simp add: diff_minus)
+  "!!x1 y1 x2 y2. HComplex x1 y1 - HComplex x2 y2 = HComplex (x1-x2) (y1-y2)"
+by (unfold HComplex_def, transfer, rule complex_diff)
 
 lemma HComplex_mult [simp]:
-  "HComplex x1 y1 * HComplex x2 y2 = HComplex (x1*x2 - y1*y2) (x1*y2 + y1*x2)"
-by (simp add: HComplex_def diff_minus hcomplex_of_hypreal_minus 
-       add_ac mult_ac right_distrib)
+  "!!x1 y1 x2 y2. HComplex x1 y1 * HComplex x2 y2 =
+   HComplex (x1*x2 - y1*y2) (x1*y2 + y1*x2)"
+by (unfold HComplex_def, transfer, rule complex_mult)
 
 (*HComplex_inverse is proved below*)
 
-lemma hcomplex_of_hypreal_eq: "hcomplex_of_hypreal r = HComplex r 0"
-by (simp add: HComplex_def)
+lemma hcomplex_of_hypreal_eq: "!!r. hcomplex_of_hypreal r = HComplex r 0"
+apply (unfold hcomplex_of_hypreal_def HComplex_def, transfer)
+apply (simp add: complex_of_real_def)
+done
 
 lemma HComplex_add_hcomplex_of_hypreal [simp]:
      "HComplex x y + hcomplex_of_hypreal r = HComplex (x+r) y"
@@ -544,226 +317,149 @@
 by (simp add: i_def hcomplex_of_hypreal_eq)
 
 lemma i_hcomplex_of_hypreal [simp]:
-     "iii * hcomplex_of_hypreal r = HComplex 0 r"
-by (simp add: HComplex_def)
+     "!!r. iii * hcomplex_of_hypreal r = HComplex 0 r"
+by (unfold iii_def hcomplex_of_hypreal_def HComplex_def, transfer, rule i_complex_of_real)
 
 lemma hcomplex_of_hypreal_i [simp]:
-     "hcomplex_of_hypreal r * iii = HComplex 0 r"
-by (simp add: mult_commute) 
+     "!!r. hcomplex_of_hypreal r * iii = HComplex 0 r"
+by (unfold iii_def hcomplex_of_hypreal_def HComplex_def, transfer, rule complex_of_real_i)
 
 
 subsection{*Conjugation*}
 
-lemma hcnj:
-  "hcnj (Abs_star(starrel `` {%n. X n})) =
-   Abs_star(starrel `` {%n. cnj(X n)})"
-apply (simp add: hcnj_def)
-apply (rule_tac f = Abs_star in arg_cong)
-apply (auto iff: starrel_iff, ultra)
-done
+lemma hcnj: "hcnj (star_n X) = star_n (%n. cnj(X n))"
+by (simp add: hcnj_def starfun)
 
-lemma hcomplex_hcnj_cancel_iff [iff]: "(hcnj x = hcnj y) = (x = y)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hcnj)
-done
+lemma hcomplex_hcnj_cancel_iff [iff]: "!!x y. (hcnj x = hcnj y) = (x = y)"
+by (unfold hcnj_def, transfer, rule complex_cnj_cancel_iff)
 
-lemma hcomplex_hcnj_hcnj [simp]: "hcnj (hcnj z) = z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hcnj)
-done
+lemma hcomplex_hcnj_hcnj [simp]: "!!z. hcnj (hcnj z) = z"
+by (unfold hcnj_def, transfer, rule complex_cnj_cnj)
 
 lemma hcomplex_hcnj_hcomplex_of_hypreal [simp]:
-     "hcnj (hcomplex_of_hypreal x) = hcomplex_of_hypreal x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: hcnj hcomplex_of_hypreal)
-done
+     "!!x. hcnj (hcomplex_of_hypreal x) = hcomplex_of_hypreal x"
+by (unfold hcnj_def hcomplex_of_hypreal_def, transfer, rule complex_cnj_complex_of_real)
+
+lemma hcomplex_hmod_hcnj [simp]: "!!z. hcmod (hcnj z) = hcmod z"
+by (unfold hcmod_def hcnj_def, transfer, rule complex_mod_cnj)
 
-lemma hcomplex_hmod_hcnj [simp]: "hcmod (hcnj z) = hcmod z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hcnj hcmod)
-done
+lemma hcomplex_hcnj_minus: "!!z. hcnj (-z) = - hcnj z"
+by (unfold hcnj_def, transfer, rule complex_cnj_minus)
 
-lemma hcomplex_hcnj_minus: "hcnj (-z) = - hcnj z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hcnj hcomplex_minus complex_cnj_minus)
-done
+lemma hcomplex_hcnj_inverse: "!!z. hcnj(inverse z) = inverse(hcnj z)"
+by (unfold hcnj_def, transfer, rule complex_cnj_inverse)
 
-lemma hcomplex_hcnj_inverse: "hcnj(inverse z) = inverse(hcnj z)"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hcnj hcomplex_inverse complex_cnj_inverse)
-done
+lemma hcomplex_hcnj_add: "!!w z. hcnj(w + z) = hcnj(w) + hcnj(z)"
+by (unfold hcnj_def, transfer, rule complex_cnj_add)
 
-lemma hcomplex_hcnj_add: "hcnj(w + z) = hcnj(w) + hcnj(z)"
-apply (rule_tac z=z in eq_Abs_star, rule_tac z=w in eq_Abs_star)
-apply (simp add: hcnj hcomplex_add complex_cnj_add)
-done
+lemma hcomplex_hcnj_diff: "!!w z. hcnj(w - z) = hcnj(w) - hcnj(z)"
+by (unfold hcnj_def, transfer, rule complex_cnj_diff)
 
-lemma hcomplex_hcnj_diff: "hcnj(w - z) = hcnj(w) - hcnj(z)"
-apply (rule_tac z=z in eq_Abs_star, rule_tac z=w in eq_Abs_star)
-apply (simp add: hcnj hcomplex_diff complex_cnj_diff)
-done
+lemma hcomplex_hcnj_mult: "!!w z. hcnj(w * z) = hcnj(w) * hcnj(z)"
+by (unfold hcnj_def, transfer, rule complex_cnj_mult)
 
-lemma hcomplex_hcnj_mult: "hcnj(w * z) = hcnj(w) * hcnj(z)"
-apply (rule_tac z=z in eq_Abs_star, rule_tac z=w in eq_Abs_star)
-apply (simp add: hcnj hcomplex_mult complex_cnj_mult)
-done
-
-lemma hcomplex_hcnj_divide: "hcnj(w / z) = (hcnj w)/(hcnj z)"
-by (simp add: divide_inverse hcomplex_hcnj_mult hcomplex_hcnj_inverse)
+lemma hcomplex_hcnj_divide: "!!w z. hcnj(w / z) = (hcnj w)/(hcnj z)"
+by (unfold hcnj_def, transfer, rule complex_cnj_divide)
 
 lemma hcnj_one [simp]: "hcnj 1 = 1"
-by (simp add: hypreal_one_def hcnj)
+by (unfold hcnj_def, transfer, rule complex_cnj_one)
 
 lemma hcomplex_hcnj_zero [simp]: "hcnj 0 = 0"
-by (simp add: hypreal_zero_def hcnj)
+by (unfold hcnj_def, transfer, rule complex_cnj_zero)
 
-lemma hcomplex_hcnj_zero_iff [iff]: "(hcnj z = 0) = (z = 0)"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hypreal_zero_def hcnj)
-done
+lemma hcomplex_hcnj_zero_iff [iff]: "!!z. (hcnj z = 0) = (z = 0)"
+by (unfold hcnj_def, transfer, rule complex_cnj_zero_iff)
 
 lemma hcomplex_mult_hcnj:
-     "z * hcnj z = hcomplex_of_hypreal (hRe(z) ^ 2 + hIm(z) ^ 2)"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hcnj hcomplex_mult hcomplex_of_hypreal hRe hIm hypreal_add
-                      hypreal_mult complex_mult_cnj numeral_2_eq_2)
+     "!!z. z * hcnj z = hcomplex_of_hypreal (hRe(z) ^ 2 + hIm(z) ^ 2)"
+apply (unfold hcnj_def hcomplex_of_hypreal_def hRe_def hIm_def)
+apply (transfer, rule complex_mult_cnj)
 done
 
 
 subsection{*More Theorems about the Function @{term hcmod}*}
 
-lemma hcomplex_hcmod_eq_zero_cancel [simp]: "(hcmod x = 0) = (x = 0)"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: hcmod hypreal_zero_def hypreal_zero_num)
-done
+lemma hcomplex_hcmod_eq_zero_cancel [simp]: "!!x. (hcmod x = 0) = (x = 0)"
+by (unfold hcmod_def, transfer, rule complex_mod_eq_zero_cancel)
 
 lemma hcmod_hcomplex_of_hypreal_of_nat [simp]:
      "hcmod (hcomplex_of_hypreal(hypreal_of_nat n)) = hypreal_of_nat n"
-apply (simp add: abs_if linorder_not_less)
-done
+by (simp add: abs_if linorder_not_less)
 
 lemma hcmod_hcomplex_of_hypreal_of_hypnat [simp]:
      "hcmod (hcomplex_of_hypreal(hypreal_of_hypnat n)) = hypreal_of_hypnat n"
-apply (simp add: abs_if linorder_not_less)
-done
+by (simp add: abs_if linorder_not_less)
 
-lemma hcmod_minus [simp]: "hcmod (-x) = hcmod(x)"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: hcmod hcomplex_minus)
-done
+lemma hcmod_minus [simp]: "!!x. hcmod (-x) = hcmod(x)"
+by (unfold hcmod_def, transfer, rule complex_mod_minus)
 
-lemma hcmod_mult_hcnj: "hcmod(z * hcnj(z)) = hcmod(z) ^ 2"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hcmod hcomplex_mult hcnj hypreal_mult complex_mod_mult_cnj numeral_2_eq_2)
-done
+lemma hcmod_mult_hcnj: "!!z. hcmod(z * hcnj(z)) = hcmod(z) ^ 2"
+by (unfold hcmod_def hcnj_def, transfer, rule complex_mod_mult_cnj)
 
-lemma hcmod_ge_zero [simp]: "(0::hypreal) \<le> hcmod x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: hcmod hypreal_zero_num hypreal_le)
-done
+lemma hcmod_ge_zero [simp]: "!!x. (0::hypreal) \<le> hcmod x"
+by (unfold hcmod_def, transfer, rule complex_mod_ge_zero)
 
 lemma hrabs_hcmod_cancel [simp]: "abs(hcmod x) = hcmod x"
 by (simp add: abs_if linorder_not_less)
 
-lemma hcmod_mult: "hcmod(x*y) = hcmod(x) * hcmod(y)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hcmod hcomplex_mult hypreal_mult complex_mod_mult)
-done
+lemma hcmod_mult: "!!x y. hcmod(x*y) = hcmod(x) * hcmod(y)"
+by (unfold hcmod_def, transfer, rule complex_mod_mult)
 
 lemma hcmod_add_squared_eq:
-     "hcmod(x + y) ^ 2 = hcmod(x) ^ 2 + hcmod(y) ^ 2 + 2 * hRe(x * hcnj y)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hcmod hcomplex_add hypreal_mult hRe hcnj hcomplex_mult
-                      numeral_2_eq_2 realpow_two [symmetric]
-                  del: realpow_Suc)
-apply (simp add: numeral_2_eq_2 [symmetric] complex_mod_add_squared_eq
-                 hypreal_add [symmetric] hypreal_mult [symmetric]
-                 star_n_def [symmetric] star_of_def [symmetric]
-                 hypreal_of_real_def [symmetric])
-done
+  "!!x y. hcmod(x + y) ^ 2 = hcmod(x) ^ 2 + hcmod(y) ^ 2 + 2 * hRe(x * hcnj y)"
+by (unfold hcmod_def hcnj_def hRe_def, transfer, rule complex_mod_add_squared_eq)
 
-lemma hcomplex_hRe_mult_hcnj_le_hcmod [simp]: "hRe(x * hcnj y) \<le> hcmod(x * hcnj y)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hcmod hcnj hcomplex_mult hRe hypreal_le)
-done
+lemma hcomplex_hRe_mult_hcnj_le_hcmod [simp]:
+  "!!x y. hRe(x * hcnj y) \<le> hcmod(x * hcnj y)"
+by (unfold hcmod_def hcnj_def hRe_def, transfer, simp)
 
-lemma hcomplex_hRe_mult_hcnj_le_hcmod2 [simp]: "hRe(x * hcnj y) \<le> hcmod(x * y)"
-apply (cut_tac x = x and y = y in hcomplex_hRe_mult_hcnj_le_hcmod)
-apply (simp add: hcmod_mult)
-done
+lemma hcomplex_hRe_mult_hcnj_le_hcmod2 [simp]:
+  "!!x y. hRe(x * hcnj y) \<le> hcmod(x * y)"
+by (unfold hcmod_def hcnj_def hRe_def, transfer, simp)
 
-lemma hcmod_triangle_squared [simp]: "hcmod (x + y) ^ 2 \<le> (hcmod(x) + hcmod(y)) ^ 2"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hcmod hcnj hcomplex_add hypreal_mult hypreal_add
-                      hypreal_le realpow_two [symmetric] numeral_2_eq_2
-            del: realpow_Suc)
-apply (simp add: numeral_2_eq_2 [symmetric])
-done
+lemma hcmod_triangle_squared [simp]:
+  "!!x y. hcmod (x + y) ^ 2 \<le> (hcmod(x) + hcmod(y)) ^ 2"
+by (unfold hcmod_def, transfer, simp)
 
-lemma hcmod_triangle_ineq [simp]: "hcmod (x + y) \<le> hcmod(x) + hcmod(y)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hcmod hcomplex_add hypreal_add hypreal_le)
-done
+lemma hcmod_triangle_ineq [simp]:
+  "!!x y. hcmod (x + y) \<le> hcmod(x) + hcmod(y)"
+by (unfold hcmod_def, transfer, simp)
 
-lemma hcmod_triangle_ineq2 [simp]: "hcmod(b + a) - hcmod b \<le> hcmod a"
-apply (cut_tac x1 = b and y1 = a and c = "-hcmod b" in hcmod_triangle_ineq [THEN add_right_mono])
-apply (simp add: add_ac)
-done
+lemma hcmod_triangle_ineq2 [simp]:
+  "!!a b. hcmod(b + a) - hcmod b \<le> hcmod a"
+by (unfold hcmod_def, transfer, simp)
 
-lemma hcmod_diff_commute: "hcmod (x - y) = hcmod (y - x)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hcmod hcomplex_diff complex_mod_diff_commute)
-done
+lemma hcmod_diff_commute: "!!x y. hcmod (x - y) = hcmod (y - x)"
+by (unfold hcmod_def, transfer, rule complex_mod_diff_commute)
 
 lemma hcmod_add_less:
-     "[| hcmod x < r; hcmod y < s |] ==> hcmod (x + y) < r + s"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (rule_tac z=r in eq_Abs_star, rule_tac z=s in eq_Abs_star)
-apply (simp add: hcmod hcomplex_add hypreal_add hypreal_less, ultra)
-apply (auto intro: complex_mod_add_less)
-done
+  "!!x y r s. [| hcmod x < r; hcmod y < s |] ==> hcmod (x + y) < r + s"
+by (unfold hcmod_def, transfer, rule complex_mod_add_less)
 
 lemma hcmod_mult_less:
-     "[| hcmod x < r; hcmod y < s |] ==> hcmod (x * y) < r * s"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (rule_tac z=r in eq_Abs_star, rule_tac z=s in eq_Abs_star)
-apply (simp add: hcmod hypreal_mult hypreal_less hcomplex_mult, ultra)
-apply (auto intro: complex_mod_mult_less)
-done
+  "!!x y r s. [| hcmod x < r; hcmod y < s |] ==> hcmod (x * y) < r * s"
+by (unfold hcmod_def, transfer, rule complex_mod_mult_less)
 
-lemma hcmod_diff_ineq [simp]: "hcmod(a) - hcmod(b) \<le> hcmod(a + b)"
-apply (rule_tac z=a in eq_Abs_star, rule_tac z=b in eq_Abs_star)
-apply (simp add: hcmod hcomplex_add hypreal_diff hypreal_le)
-done
+lemma hcmod_diff_ineq [simp]: "!!a b. hcmod(a) - hcmod(b) \<le> hcmod(a + b)"
+by (unfold hcmod_def, transfer, simp)
 
 
 subsection{*A Few Nonlinear Theorems*}
 
-lemma hcpow:
-  "Abs_star(starrel``{%n. X n}) hcpow
-   Abs_star(starrel``{%n. Y n}) =
-   Abs_star(starrel``{%n. X n ^ Y n})"
-apply (simp add: hcpow_def)
-apply (rule_tac f = Abs_star in arg_cong)
-apply (auto iff: starrel_iff, ultra)
-done
+lemma hcpow: "star_n X hcpow star_n Y = star_n (%n. X n ^ Y n)"
+by (simp add: hcpow_def Ifun2_of_def star_of_def Ifun_star_n)
 
 lemma hcomplex_of_hypreal_hyperpow:
-     "hcomplex_of_hypreal (x pow n) = (hcomplex_of_hypreal x) hcpow n"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (simp add: hcomplex_of_hypreal hyperpow hcpow complex_of_real_pow)
+     "!!x n. hcomplex_of_hypreal (x pow n) = (hcomplex_of_hypreal x) hcpow n"
+apply (unfold hcomplex_of_hypreal_def hyperpow_def hcpow_def)
+apply (transfer, rule complex_of_real_pow)
 done
 
-lemma hcmod_hcpow: "hcmod(x hcpow n) = hcmod(x) pow n"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (simp add: hcpow hyperpow hcmod complex_mod_complexpow)
-done
+lemma hcmod_hcpow: "!!x n. hcmod(x hcpow n) = hcmod(x) pow n"
+by (unfold hcmod_def hcpow_def hyperpow_def, transfer, rule complex_mod_complexpow)
 
-lemma hcmod_hcomplex_inverse: "hcmod(inverse x) = inverse(hcmod x)"
-apply (case_tac "x = 0", simp)
-apply (rule_tac c1 = "hcmod x" in hypreal_mult_left_cancel [THEN iffD1])
-apply (auto simp add: hcmod_mult [symmetric])
-done
+lemma hcmod_hcomplex_inverse: "!!x. hcmod(inverse x) = inverse(hcmod x)"
+by (unfold hcmod_def, transfer, rule complex_mod_inverse)
 
 lemma hcmod_divide: "hcmod(x/y) = hcmod(x)/(hcmod y)"
 by (simp add: divide_inverse hcmod_mult hcmod_hcomplex_inverse)
@@ -777,15 +473,6 @@
 lemma hcomplexpow_Suc [simp]: "z ^ (Suc n) = (z::hcomplex) * (z ^ n)"
 by (rule power_Suc)
 
-(*
-instance hcomplex :: recpower
-proof
-  fix z :: hcomplex
-  fix n :: nat
-  show "z^0 = 1" by simp
-  show "z^(Suc n) = z * (z^n)" by simp
-qed
-*)
 lemma hcomplexpow_i_squared [simp]: "iii ^ 2 = - 1"
 by (simp add: power2_eq_square)
 
@@ -807,74 +494,59 @@
 done
 
 lemma hcpow_minus:
-     "(-x::hcomplex) hcpow n =
-      (if ( *pNat* even) n then (x hcpow n) else -(x hcpow n))"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (auto simp add: hcpow hyperpow starPNat hcomplex_minus, ultra)
-apply (auto simp add: neg_power_if, ultra)
-done
+     "!!x n. (-x::hcomplex) hcpow n =
+      (if ( *p* even) n then (x hcpow n) else -(x hcpow n))"
+by (unfold hcpow_def, transfer, rule neg_power_if)
 
-lemma hcpow_mult: "((r::hcomplex) * s) hcpow n = (r hcpow n) * (s hcpow n)"
-apply (rule_tac z=r in eq_Abs_star, rule_tac z=s in eq_Abs_star)
-apply (rule_tac z=n in eq_Abs_star)
-apply (simp add: hcpow hypreal_mult hcomplex_mult power_mult_distrib)
-done
+lemma hcpow_mult:
+  "!!r s n. ((r::hcomplex) * s) hcpow n = (r hcpow n) * (s hcpow n)"
+by (unfold hcpow_def, transfer, rule power_mult_distrib)
 
 lemma hcpow_zero [simp]: "0 hcpow (n + 1) = 0"
-apply (simp add: hypreal_zero_def hypnat_one_def, rule_tac z=n in eq_Abs_star)
-apply (simp add: hcpow hypnat_add)
+apply (simp add: star_n_zero_num star_n_one_num, cases n)
+apply (simp add: hcpow star_n_add)
 done
 
 lemma hcpow_zero2 [simp]: "0 hcpow (hSuc n) = 0"
 by (simp add: hSuc_def)
 
 lemma hcpow_not_zero [simp,intro]: "r \<noteq> 0 ==> r hcpow n \<noteq> (0::hcomplex)"
-apply (rule_tac z=r in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (auto simp add: hcpow hypreal_zero_def, ultra)
+apply (cases r, cases n)
+apply (auto simp add: hcpow star_n_zero_num star_n_eq_iff, ultra)
 done
 
 lemma hcpow_zero_zero: "r hcpow n = (0::hcomplex) ==> r = 0"
 by (blast intro: ccontr dest: hcpow_not_zero)
 
-lemma hcomplex_divide:
-  "Abs_star(starrel``{%n. X n::complex}) / Abs_star(starrel``{%n. Y n}) =
-   Abs_star(starrel``{%n. X n / Y n})"
-by (simp add: divide_inverse complex_divide_def hcomplex_inverse hcomplex_mult)
-
-
-
+lemma star_n_divide: "star_n X / star_n Y = star_n (%n. X n / Y n)"
+by (simp add: star_divide_def Ifun2_of_def star_of_def Ifun_star_n)
 
 subsection{*The Function @{term hsgn}*}
 
-lemma hsgn:
-  "hsgn (Abs_star(starrel `` {%n. X n})) =
-      Abs_star(starrel `` {%n. sgn (X n)})"
-apply (simp add: hsgn_def)
-apply (rule_tac f = Abs_star in arg_cong)
-apply (auto iff: starrel_iff, ultra)
-done
+lemma hsgn: "hsgn (star_n X) = star_n (%n. sgn (X n))"
+by (simp add: hsgn_def starfun)
 
 lemma hsgn_zero [simp]: "hsgn 0 = 0"
-by (simp add: hypreal_zero_def hsgn)
+by (simp add: star_n_zero_num hsgn)
 
 lemma hsgn_one [simp]: "hsgn 1 = 1"
-by (simp add: hypreal_one_def hsgn)
+by (simp add: star_n_one_num hsgn)
 
 lemma hsgn_minus: "hsgn (-z) = - hsgn(z)"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hsgn hcomplex_minus sgn_minus)
+apply (cases z)
+apply (simp add: hsgn star_n_minus sgn_minus)
 done
 
 lemma hsgn_eq: "hsgn z = z / hcomplex_of_hypreal (hcmod z)"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hsgn hcomplex_divide hcomplex_of_hypreal hcmod sgn_eq)
+apply (cases z)
+apply (simp add: hsgn star_n_divide hcomplex_of_hypreal hcmod sgn_eq)
 done
 
 
 lemma hcmod_i: "hcmod (HComplex x y) = ( *f* sqrt) (x ^ 2 + y ^ 2)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star) 
+apply (cases x, cases y) 
 apply (simp add: HComplex_eq_Abs_star_Complex starfun 
-                 hypreal_mult hypreal_add hcmod numeral_2_eq_2)
+                 star_n_mult star_n_add hcmod numeral_2_eq_2)
 done
 
 lemma hcomplex_eq_cancel_iff1 [simp]:
@@ -898,13 +570,13 @@
 by (simp add: i_eq_HComplex_0_1) 
 
 lemma hRe_hsgn [simp]: "hRe(hsgn z) = hRe(z)/hcmod z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hsgn hcmod hRe hypreal_divide)
+apply (cases z)
+apply (simp add: hsgn hcmod hRe star_n_divide)
 done
 
 lemma hIm_hsgn [simp]: "hIm(hsgn z) = hIm(z)/hcmod z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: hsgn hcmod hIm hypreal_divide)
+apply (cases z)
+apply (simp add: hsgn hcmod hIm star_n_divide)
 done
 
 (*????move to RealDef????*)
@@ -912,68 +584,46 @@
 by (auto intro: real_sum_squares_cancel iff: real_add_eq_0_iff)
 
 lemma hcomplex_inverse_complex_split:
-     "inverse(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y) =
+     "!!x y. inverse(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y) =
       hcomplex_of_hypreal(x/(x ^ 2 + y ^ 2)) -
       iii * hcomplex_of_hypreal(y/(x ^ 2 + y ^ 2))"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hcomplex_of_hypreal hcomplex_mult hcomplex_add iii_def
-         starfun hypreal_mult hypreal_add hcomplex_inverse hypreal_divide
-         hcomplex_diff numeral_2_eq_2 complex_of_real_def i_def)
-apply (simp add: diff_minus) 
+apply (unfold hcomplex_of_hypreal_def iii_def)
+apply (transfer, rule complex_inverse_complex_split)
 done
 
 lemma HComplex_inverse:
-     "inverse (HComplex x y) =
+     "!!x y. inverse (HComplex x y) =
       HComplex (x/(x ^ 2 + y ^ 2)) (-y/(x ^ 2 + y ^ 2))"
-by (simp only: HComplex_def hcomplex_inverse_complex_split, simp)
-
-
+by (unfold HComplex_def, transfer, rule complex_inverse)
 
 lemma hRe_mult_i_eq[simp]:
-    "hRe (iii * hcomplex_of_hypreal y) = 0"
-apply (simp add: iii_def, rule_tac z=y in eq_Abs_star)
-apply (simp add: hcomplex_of_hypreal hcomplex_mult hRe hypreal_zero_num)
-done
+    "!!y. hRe (iii * hcomplex_of_hypreal y) = 0"
+by (unfold hRe_def iii_def hcomplex_of_hypreal_def, transfer, simp)
 
 lemma hIm_mult_i_eq [simp]:
-    "hIm (iii * hcomplex_of_hypreal y) = y"
-apply (simp add: iii_def, rule_tac z=y in eq_Abs_star)
-apply (simp add: hcomplex_of_hypreal hcomplex_mult hIm hypreal_zero_num)
-done
+    "!!y. hIm (iii * hcomplex_of_hypreal y) = y"
+by (unfold hIm_def iii_def hcomplex_of_hypreal_def, transfer, simp)
 
-lemma hcmod_mult_i [simp]: "hcmod (iii * hcomplex_of_hypreal y) = abs y"
-apply (rule_tac z=y in eq_Abs_star)
-apply (simp add: hcomplex_of_hypreal hcmod hypreal_hrabs iii_def hcomplex_mult)
-done
+lemma hcmod_mult_i [simp]: "!!y. hcmod (iii * hcomplex_of_hypreal y) = abs y"
+by (unfold hcmod_def iii_def hcomplex_of_hypreal_def, transfer, simp)
 
 lemma hcmod_mult_i2 [simp]: "hcmod (hcomplex_of_hypreal y * iii) = abs y"
-by (simp only: hcmod_mult_i hcomplex_mult_commute)
+by (simp only: hcmod_mult_i mult_commute)
 
 (*---------------------------------------------------------------------------*)
 (*  harg                                                                     *)
 (*---------------------------------------------------------------------------*)
 
-lemma harg:
-  "harg (Abs_star(starrel `` {%n. X n})) =
-      Abs_star(starrel `` {%n. arg (X n)})"
-apply (simp add: harg_def)
-apply (rule_tac f = Abs_star in arg_cong)
-apply (auto iff: starrel_iff, ultra)
-done
+lemma harg: "harg (star_n X) = star_n (%n. arg (X n))"
+by (simp add: harg_def starfun)
 
 lemma cos_harg_i_mult_zero_pos:
-     "0 < y ==> ( *f* cos) (harg(HComplex 0 y)) = 0"
-apply (rule_tac z=y in eq_Abs_star)
-apply (simp add: HComplex_def hcomplex_of_hypreal iii_def hcomplex_mult
-                hcomplex_add hypreal_zero_num hypreal_less starfun harg, ultra)
-done
+     "!!y. 0 < y ==> ( *f* cos) (harg(HComplex 0 y)) = 0"
+by (unfold harg_def HComplex_def, transfer, rule cos_arg_i_mult_zero_pos)
 
 lemma cos_harg_i_mult_zero_neg:
-     "y < 0 ==> ( *f* cos) (harg(HComplex 0 y)) = 0"
-apply (rule_tac z=y in eq_Abs_star)
-apply (simp add: HComplex_def hcomplex_of_hypreal iii_def hcomplex_mult
-                 hcomplex_add hypreal_zero_num hypreal_less starfun harg, ultra)
-done
+     "!!y. y < 0 ==> ( *f* cos) (harg(HComplex 0 y)) = 0"
+by (unfold harg_def HComplex_def, transfer, rule cos_arg_i_mult_zero_neg)
 
 lemma cos_harg_i_mult_zero [simp]:
      "y \<noteq> 0 ==> ( *f* cos) (harg(HComplex 0 y)) = 0"
@@ -981,10 +631,8 @@
                    cos_harg_i_mult_zero_pos cos_harg_i_mult_zero_neg)
 
 lemma hcomplex_of_hypreal_zero_iff [simp]:
-     "(hcomplex_of_hypreal y = 0) = (y = 0)"
-apply (rule_tac z=y in eq_Abs_star)
-apply (simp add: hcomplex_of_hypreal hypreal_zero_num hypreal_zero_def)
-done
+     "!!y. (hcomplex_of_hypreal y = 0) = (y = 0)"
+by (unfold hcomplex_of_hypreal_def, transfer, simp)
 
 
 subsection{*Polar Form for Nonstandard Complex Numbers*}
@@ -995,41 +643,29 @@
 
 lemma lemma_hypreal_P_EX2:
      "(\<exists>(x::hypreal) y. P x y) =
-      (\<exists>f g. P (Abs_star(starrel `` {f})) (Abs_star(starrel `` {g})))"
+      (\<exists>f g. P (star_n f) (star_n g))"
 apply auto
-apply (rule_tac z = x in eq_Abs_star)
-apply (rule_tac z = y in eq_Abs_star, auto)
+apply (rule_tac x = x in star_cases)
+apply (rule_tac x = y in star_cases, auto)
 done
 
 lemma hcomplex_split_polar:
-  "\<exists>r a. z = hcomplex_of_hypreal r * (HComplex(( *f* cos) a)(( *f* sin) a))"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: lemma_hypreal_P_EX2 hcomplex_of_hypreal iii_def starfun hcomplex_add hcomplex_mult HComplex_def)
-apply (cut_tac z = x in complex_split_polar2)
-apply (drule choice, safe)+
-apply (rule_tac x = f in exI)
-apply (rule_tac x = fa in exI, auto)
-done
+  "!!z. \<exists>r a. z = hcomplex_of_hypreal r * (HComplex(( *f* cos) a)(( *f* sin) a))"
+by (unfold hcomplex_of_hypreal_def HComplex_def, transfer, rule complex_split_polar)
 
-lemma hcis:
-  "hcis (Abs_star(starrel `` {%n. X n})) =
-      Abs_star(starrel `` {%n. cis (X n)})"
-apply (simp add: hcis_def)
-apply (rule_tac f = Abs_star in arg_cong, auto iff: starrel_iff, ultra)
-done
+lemma hcis: "hcis (star_n X) = star_n (%n. cis (X n))"
+by (simp add: hcis_def starfun)
 
 lemma hcis_eq:
    "hcis a =
     (hcomplex_of_hypreal(( *f* cos) a) +
     iii * hcomplex_of_hypreal(( *f* sin) a))"
-apply (rule_tac z=a in eq_Abs_star)
-apply (simp add: starfun hcis hcomplex_of_hypreal iii_def hcomplex_mult hcomplex_add cis_def)
+apply (cases a)
+apply (simp add: starfun hcis hcomplex_of_hypreal iii_def star_of_def star_n_mult star_n_add cis_def star_n_eq_iff)
 done
 
-lemma hrcis:
-  "hrcis (Abs_star(starrel `` {%n. X n})) (Abs_star(starrel `` {%n. Y n})) =
-      Abs_star(starrel `` {%n. rcis (X n) (Y n)})"
-by (simp add: hrcis_def hcomplex_of_hypreal hcomplex_mult hcis rcis_def)
+lemma hrcis: "hrcis (star_n X) (star_n Y) = star_n (%n. rcis (X n) (Y n))"
+by (simp add: hrcis_def hcomplex_of_hypreal star_n_mult hcis rcis_def)
 
 lemma hrcis_Ex: "\<exists>r a. z = hrcis r a"
 apply (simp add: hrcis_def hcis_eq hcomplex_of_hypreal_mult_HComplex [symmetric])
@@ -1054,11 +690,8 @@
 
 
 lemma hcmod_unit_one [simp]:
-     "hcmod (HComplex (( *f* cos) a) (( *f* sin) a)) = 1"
-apply (rule_tac z=a in eq_Abs_star) 
-apply (simp add: HComplex_def iii_def starfun hcomplex_of_hypreal 
-                 hcomplex_mult hcmod hcomplex_add hypreal_one_def)
-done
+     "!!a. hcmod (HComplex (( *f* cos) a) (( *f* sin) a)) = 1"
+by (unfold hcmod_def HComplex_def, transfer, simp)
 
 lemma hcmod_complex_polar [simp]:
      "hcmod (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) =
@@ -1079,39 +712,39 @@
 
 lemma hrcis_mult:
   "hrcis r1 a * hrcis r2 b = hrcis (r1*r2) (a + b)"
-apply (simp add: hrcis_def, rule_tac z=r1 in eq_Abs_star, rule_tac z=r2 in eq_Abs_star, rule_tac z=a in eq_Abs_star, rule_tac z=b in eq_Abs_star)
-apply (simp add: hrcis hcis hypreal_add hypreal_mult hcomplex_of_hypreal
-                      hcomplex_mult cis_mult [symmetric])
+apply (simp add: hrcis_def, rule_tac z=r1 in eq_Abs_star, rule_tac z=r2 in eq_Abs_star, cases a, cases b)
+apply (simp add: hrcis hcis star_n_add star_n_mult hcomplex_of_hypreal
+                      star_n_mult cis_mult [symmetric])
 done
 
 lemma hcis_mult: "hcis a * hcis b = hcis (a + b)"
-apply (rule_tac z=a in eq_Abs_star, rule_tac z=b in eq_Abs_star)
-apply (simp add: hcis hcomplex_mult hypreal_add cis_mult)
+apply (cases a, cases b)
+apply (simp add: hcis star_n_mult star_n_add cis_mult)
 done
 
 lemma hcis_zero [simp]: "hcis 0 = 1"
-by (simp add: hypreal_one_def hcis hypreal_zero_num)
+by (simp add: star_n_one_num hcis star_n_zero_num)
 
 lemma hrcis_zero_mod [simp]: "hrcis 0 a = 0"
-apply (simp add: hypreal_zero_def, rule_tac z=a in eq_Abs_star)
-apply (simp add: hrcis hypreal_zero_num)
+apply (simp add: star_n_zero_num, cases a)
+apply (simp add: hrcis star_n_zero_num)
 done
 
 lemma hrcis_zero_arg [simp]: "hrcis r 0 = hcomplex_of_hypreal r"
-apply (rule_tac z=r in eq_Abs_star)
-apply (simp add: hrcis hypreal_zero_num hcomplex_of_hypreal)
+apply (cases r)
+apply (simp add: hrcis star_n_zero_num hcomplex_of_hypreal)
 done
 
 lemma hcomplex_i_mult_minus [simp]: "iii * (iii * x) = - x"
-by (simp add: hcomplex_mult_assoc [symmetric])
+by (simp add: mult_assoc [symmetric])
 
 lemma hcomplex_i_mult_minus2 [simp]: "iii * iii * x = - x"
 by simp
 
 lemma hcis_hypreal_of_nat_Suc_mult:
    "hcis (hypreal_of_nat (Suc n) * a) = hcis a * hcis (hypreal_of_nat n * a)"
-apply (rule_tac z=a in eq_Abs_star)
-apply (simp add: hypreal_of_nat hcis hypreal_mult hcomplex_mult cis_real_of_nat_Suc_mult)
+apply (cases a)
+apply (simp add: hypreal_of_nat hcis star_n_mult star_n_mult cis_real_of_nat_Suc_mult)
 done
 
 lemma NSDeMoivre: "(hcis a) ^ n = hcis (hypreal_of_nat n * a)"
@@ -1122,13 +755,13 @@
 lemma hcis_hypreal_of_hypnat_Suc_mult:
      "hcis (hypreal_of_hypnat (n + 1) * a) =
       hcis a * hcis (hypreal_of_hypnat n * a)"
-apply (rule_tac z=a in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (simp add: hcis hypreal_of_hypnat hypnat_add hypnat_one_def hypreal_mult hcomplex_mult cis_real_of_nat_Suc_mult)
+apply (cases a, cases n)
+apply (simp add: hcis hypreal_of_hypnat star_n_add star_n_one_num star_n_mult star_n_mult cis_real_of_nat_Suc_mult)
 done
 
 lemma NSDeMoivre_ext: "(hcis a) hcpow n = hcis (hypreal_of_hypnat n * a)"
-apply (rule_tac z=a in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (simp add: hcis hypreal_of_hypnat hypreal_mult hcpow DeMoivre)
+apply (cases a, cases n)
+apply (simp add: hcis hypreal_of_hypnat star_n_mult hcpow DeMoivre)
 done
 
 lemma DeMoivre2:
@@ -1142,23 +775,23 @@
 done
 
 lemma hcis_inverse [simp]: "inverse(hcis a) = hcis (-a)"
-apply (rule_tac z=a in eq_Abs_star)
-apply (simp add: hcomplex_inverse hcis hypreal_minus)
+apply (cases a)
+apply (simp add: star_n_inverse2 hcis star_n_minus)
 done
 
 lemma hrcis_inverse: "inverse(hrcis r a) = hrcis (inverse r) (-a)"
-apply (rule_tac z=a in eq_Abs_star, rule_tac z=r in eq_Abs_star)
-apply (simp add: hcomplex_inverse hrcis hypreal_minus hypreal_inverse rcis_inverse, ultra)
+apply (cases a, cases r)
+apply (simp add: star_n_inverse2 hrcis star_n_minus rcis_inverse star_n_eq_iff, ultra)
 apply (simp add: real_divide_def)
 done
 
 lemma hRe_hcis [simp]: "hRe(hcis a) = ( *f* cos) a"
-apply (rule_tac z=a in eq_Abs_star)
+apply (cases a)
 apply (simp add: hcis starfun hRe)
 done
 
 lemma hIm_hcis [simp]: "hIm(hcis a) = ( *f* sin) a"
-apply (rule_tac z=a in eq_Abs_star)
+apply (cases a)
 apply (simp add: hcis starfun hIm)
 done
 
@@ -1175,8 +808,8 @@
 by (simp add: NSDeMoivre_ext)
 
 lemma hexpi_add: "hexpi(a + b) = hexpi(a) * hexpi(b)"
-apply (simp add: hexpi_def, rule_tac z=a in eq_Abs_star, rule_tac z=b in eq_Abs_star)
-apply (simp add: hcis hRe hIm hcomplex_add hcomplex_mult hypreal_mult starfun hcomplex_of_hypreal cis_mult [symmetric] complex_Im_add complex_Re_add exp_add complex_of_real_mult)
+apply (simp add: hexpi_def, cases a, cases b)
+apply (simp add: hcis hRe hIm star_n_add star_n_mult star_n_mult starfun hcomplex_of_hypreal cis_mult [symmetric] complex_Im_add complex_Re_add exp_add complex_of_real_mult star_n_eq_iff)
 done
 
 
@@ -1184,93 +817,61 @@
   type @{typ complex} to to @{typ hcomplex}*}
 
 lemma inj_hcomplex_of_complex: "inj(hcomplex_of_complex)"
-apply (rule inj_onI, rule ccontr)
-apply (simp add: hcomplex_of_complex_def)
-done
+by (rule inj_onI, simp)
 
 lemma hcomplex_of_complex_i: "iii = hcomplex_of_complex ii"
-by (simp add: iii_def hcomplex_of_complex_def star_of_def star_n_def)
+by (simp add: iii_def star_of_def star_n_def)
 
-lemma hcomplex_of_complex_add [simp]:
+lemma hcomplex_of_complex_add:
      "hcomplex_of_complex (z1 + z2) = hcomplex_of_complex z1 + hcomplex_of_complex z2"
-by (simp add: hcomplex_of_complex_def hcomplex_add)
+by simp
 
-lemma hcomplex_of_complex_mult [simp]:
+lemma hcomplex_of_complex_mult:
      "hcomplex_of_complex (z1 * z2) = hcomplex_of_complex z1 * hcomplex_of_complex z2"
-by (simp add: hcomplex_of_complex_def hcomplex_mult)
+by simp
 
-lemma hcomplex_of_complex_eq_iff [simp]:
+lemma hcomplex_of_complex_eq_iff:
      "(hcomplex_of_complex z1 = hcomplex_of_complex z2) = (z1 = z2)"
-by (simp add: hcomplex_of_complex_def)
-
-
-lemma hcomplex_of_complex_minus [simp]:
-     "hcomplex_of_complex (-r) = - hcomplex_of_complex  r"
-by (simp add: hcomplex_of_complex_def hcomplex_minus)
-
-lemma hcomplex_of_complex_one [simp]: "hcomplex_of_complex 1 = 1"
-by (simp add: hcomplex_of_complex_def hypreal_one_def)
-
-lemma hcomplex_of_complex_zero [simp]: "hcomplex_of_complex 0 = 0"
-by (simp add: hcomplex_of_complex_def hypreal_zero_def)
+by simp
 
-lemma hcomplex_of_complex_zero_iff [simp]:
+lemma hcomplex_of_complex_minus:
+     "hcomplex_of_complex (-r) = - hcomplex_of_complex  r"
+by simp
+
+lemma hcomplex_of_complex_one: "hcomplex_of_complex 1 = 1"
+by simp
+
+lemma hcomplex_of_complex_zero: "hcomplex_of_complex 0 = 0"
+by simp
+
+lemma hcomplex_of_complex_zero_iff:
      "(hcomplex_of_complex r = 0) = (r = 0)"
-by (auto intro: FreeUltrafilterNat_P 
-         simp add: hcomplex_of_complex_def star_of_def star_n_def hypreal_zero_def)
+by simp
 
-lemma hcomplex_of_complex_inverse [simp]:
+lemma hcomplex_of_complex_inverse:
      "hcomplex_of_complex (inverse r) = inverse (hcomplex_of_complex r)"
-proof cases
-  assume "r=0" thus ?thesis by simp
-next
-  assume nz: "r\<noteq>0" 
-  show ?thesis
-  proof (rule hcomplex_mult_left_cancel [THEN iffD1]) 
-    show "hcomplex_of_complex r \<noteq> 0"
-      by (simp add: nz) 
-  next
-    have "hcomplex_of_complex r * hcomplex_of_complex (inverse r) =
-          hcomplex_of_complex (r * inverse r)"
-      by simp
-    also have "... = hcomplex_of_complex r * inverse (hcomplex_of_complex r)" 
-      by (simp add: nz)
-    finally show "hcomplex_of_complex r * hcomplex_of_complex (inverse r) =
-                  hcomplex_of_complex r * inverse (hcomplex_of_complex r)" .
-  qed
-qed
+by simp
 
-lemma hcomplex_of_complex_divide [simp]:
+lemma hcomplex_of_complex_divide:
      "hcomplex_of_complex (z1 / z2) = 
       hcomplex_of_complex z1 / hcomplex_of_complex z2"
-by (simp add: divide_inverse)
+by simp
 
 lemma hRe_hcomplex_of_complex:
    "hRe (hcomplex_of_complex z) = hypreal_of_real (Re z)"
-by (simp add: hcomplex_of_complex_def hypreal_of_real_def star_of_def star_n_def hRe)
+by (simp add: star_of_def hRe)
 
 lemma hIm_hcomplex_of_complex:
    "hIm (hcomplex_of_complex z) = hypreal_of_real (Im z)"
-by (simp add: hcomplex_of_complex_def hypreal_of_real_def star_of_def star_n_def hIm)
+by (simp add: star_of_def hIm)
 
 lemma hcmod_hcomplex_of_complex:
      "hcmod (hcomplex_of_complex x) = hypreal_of_real (cmod x)"
-by (simp add: hypreal_of_real_def star_of_def star_n_def hcomplex_of_complex_def hcmod)
+by (simp add: star_of_def hcmod)
 
 
 subsection{*Numerals and Arithmetic*}
 
-(*
-instance hcomplex :: number ..
-
-defs (overloaded)
-  hcomplex_number_of_def: "(number_of w :: hcomplex) == of_int (Rep_Bin w)"
-    --{*the type constraint is essential!*}
-
-instance hcomplex :: number_ring
-by (intro_classes, simp add: hcomplex_number_of_def) 
-*)
-
 lemma hcomplex_number_of_def: "(number_of w :: hcomplex) == of_int (Rep_Bin w)"
 apply (rule eq_reflection)
 apply (unfold star_number_def star_of_int_def)
@@ -1278,37 +879,27 @@
 apply (rule number_of_eq)
 done
 
-lemma hcomplex_of_complex_of_nat [simp]:
+lemma hcomplex_of_complex_of_nat:
      "hcomplex_of_complex (of_nat n) = of_nat n"
-by (simp add: hcomplex_of_complex_def)
+by (rule star_of_of_nat)
 
-lemma hcomplex_of_complex_of_int [simp]:
+lemma hcomplex_of_complex_of_int:
      "hcomplex_of_complex (of_int z) = of_int z"
-proof (cases z)
-  case (1 n)
-    thus ?thesis by simp
-next
-  case (2 n)
-    thus ?thesis 
-      by (simp only: of_int_minus hcomplex_of_complex_minus, simp)
-qed
+by (rule star_of_of_int)
 
-
-text{*Collapse applications of @{term hcomplex_of_complex} to @{term number_of}*}
-lemma hcomplex_number_of [simp]:
+lemma hcomplex_number_of:
      "hcomplex_of_complex (number_of w) = number_of w"
-by (simp add: hcomplex_number_of_def complex_number_of_def) 
+by (rule star_of_number_of)
 
 lemma hcomplex_of_hypreal_eq_hcomplex_of_complex: 
      "hcomplex_of_hypreal (hypreal_of_real x) =  
       hcomplex_of_complex (complex_of_real x)"
-by (simp add: hypreal_of_real_def hcomplex_of_hypreal hcomplex_of_complex_def 
-              star_of_def star_n_def
+by (simp add: hcomplex_of_hypreal star_of_def
               complex_of_real_def)
 
 lemma hcomplex_hypreal_number_of: 
   "hcomplex_of_complex (number_of w) = hcomplex_of_hypreal(number_of w)"
-by (simp only: complex_number_of [symmetric] hypreal_number_of [symmetric] 
+by (simp only: complex_number_of [symmetric] star_of_number_of [symmetric] 
                hcomplex_of_hypreal_eq_hcomplex_of_complex)
 
 text{*This theorem is necessary because theorems such as
@@ -1327,32 +918,24 @@
 Goal "z + hcnj z =  
       hcomplex_of_hypreal (2 * hRe(z))"
 by (res_inst_tac [("z","z")] eq_Abs_star 1);
-by (auto_tac (claset(),HOL_ss addsimps [hRe,hcnj,hcomplex_add,
+by (auto_tac (claset(),HOL_ss addsimps [hRe,hcnj,star_n_add,
     hypreal_mult,hcomplex_of_hypreal,complex_add_cnj]));
-qed "hcomplex_add_hcnj";
+qed "star_n_add_hcnj";
 
 Goal "z - hcnj z = \
 \     hcomplex_of_hypreal (hypreal_of_real #2 * hIm(z)) * iii";
 by (res_inst_tac [("z","z")] eq_Abs_star 1);
 by (auto_tac (claset(),simpset() addsimps [hIm,hcnj,hcomplex_diff,
     hypreal_of_real_def,hypreal_mult,hcomplex_of_hypreal,
-    complex_diff_cnj,iii_def,hcomplex_mult]));
+    complex_diff_cnj,iii_def,star_n_mult]));
 qed "hcomplex_diff_hcnj";
 *)
 
 
-lemma hcomplex_hcnj_num_zero_iff: "(hcnj z = 0) = (z = 0)"
+lemma hcomplex_hcnj_num_zero_iff [simp]: "(hcnj z = 0) = (z = 0)"
 apply (auto simp add: hcomplex_hcnj_zero_iff)
 done
-declare hcomplex_hcnj_num_zero_iff [simp]
 
-lemma hcomplex_zero_num: "0 = Abs_star (starrel `` {%n. 0})"
-apply (simp add: hypreal_zero_def)
-done
-
-lemma hcomplex_one_num: "1 =  Abs_star (starrel `` {%n. 1})"
-apply (simp add: hypreal_one_def)
-done
 
 (*** Real and imaginary stuff ***)
 
@@ -1461,14 +1044,6 @@
 
 ML
 {*
-(* val hcomplex_zero_def = thm"hcomplex_zero_def"; *)
-(* val hcomplex_one_def = thm"hcomplex_one_def"; *)
-(* val hcomplex_minus_def = thm"hcomplex_minus_def"; *)
-(* val hcomplex_diff_def = thm"hcomplex_diff_def"; *)
-(* val hcomplex_divide_def = thm"hcomplex_divide_def"; *)
-(* val hcomplex_mult_def = thm"hcomplex_mult_def"; *)
-(* val hcomplex_add_def = thm"hcomplex_add_def"; *)
-val hcomplex_of_complex_def = thm"hcomplex_of_complex_def";
 val iii_def = thm"iii_def";
 
 val hRe = thm"hRe";
@@ -1480,33 +1055,15 @@
 val hcomplex_hIm_one = thm"hcomplex_hIm_one";
 val inj_hcomplex_of_complex = thm"inj_hcomplex_of_complex";
 val hcomplex_of_complex_i = thm"hcomplex_of_complex_i";
-val hcomplex_add = thm"hcomplex_add";
-val hcomplex_add_commute = thm"hcomplex_add_commute";
-val hcomplex_add_assoc = thm"hcomplex_add_assoc";
-val hcomplex_add_zero_left = thm"hcomplex_add_zero_left";
-val hcomplex_add_zero_right = thm"hcomplex_add_zero_right";
+val star_n_add = thm"star_n_add";
 val hRe_add = thm"hRe_add";
 val hIm_add = thm"hIm_add";
-(* val hcomplex_minus_congruent = thm"hcomplex_minus_congruent"; *)
-val hcomplex_minus = thm"hcomplex_minus";
-val hcomplex_add_minus_left = thm"hcomplex_add_minus_left";
 val hRe_minus = thm"hRe_minus";
 val hIm_minus = thm"hIm_minus";
 val hcomplex_add_minus_eq_minus = thm"hcomplex_add_minus_eq_minus";
-val hcomplex_diff = thm"hcomplex_diff";
 val hcomplex_diff_eq_eq = thm"hcomplex_diff_eq_eq";
-val hcomplex_mult = thm"hcomplex_mult";
-val hcomplex_mult_commute = thm"hcomplex_mult_commute";
-val hcomplex_mult_assoc = thm"hcomplex_mult_assoc";
-val hcomplex_mult_one_left = thm"hcomplex_mult_one_left";
-val hcomplex_mult_one_right = thm"hcomplex_mult_one_right";
-val hcomplex_mult_zero_left = thm"hcomplex_mult_zero_left";
 val hcomplex_mult_minus_one = thm"hcomplex_mult_minus_one";
 val hcomplex_mult_minus_one_right = thm"hcomplex_mult_minus_one_right";
-val hcomplex_add_mult_distrib = thm"hcomplex_add_mult_distrib";
-val hcomplex_zero_not_eq_one = thm"hcomplex_zero_not_eq_one";
-val hcomplex_inverse = thm"hcomplex_inverse";
-val hcomplex_mult_inv_left = thm"hcomplex_mult_inv_left";
 val hcomplex_mult_left_cancel = thm"hcomplex_mult_left_cancel";
 val hcomplex_mult_right_cancel = thm"hcomplex_mult_right_cancel";
 val hcomplex_add_divide_distrib = thm"hcomplex_add_divide_distrib";
@@ -1579,7 +1136,7 @@
 val hcomplex_i_mult_eq = thm"hcomplex_i_mult_eq";
 val hcomplexpow_i_squared = thm"hcomplexpow_i_squared";
 val hcomplex_i_not_zero = thm"hcomplex_i_not_zero";
-val hcomplex_divide = thm"hcomplex_divide";
+val star_n_divide = thm"star_n_divide";
 val hsgn = thm"hsgn";
 val hsgn_zero = thm"hsgn_zero";
 val hsgn_one = thm"hsgn_one";
--- a/src/HOL/Hyperreal/HLog.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/HLog.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -12,7 +12,7 @@
 
 (* should be in NSA.ML *)
 lemma epsilon_ge_zero [simp]: "0 \<le> epsilon"
-by (simp add: epsilon_def star_n_def hypreal_zero_num hypreal_le)
+by (simp add: epsilon_def star_n_zero_num star_n_le)
 
 lemma hpfinite_witness: "epsilon : {x. 0 \<le> x & x : HFinite}"
 by auto
@@ -21,82 +21,52 @@
 constdefs
 
     powhr  :: "[hypreal,hypreal] => hypreal"     (infixr "powhr" 80)
-    "x powhr a == ( *f* exp) (a * ( *f* ln) x)"
+    "x powhr a == Ifun2_of (op powr) x a"
   
     hlog :: "[hypreal,hypreal] => hypreal"
-    "hlog a x == Abs_star(\<Union>A \<in> Rep_star(a).\<Union>X \<in> Rep_star(x).
-			     starrel `` {%n. log (A n) (X n)})"
+    "hlog a x == Ifun2_of log a x"
 
 
-lemma powhr: 
-    "(Abs_star(starrel `` {X})) powhr (Abs_star(starrel `` {Y})) =  
-     Abs_star(starrel `` {%n.  (X n) powr (Y n)})"
-by (simp add: powhr_def starfun hypreal_mult powr_def)
+lemma powhr: "(star_n X) powhr (star_n Y) = star_n (%n. (X n) powr (Y n))"
+by (simp add: powhr_def Ifun2_of_def star_of_def Ifun_star_n)
 
-lemma powhr_one_eq_one [simp]: "1 powhr a = 1"
-apply (rule_tac z=a in eq_Abs_star)
-apply (simp add: powhr hypreal_one_num)
-done
+lemma powhr_one_eq_one [simp]: "!!a. 1 powhr a = 1"
+by (unfold powhr_def, transfer, simp)
 
 lemma powhr_mult:
-     "[| 0 < x; 0 < y |] ==> (x * y) powhr a = (x powhr a) * (y powhr a)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star, rule_tac z=a in eq_Abs_star)
-apply (simp add: powhr hypreal_zero_num hypreal_mult hypreal_less, ultra)
-apply (simp add: powr_mult) 
-done
+  "!!a x y. [| 0 < x; 0 < y |] ==> (x * y) powhr a = (x powhr a) * (y powhr a)"
+by (unfold powhr_def, transfer, rule powr_mult)
 
-lemma powhr_gt_zero [simp]: "0 < x powhr a"
-apply (rule_tac z=a in eq_Abs_star, rule_tac z=x in eq_Abs_star)
-apply (simp add: hypreal_zero_def powhr hypreal_less hypreal_zero_num)
-done
+lemma powhr_gt_zero [simp]: "!!a x. 0 < x powhr a"
+by (unfold powhr_def, transfer, simp)
 
 lemma powhr_not_zero [simp]: "x powhr a \<noteq> 0"
 by (rule powhr_gt_zero [THEN hypreal_not_refl2, THEN not_sym])
 
-lemma hypreal_divide: 
-     "(Abs_star(starrel `` {X}))/(Abs_star(starrel `` {Y})) =  
-      (Abs_star(starrel `` {%n. (X n::real)/(Y n)}))"
-by (simp add: divide_inverse hypreal_zero_num hypreal_inverse hypreal_mult) 
+lemma powhr_divide:
+  "!!a x y. [| 0 < x; 0 < y |] ==> (x / y) powhr a = (x powhr a)/(y powhr a)"
+by (unfold powhr_def, transfer, rule powr_divide)
 
-lemma powhr_divide:
-     "[| 0 < x; 0 < y |] ==> (x / y) powhr a = (x powhr a)/(y powhr a)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star, rule_tac z=a in eq_Abs_star)
-apply (simp add: powhr hypreal_divide hypreal_zero_num hypreal_less, ultra)
-apply (simp add: powr_divide)
-done
+lemma powhr_add: "!!a b x. x powhr (a + b) = (x powhr a) * (x powhr b)"
+by (unfold powhr_def, transfer, rule powr_add)
 
-lemma powhr_add: "x powhr (a + b) = (x powhr a) * (x powhr b)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=b in eq_Abs_star, rule_tac z=a in eq_Abs_star)
-apply (simp add: powhr hypreal_add hypreal_mult powr_add)
-done
-
-lemma powhr_powhr: "(x powhr a) powhr b = x powhr (a * b)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=b in eq_Abs_star, rule_tac z=a in eq_Abs_star)
-apply (simp add: powhr hypreal_mult powr_powr)
-done
+lemma powhr_powhr: "!!a b x. (x powhr a) powhr b = x powhr (a * b)"
+by (unfold powhr_def, transfer, rule powr_powr)
 
-lemma powhr_powhr_swap: "(x powhr a) powhr b = (x powhr b) powhr a"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=b in eq_Abs_star, rule_tac z=a in eq_Abs_star)
-apply (simp add: powhr powr_powr_swap)
-done
+lemma powhr_powhr_swap: "!!a b x. (x powhr a) powhr b = (x powhr b) powhr a"
+by (unfold powhr_def, transfer, rule powr_powr_swap)
 
-lemma powhr_minus: "x powhr (-a) = inverse (x powhr a)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=a in eq_Abs_star)
-apply (simp add: powhr hypreal_minus hypreal_inverse hypreal_less powr_minus)
-done
+lemma powhr_minus: "!!a x. x powhr (-a) = inverse (x powhr a)"
+by (unfold powhr_def, transfer, rule powr_minus)
 
 lemma powhr_minus_divide: "x powhr (-a) = 1/(x powhr a)"
 by (simp add: divide_inverse powhr_minus)
 
-lemma powhr_less_mono: "[| a < b; 1 < x |] ==> x powhr a < x powhr b"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=a in eq_Abs_star, rule_tac z=b in eq_Abs_star)
-apply (simp add: powhr hypreal_one_num hypreal_less, ultra)
-done
+lemma powhr_less_mono: "!!a b x. [| a < b; 1 < x |] ==> x powhr a < x powhr b"
+by (unfold powhr_def, transfer, simp)
 
-lemma powhr_less_cancel: "[| x powhr a < x powhr b; 1 < x |] ==> a < b"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=a in eq_Abs_star, rule_tac z=b in eq_Abs_star)
-apply (simp add: powhr hypreal_one_num hypreal_less, ultra)
-done
+lemma powhr_less_cancel: "!!a b x. [| x powhr a < x powhr b; 1 < x |] ==> a < b"
+by (unfold powhr_def, transfer, simp)
 
 lemma powhr_less_cancel_iff [simp]:
      "1 < x ==> (x powhr a < x powhr b) = (a < b)"
@@ -106,57 +76,38 @@
      "1 < x ==> (x powhr a \<le> x powhr b) = (a \<le> b)"
 by (simp add: linorder_not_less [symmetric])
 
-lemma hlog: 
-     "hlog (Abs_star(starrel `` {X})) (Abs_star(starrel `` {Y})) =  
-      Abs_star(starrel `` {%n. log (X n) (Y n)})"
-apply (simp add: hlog_def)
-apply (rule arg_cong [where f=Abs_star], auto, ultra)
-done
+lemma hlog:
+     "hlog (star_n X) (star_n Y) =  
+      star_n (%n. log (X n) (Y n))"
+by (simp add: hlog_def Ifun2_of_def star_of_def Ifun_star_n)
 
-lemma hlog_starfun_ln: "( *f* ln) x = hlog (( *f* exp) 1) x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfun hlog log_ln hypreal_one_num)
-done
+lemma hlog_starfun_ln: "!!x. ( *f* ln) x = hlog (( *f* exp) 1) x"
+by (unfold hlog_def, transfer, rule log_ln)
 
 lemma powhr_hlog_cancel [simp]:
-     "[| 0 < a; a \<noteq> 1; 0 < x |] ==> a powhr (hlog a x) = x"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=a in eq_Abs_star)
-apply (simp add: hlog powhr hypreal_zero_num hypreal_less hypreal_one_num, ultra)
-done
+     "!!a x. [| 0 < a; a \<noteq> 1; 0 < x |] ==> a powhr (hlog a x) = x"
+by (unfold powhr_def hlog_def, transfer, simp)
 
 lemma hlog_powhr_cancel [simp]:
-     "[| 0 < a; a \<noteq> 1 |] ==> hlog a (a powhr y) = y"
-apply (rule_tac z=y in eq_Abs_star, rule_tac z=a in eq_Abs_star)
-apply (simp add: hlog powhr hypreal_zero_num hypreal_less hypreal_one_num, ultra)
-apply (auto intro: log_powr_cancel) 
-done
+     "!!a y. [| 0 < a; a \<noteq> 1 |] ==> hlog a (a powhr y) = y"
+by (unfold powhr_def hlog_def, transfer, simp)
 
 lemma hlog_mult:
-     "[| 0 < a; a \<noteq> 1; 0 < x; 0 < y  |]  
+     "!!a x y. [| 0 < a; a \<noteq> 1; 0 < x; 0 < y  |]  
       ==> hlog a (x * y) = hlog a x + hlog a y"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star, rule_tac z=a in eq_Abs_star)
-apply (simp add: hlog powhr hypreal_zero_num hypreal_one_num hypreal_less hypreal_add hypreal_mult, ultra)
-apply (simp add: log_mult)
-done
+by (unfold powhr_def hlog_def, transfer, rule log_mult)
 
 lemma hlog_as_starfun:
-     "[| 0 < a; a \<noteq> 1 |] ==> hlog a x = ( *f* ln) x / ( *f* ln) a"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=a in eq_Abs_star)
-apply (simp add: hlog starfun hypreal_zero_num hypreal_one_num hypreal_divide log_def)
-done
+     "!!a x. [| 0 < a; a \<noteq> 1 |] ==> hlog a x = ( *f* ln) x / ( *f* ln) a"
+by (unfold hlog_def, transfer, simp add: log_def)
 
 lemma hlog_eq_div_starfun_ln_mult_hlog:
-     "[| 0 < a; a \<noteq> 1; 0 < b; b \<noteq> 1; 0 < x |]  
+     "!!a b x. [| 0 < a; a \<noteq> 1; 0 < b; b \<noteq> 1; 0 < x |]  
       ==> hlog a x = (( *f* ln) b/( *f*ln) a) * hlog b x"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=a in eq_Abs_star, rule_tac z=b in eq_Abs_star)
-apply (simp add: hlog starfun hypreal_zero_num hypreal_one_num hypreal_less hypreal_divide hypreal_mult, ultra)
-apply (auto dest: log_eq_div_ln_mult_log) 
-done
+by (unfold hlog_def, transfer, rule log_eq_div_ln_mult_log)
 
-lemma powhr_as_starfun: "x powhr a = ( *f* exp) (a * ( *f* ln) x)"
-apply (rule_tac z=a in eq_Abs_star, rule_tac z=x in eq_Abs_star)
-apply (simp add: powhr starfun hypreal_mult powr_def)
-done
+lemma powhr_as_starfun: "!!a x. x powhr a = ( *f* exp) (a * ( *f* ln) x)"
+by (unfold powhr_def, transfer, simp add: powr_def)
 
 lemma HInfinite_powhr:
      "[| x : HInfinite; 0 < x; a : HFinite - Infinitesimal;  
@@ -179,16 +130,11 @@
      "[| a : HInfinite; 0 < a |] ==> hlog a x = ( *f* ln) x / ( *f* ln) a"
 by (rule hlog_as_starfun, auto)
 
-lemma hlog_one [simp]: "hlog a 1 = 0"
-apply (rule_tac z=a in eq_Abs_star)
-apply (simp add: hypreal_one_num hypreal_zero_num hlog)
-done
+lemma hlog_one [simp]: "!!a. hlog a 1 = 0"
+by (unfold hlog_def, transfer, simp)
 
-lemma hlog_eq_one [simp]: "[| 0 < a; a \<noteq> 1 |] ==> hlog a a = 1"
-apply (rule_tac z=a in eq_Abs_star)
-apply (simp add: hypreal_one_num hypreal_zero_num hlog hypreal_less, ultra)
-apply (auto intro: log_eq_one) 
-done
+lemma hlog_eq_one [simp]: "!!a. [| 0 < a; a \<noteq> 1 |] ==> hlog a a = 1"
+by (unfold hlog_def, transfer, rule log_eq_one)
 
 lemma hlog_inverse:
      "[| 0 < a; a \<noteq> 1; 0 < x |] ==> hlog a (inverse x) = - hlog a x"
@@ -201,10 +147,8 @@
 by (simp add: hlog_mult hlog_inverse divide_inverse)
 
 lemma hlog_less_cancel_iff [simp]:
-     "[| 1 < a; 0 < x; 0 < y |] ==> (hlog a x < hlog a y) = (x < y)"
-apply (rule_tac z=a in eq_Abs_star, rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (auto simp add: hlog hypreal_less hypreal_zero_num hypreal_one_num, ultra+)
-done
+     "!!a x y. [| 1 < a; 0 < x; 0 < y |] ==> (hlog a x < hlog a y) = (x < y)"
+by (unfold hlog_def, transfer, simp)
 
 lemma hlog_le_cancel_iff [simp]:
      "[| 1 < a; 0 < x; 0 < y |] ==> (hlog a x \<le> hlog a y) = (x \<le> y)"
@@ -218,7 +162,6 @@
 val powhr_mult = thm "powhr_mult";
 val powhr_gt_zero = thm "powhr_gt_zero";
 val powhr_not_zero = thm "powhr_not_zero";
-val hypreal_divide = thm "hypreal_divide";
 val powhr_divide = thm "powhr_divide";
 val powhr_add = thm "powhr_add";
 val powhr_powhr = thm "powhr_powhr";
--- a/src/HOL/Hyperreal/HSeries.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/HSeries.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -13,9 +13,13 @@
 
 constdefs 
   sumhr :: "(hypnat * hypnat * (nat=>real)) => hypreal"
+   "sumhr == 
+      %(M,N,f). Ifun2_of (%m n. setsum f {m..<n}) M N"
+(*
    "sumhr p == 
       (%(M,N,f). Abs_star(\<Union>X \<in> Rep_star(M). \<Union>Y \<in> Rep_star(N).  
                              starrel ``{%n::nat. setsum f {X n..<Y n}})) p"
+*)
 
   NSsums  :: "[nat=>real,real] => bool"     (infixr "NSsums" 80)
    "f NSsums s  == (%n. setsum f {0..<n}) ----NS> s"
@@ -27,72 +31,68 @@
    "NSsuminf f == (@s. f NSsums s)"
 
 
-lemma sumhr: 
-     "sumhr(Abs_star(starrel``{%n. M n}),  
-            Abs_star(starrel``{%n. N n}), f) =  
-      Abs_star(starrel `` {%n. setsum f {M n..<N n}})"
-apply (simp add: sumhr_def)
-apply (rule arg_cong [where f = Abs_star])
-apply (auto, ultra)
-done
+lemma sumhr:
+     "sumhr(star_n M, star_n N, f) =  
+      star_n (%n. setsum f {M n..<N n})"
+by (simp add: sumhr_def Ifun2_of_def star_of_def Ifun_star_n)
 
 text{*Base case in definition of @{term sumr}*}
 lemma sumhr_zero [simp]: "sumhr (m,0,f) = 0"
-apply (rule_tac z=m in eq_Abs_star)
-apply (simp add: hypnat_zero_def sumhr symmetric hypreal_zero_def)
+apply (cases m)
+apply (simp add: star_n_zero_num sumhr symmetric)
 done
 
 text{*Recursive case in definition of @{term sumr}*}
 lemma sumhr_if: 
      "sumhr(m,n+1,f) = 
-      (if n + 1 \<le> m then 0 else sumhr(m,n,f) + ( *fNat* f) n)"
-apply (rule_tac z=m in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (auto simp add: hypnat_one_def sumhr hypnat_add hypnat_le starfunNat
-           hypreal_add hypreal_zero_def,   ultra+)
+      (if n + 1 \<le> m then 0 else sumhr(m,n,f) + ( *f* f) n)"
+apply (cases m, cases n)
+apply (auto simp add: star_n_one_num sumhr star_n_add star_n_le starfun
+           star_n_zero_num star_n_eq_iff, ultra+)
 done
 
 lemma sumhr_Suc_zero [simp]: "sumhr (n + 1, n, f) = 0"
-apply (rule_tac z=n in eq_Abs_star)
-apply (simp add: hypnat_one_def sumhr hypnat_add hypreal_zero_def)
+apply (cases n)
+apply (simp add: star_n_one_num sumhr star_n_add star_n_zero_num)
 done
 
 lemma sumhr_eq_bounds [simp]: "sumhr (n,n,f) = 0"
-apply (rule_tac z=n in eq_Abs_star)
-apply (simp add: sumhr hypreal_zero_def)
+apply (cases n)
+apply (simp add: sumhr star_n_zero_num)
 done
 
-lemma sumhr_Suc [simp]: "sumhr (m,m + 1,f) = ( *fNat* f) m"
-apply (rule_tac z=m in eq_Abs_star)
-apply (simp add: sumhr hypnat_one_def  hypnat_add starfunNat)
+lemma sumhr_Suc [simp]: "sumhr (m,m + 1,f) = ( *f* f) m"
+apply (cases m)
+apply (simp add: sumhr star_n_one_num star_n_add starfun)
 done
 
 lemma sumhr_add_lbound_zero [simp]: "sumhr(m+k,k,f) = 0"
-apply (rule_tac z=m in eq_Abs_star, rule_tac z=k in eq_Abs_star)
-apply (simp add: sumhr hypnat_add hypreal_zero_def)
+apply (cases m, cases k)
+apply (simp add: sumhr star_n_add star_n_zero_num)
 done
 
 lemma sumhr_add: "sumhr (m,n,f) + sumhr(m,n,g) = sumhr(m,n,%i. f i + g i)"
-apply (rule_tac z=m in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (simp add: sumhr hypreal_add setsum_addf)
+apply (cases m, cases n)
+apply (simp add: sumhr star_n_add setsum_addf)
 done
 
 lemma sumhr_mult: "hypreal_of_real r * sumhr(m,n,f) = sumhr(m,n,%n. r * f n)"
-apply (rule_tac z=m in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (simp add: sumhr hypreal_of_real_def star_of_def star_n_def hypreal_mult setsum_mult)
+apply (cases m, cases n)
+apply (simp add: sumhr star_of_def star_n_mult setsum_mult)
 done
 
 lemma sumhr_split_add: "n < p ==> sumhr(0,n,f) + sumhr(n,p,f) = sumhr(0,p,f)"
-apply (rule_tac z=n in eq_Abs_star, rule_tac z=p in eq_Abs_star)
+apply (cases n, cases p)
 apply (auto elim!: FreeUltrafilterNat_subset simp 
-            add: hypnat_zero_def sumhr hypreal_add hypnat_less setsum_add_nat_ivl)
+            add: star_n_zero_num sumhr star_n_add star_n_less setsum_add_nat_ivl star_n_eq_iff)
 done
 
 lemma sumhr_split_diff: "n<p ==> sumhr(0,p,f) - sumhr(0,n,f) = sumhr(n,p,f)"
 by (drule_tac f1 = f in sumhr_split_add [symmetric], simp)
 
 lemma sumhr_hrabs: "abs(sumhr(m,n,f)) \<le> sumhr(m,n,%i. abs(f i))"
-apply (rule_tac z=n in eq_Abs_star, rule_tac z=m in eq_Abs_star)
-apply (simp add: sumhr hypreal_le hypreal_hrabs setsum_abs)
+apply (cases n, cases m)
+apply (simp add: sumhr star_n_le star_n_abs setsum_abs)
 done
 
 text{* other general version also needed *}
@@ -105,27 +105,26 @@
 
 lemma sumhr_const:
      "sumhr(0, n, %i. r) = hypreal_of_hypnat n * hypreal_of_real r"
-apply (rule_tac z=n in eq_Abs_star)
-apply (simp add: sumhr hypnat_zero_def hypreal_of_real_def hypreal_of_hypnat 
-                 star_of_def star_n_def
-                 hypreal_mult real_of_nat_def)
+apply (cases n)
+apply (simp add: sumhr star_n_zero_num hypreal_of_hypnat 
+                 star_of_def star_n_mult real_of_nat_def)
 done
 
 lemma sumhr_less_bounds_zero [simp]: "n < m ==> sumhr(m,n,f) = 0"
-apply (rule_tac z=m in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (auto elim: FreeUltrafilterNat_subset 
-            simp add: sumhr hypnat_less hypreal_zero_def)
+apply (cases m, cases n)
+apply (auto elim: FreeUltrafilterNat_subset
+            simp add: sumhr star_n_less star_n_zero_num star_n_eq_iff)
 done
 
 lemma sumhr_minus: "sumhr(m, n, %i. - f i) = - sumhr(m, n, f)"
-apply (rule_tac z=m in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (simp add: sumhr hypreal_minus setsum_negf)
+apply (cases m, cases n)
+apply (simp add: sumhr star_n_minus setsum_negf)
 done
 
 lemma sumhr_shift_bounds:
      "sumhr(m+hypnat_of_nat k,n+hypnat_of_nat k,f) = sumhr(m,n,%i. f(i + k))"
-apply (rule_tac z=m in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (simp add: sumhr hypnat_add setsum_shift_bounds_nat_ivl hypnat_of_nat_eq)
+apply (cases m, cases n)
+apply (simp add: sumhr star_n_add setsum_shift_bounds_nat_ivl hypnat_of_nat_eq)
 done
 
 
@@ -135,29 +134,29 @@
  (such as @{term whn})*}
 lemma sumhr_hypreal_of_hypnat_omega: 
       "sumhr(0,whn,%i. 1) = hypreal_of_hypnat whn"
-by (simp add: hypnat_omega_def hypnat_zero_def sumhr hypreal_of_hypnat
+by (simp add: hypnat_omega_def star_n_zero_num sumhr hypreal_of_hypnat
               real_of_nat_def)
 
 lemma sumhr_hypreal_omega_minus_one: "sumhr(0, whn, %i. 1) = omega - 1"
-by (simp add: hypnat_omega_def hypnat_zero_def omega_def hypreal_one_def
-              sumhr hypreal_diff real_of_nat_def star_n_def)
+by (simp add: hypnat_omega_def star_n_zero_num omega_def star_n_one_num
+              sumhr star_n_diff real_of_nat_def)
 
 lemma sumhr_minus_one_realpow_zero [simp]: 
      "sumhr(0, whn + whn, %i. (-1) ^ (i+1)) = 0"
 by (simp del: realpow_Suc 
-         add: sumhr hypnat_add nat_mult_2 [symmetric] hypreal_zero_def 
-              hypnat_zero_def hypnat_omega_def)
+         add: sumhr star_n_add nat_mult_2 [symmetric] star_n_zero_num 
+              star_n_zero_num hypnat_omega_def)
 
 lemma sumhr_interval_const:
      "(\<forall>n. m \<le> Suc n --> f n = r) & m \<le> na  
       ==> sumhr(hypnat_of_nat m,hypnat_of_nat na,f) =  
           (hypreal_of_nat (na - m) * hypreal_of_real r)"
-by(simp add: sumhr hypreal_of_nat_eq hypnat_of_nat_eq hypreal_of_real_def
-             real_of_nat_def star_of_def star_n_def hypreal_mult cong: setsum_ivl_cong)
+by(simp add: sumhr hypreal_of_nat_eq hypnat_of_nat_eq
+             real_of_nat_def star_of_def star_n_mult cong: setsum_ivl_cong)
 
-lemma starfunNat_sumr: "( *fNat* (%n. setsum f {0..<n})) N = sumhr(0,N,f)"
-apply (rule_tac z=N in eq_Abs_star)
-apply (simp add: hypnat_zero_def starfunNat sumhr)
+lemma starfunNat_sumr: "( *f* (%n. setsum f {0..<n})) N = sumhr(0,N,f)"
+apply (cases N)
+apply (simp add: star_n_zero_num starfun sumhr)
 done
 
 lemma sumhr_hrabs_approx [simp]: "sumhr(0, M, f) @= sumhr(0, N, f)  
--- a/src/HOL/Hyperreal/HTranscendental.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/HTranscendental.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -48,23 +48,19 @@
 subsection{*Nonstandard Extension of Square Root Function*}
 
 lemma STAR_sqrt_zero [simp]: "( *f* sqrt) 0 = 0"
-by (simp add: starfun hypreal_zero_num)
+by (simp add: starfun star_n_zero_num)
 
 lemma STAR_sqrt_one [simp]: "( *f* sqrt) 1 = 1"
-by (simp add: starfun hypreal_one_num)
+by (simp add: starfun star_n_one_num)
 
 lemma hypreal_sqrt_pow2_iff: "(( *f* sqrt)(x) ^ 2 = x) = (0 \<le> x)"
-apply (rule_tac z=x in eq_Abs_star)
-apply (auto simp add: hypreal_le hypreal_zero_num starfun hrealpow 
+apply (cases x)
+apply (auto simp add: star_n_le star_n_zero_num starfun hrealpow star_n_eq_iff
             simp del: hpowr_Suc realpow_Suc)
 done
 
-lemma hypreal_sqrt_gt_zero_pow2: "0 < x ==> ( *f* sqrt) (x) ^ 2 = x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (auto intro: FreeUltrafilterNat_subset 
-            simp add: hypreal_less starfun hrealpow hypreal_zero_num 
-            simp del: hpowr_Suc realpow_Suc)
-done
+lemma hypreal_sqrt_gt_zero_pow2: "!!x. 0 < x ==> ( *f* sqrt) (x) ^ 2 = x"
+by (transfer, simp)
 
 lemma hypreal_sqrt_pow2_gt_zero: "0 < x ==> 0 < ( *f* sqrt) (x) ^ 2"
 by (frule hypreal_sqrt_gt_zero_pow2, auto)
@@ -81,9 +77,9 @@
 done
 
 lemma hypreal_sqrt_mult_distrib: 
-    "[|0 < x; 0 <y |] ==> ( *f* sqrt)(x*y) =  ( *f* sqrt)(x) * ( *f* sqrt)(y)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: hypreal_zero_def starfun hypreal_mult hypreal_less hypreal_zero_num, ultra)
+    "!!x y. [|0 < x; 0 <y |] ==>
+      ( *f* sqrt)(x*y) = ( *f* sqrt)(x) * ( *f* sqrt)(y)"
+apply transfer
 apply (auto intro: real_sqrt_mult_distrib) 
 done
 
@@ -108,42 +104,34 @@
 lemma hypreal_sqrt_sum_squares [simp]:
      "(( *f* sqrt)(x*x + y*y + z*z) @= 0) = (x*x + y*y + z*z @= 0)"
 apply (rule hypreal_sqrt_approx_zero2)
-apply (rule hypreal_le_add_order)+
+apply (rule add_nonneg_nonneg)+
 apply (auto simp add: zero_le_square)
 done
 
 lemma hypreal_sqrt_sum_squares2 [simp]:
      "(( *f* sqrt)(x*x + y*y) @= 0) = (x*x + y*y @= 0)"
 apply (rule hypreal_sqrt_approx_zero2)
-apply (rule hypreal_le_add_order)
+apply (rule add_nonneg_nonneg)
 apply (auto simp add: zero_le_square)
 done
 
-lemma hypreal_sqrt_gt_zero: "0 < x ==> 0 < ( *f* sqrt)(x)"
-apply (rule_tac z=x in eq_Abs_star)
-apply (auto simp add: starfun hypreal_zero_def hypreal_less hypreal_zero_num, ultra)
+lemma hypreal_sqrt_gt_zero: "!!x. 0 < x ==> 0 < ( *f* sqrt)(x)"
+apply transfer
 apply (auto intro: real_sqrt_gt_zero)
 done
 
 lemma hypreal_sqrt_ge_zero: "0 \<le> x ==> 0 \<le> ( *f* sqrt)(x)"
 by (auto intro: hypreal_sqrt_gt_zero simp add: order_le_less)
 
-lemma hypreal_sqrt_hrabs [simp]: "( *f* sqrt)(x ^ 2) = abs(x)"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfun hypreal_hrabs hypreal_mult numeral_2_eq_2)
-done
+lemma hypreal_sqrt_hrabs [simp]: "!!x. ( *f* sqrt)(x ^ 2) = abs(x)"
+by (transfer, simp)
 
-lemma hypreal_sqrt_hrabs2 [simp]: "( *f* sqrt)(x*x) = abs(x)"
-apply (rule hrealpow_two [THEN subst])
-apply (rule numeral_2_eq_2 [THEN subst])
-apply (rule hypreal_sqrt_hrabs)
-done
+lemma hypreal_sqrt_hrabs2 [simp]: "!!x. ( *f* sqrt)(x*x) = abs(x)"
+by (transfer, simp)
 
 lemma hypreal_sqrt_hyperpow_hrabs [simp]:
-     "( *f* sqrt)(x pow (hypnat_of_nat 2)) = abs(x)"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfun hypreal_hrabs hypnat_of_nat_eq hyperpow)
-done
+     "!!x. ( *f* sqrt)(x pow (hypnat_of_nat 2)) = abs(x)"
+by (unfold hyperpow_def, transfer, simp)
 
 lemma star_sqrt_HFinite: "\<lbrakk>x \<in> HFinite; 0 \<le> x\<rbrakk> \<Longrightarrow> ( *f* sqrt) x \<in> HFinite"
 apply (rule HFinite_square_iff [THEN iffD1])
@@ -160,11 +148,8 @@
 apply (auto simp add: st_hrabs st_zero_le star_sqrt_HFinite)
 done
 
-lemma hypreal_sqrt_sum_squares_ge1 [simp]: "x \<le> ( *f* sqrt)(x ^ 2 + y ^ 2)"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: starfun hypreal_add hrealpow hypreal_le 
-            del: hpowr_Suc realpow_Suc)
-done
+lemma hypreal_sqrt_sum_squares_ge1 [simp]: "!!x y. x \<le> ( *f* sqrt)(x ^ 2 + y ^ 2)"
+by (transfer, simp)
 
 lemma HFinite_hypreal_sqrt:
      "[| 0 \<le> x; x \<in> HFinite |] ==> ( *f* sqrt) x \<in> HFinite"
@@ -189,7 +174,7 @@
 lemma HFinite_sqrt_sum_squares [simp]:
      "(( *f* sqrt)(x*x + y*y) \<in> HFinite) = (x*x + y*y \<in> HFinite)"
 apply (rule HFinite_hypreal_sqrt_iff)
-apply (rule hypreal_le_add_order)
+apply (rule add_nonneg_nonneg)
 apply (auto simp add: zero_le_square)
 done
 
@@ -216,7 +201,7 @@
 lemma Infinitesimal_sqrt_sum_squares [simp]:
      "(( *f* sqrt)(x*x + y*y) \<in> Infinitesimal) = (x*x + y*y \<in> Infinitesimal)"
 apply (rule Infinitesimal_hypreal_sqrt_iff)
-apply (rule hypreal_le_add_order)
+apply (rule add_nonneg_nonneg)
 apply (auto simp add: zero_le_square)
 done
 
@@ -243,35 +228,35 @@
 lemma HInfinite_sqrt_sum_squares [simp]:
      "(( *f* sqrt)(x*x + y*y) \<in> HInfinite) = (x*x + y*y \<in> HInfinite)"
 apply (rule HInfinite_hypreal_sqrt_iff)
-apply (rule hypreal_le_add_order)
+apply (rule add_nonneg_nonneg)
 apply (auto simp add: zero_le_square)
 done
 
 lemma HFinite_exp [simp]:
      "sumhr (0, whn, %n. inverse (real (fact n)) * x ^ n) \<in> HFinite"
 by (auto intro!: NSBseq_HFinite_hypreal NSconvergent_NSBseq 
-         simp add: starfunNat_sumr [symmetric] starfunNat hypnat_omega_def
+         simp add: starfunNat_sumr [symmetric] starfun hypnat_omega_def
                    convergent_NSconvergent_iff [symmetric] 
                    summable_convergent_sumr_iff [symmetric] summable_exp)
 
 lemma exphr_zero [simp]: "exphr 0 = 1"
 apply (simp add: exphr_def sumhr_split_add
                    [OF hypnat_one_less_hypnat_omega, symmetric]) 
-apply (simp add: sumhr hypnat_zero_def starfunNat hypnat_one_def hypnat_add
-                 hypnat_omega_def hypreal_add 
-            del: hypnat_add_zero_left)
-apply (simp add: hypreal_one_num [symmetric])
+apply (simp add: sumhr star_n_zero_num starfun star_n_one_num star_n_add
+                 hypnat_omega_def
+            del: OrderedGroup.add_0)
+apply (simp add: star_n_one_num [symmetric])
 done
 
 lemma coshr_zero [simp]: "coshr 0 = 1"
 apply (simp add: coshr_def sumhr_split_add
                    [OF hypnat_one_less_hypnat_omega, symmetric]) 
-apply (simp add: sumhr hypnat_zero_def hypnat_one_def hypnat_omega_def)
-apply (simp add: hypreal_one_def [symmetric] hypreal_zero_def [symmetric])
+apply (simp add: sumhr star_n_zero_num star_n_one_num hypnat_omega_def)
+apply (simp add: star_n_one_num [symmetric] star_n_zero_num [symmetric])
 done
 
 lemma STAR_exp_zero_approx_one [simp]: "( *f* exp) 0 @= 1"
-by (simp add: hypreal_zero_def hypreal_one_def starfun hypreal_one_num)
+by (simp add: star_n_zero_num star_n_one_num starfun)
 
 lemma STAR_exp_Infinitesimal: "x \<in> Infinitesimal ==> ( *f* exp) x @= 1"
 apply (case_tac "x = 0")
@@ -289,10 +274,8 @@
 lemma STAR_exp_epsilon [simp]: "( *f* exp) epsilon @= 1"
 by (auto intro: STAR_exp_Infinitesimal)
 
-lemma STAR_exp_add: "( *f* exp)(x + y) = ( *f* exp) x * ( *f* exp) y"
-apply (rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (simp add: starfun hypreal_add hypreal_mult exp_add)
-done
+lemma STAR_exp_add: "!!x y. ( *f* exp)(x + y) = ( *f* exp) x * ( *f* exp) y"
+by (transfer, rule exp_add)
 
 lemma exphr_hypreal_of_real_exp_eq: "exphr x = hypreal_of_real (exp x)"
 apply (simp add: exphr_def)
@@ -300,7 +283,7 @@
 apply (rule approx_st_eq, auto)
 apply (rule approx_minus_iff [THEN iffD2])
 apply (simp only: mem_infmal_iff [symmetric])
-apply (auto simp add: mem_infmal_iff [symmetric] hypreal_of_real_def star_of_def star_n_def hypnat_zero_def hypnat_omega_def sumhr hypreal_minus hypreal_add)
+apply (auto simp add: mem_infmal_iff [symmetric] star_of_def star_n_zero_num hypnat_omega_def sumhr star_n_minus star_n_add)
 apply (rule NSLIMSEQ_zero_Infinitesimal_hypreal)
 apply (insert exp_converges [of x]) 
 apply (simp add: sums_def) 
@@ -308,11 +291,8 @@
 apply (simp add: LIMSEQ_NSLIMSEQ_iff)
 done
 
-lemma starfun_exp_ge_add_one_self [simp]: "0 \<le> x ==> (1 + x) \<le> ( *f* exp) x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfun hypreal_add hypreal_le hypreal_zero_num hypreal_one_num, ultra)
-apply (erule exp_ge_add_one_self_aux)
-done
+lemma starfun_exp_ge_add_one_self [simp]: "!!x. 0 \<le> x ==> (1 + x) \<le> ( *f* exp) x"
+by (transfer, rule exp_ge_add_one_self_aux)
 
 (* exp (oo) is infinite *)
 lemma starfun_exp_HInfinite:
@@ -322,10 +302,8 @@
 apply (rule order_trans [of _ "1+x"], auto) 
 done
 
-lemma starfun_exp_minus: "( *f* exp) (-x) = inverse(( *f* exp) x)"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfun hypreal_inverse hypreal_minus exp_minus)
-done
+lemma starfun_exp_minus: "!!x. ( *f* exp) (-x) = inverse(( *f* exp) x)"
+by (transfer, rule exp_minus)
 
 (* exp (-oo) is infinitesimal *)
 lemma starfun_exp_Infinitesimal:
@@ -336,10 +314,8 @@
             simp add: starfun_exp_minus HInfinite_minus_iff)
 done
 
-lemma starfun_exp_gt_one [simp]: "0 < x ==> 1 < ( *f* exp) x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfun hypreal_one_num hypreal_zero_num hypreal_less, ultra)
-done
+lemma starfun_exp_gt_one [simp]: "!!x. 0 < x ==> 1 < ( *f* exp) x"
+by (transfer, simp)
 
 (* needs derivative of inverse function
    TRY a NS one today!!!
@@ -354,38 +330,26 @@
 *)
 
 
-lemma starfun_ln_exp [simp]: "( *f* ln) (( *f* exp) x) = x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfun)
-done
+lemma starfun_ln_exp [simp]: "!!x. ( *f* ln) (( *f* exp) x) = x"
+by (transfer, simp)
 
-lemma starfun_exp_ln_iff [simp]: "(( *f* exp)(( *f* ln) x) = x) = (0 < x)"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfun hypreal_zero_num hypreal_less)
-done
+lemma starfun_exp_ln_iff [simp]: "!!x. (( *f* exp)(( *f* ln) x) = x) = (0 < x)"
+by (transfer, simp)
 
 lemma starfun_exp_ln_eq: "( *f* exp) u = x ==> ( *f* ln) x = u"
-by (auto simp add: starfun)
+by auto
 
-lemma starfun_ln_less_self [simp]: "0 < x ==> ( *f* ln) x < x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfun hypreal_zero_num hypreal_less, ultra)
-done
+lemma starfun_ln_less_self [simp]: "!!x. 0 < x ==> ( *f* ln) x < x"
+by (transfer, simp)
 
-lemma starfun_ln_ge_zero [simp]: "1 \<le> x ==> 0 \<le> ( *f* ln) x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfun hypreal_zero_num hypreal_le hypreal_one_num, ultra)
-done
+lemma starfun_ln_ge_zero [simp]: "!!x. 1 \<le> x ==> 0 \<le> ( *f* ln) x"
+by (transfer, simp)
 
-lemma starfun_ln_gt_zero [simp]: "1 < x ==> 0 < ( *f* ln) x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfun hypreal_zero_num hypreal_less hypreal_one_num, ultra)
-done
+lemma starfun_ln_gt_zero [simp]: "!!x .1 < x ==> 0 < ( *f* ln) x"
+by (transfer, simp)
 
-lemma starfun_ln_not_eq_zero [simp]: "[| 0 < x; x \<noteq> 1 |] ==> ( *f* ln) x \<noteq> 0"
-apply (rule_tac z=x in eq_Abs_star)
-apply (auto simp add: starfun hypreal_zero_num hypreal_less hypreal_one_num, ultra)
-done
+lemma starfun_ln_not_eq_zero [simp]: "!!x. [| 0 < x; x \<noteq> 1 |] ==> ( *f* ln) x \<noteq> 0"
+by (transfer, simp)
 
 lemma starfun_ln_HFinite: "[| x \<in> HFinite; 1 \<le> x |] ==> ( *f* ln) x \<in> HFinite"
 apply (rule HFinite_bounded)
@@ -393,16 +357,13 @@
 apply (simp_all add: starfun_ln_less_self order_less_imp_le)
 done
 
-lemma starfun_ln_inverse: "0 < x ==> ( *f* ln) (inverse x) = -( *f* ln) x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfun hypreal_zero_num hypreal_minus hypreal_inverse hypreal_less, ultra)
-apply (simp add: ln_inverse)
-done
+lemma starfun_ln_inverse: "!!x. 0 < x ==> ( *f* ln) (inverse x) = -( *f* ln) x"
+by (transfer, rule ln_inverse)
 
 lemma starfun_exp_HFinite: "x \<in> HFinite ==> ( *f* exp) x \<in> HFinite"
-apply (rule_tac z=x in eq_Abs_star)
+apply (cases x)
 apply (auto simp add: starfun HFinite_FreeUltrafilterNat_iff)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl, auto)
+apply (rule bexI [OF _ Rep_star_star_n], auto)
 apply (rule_tac x = "exp u" in exI)
 apply (ultra, arith)
 done
@@ -447,10 +408,8 @@
 apply (auto simp add: starfun_ln_inverse HInfinite_minus_iff)
 done
 
-lemma starfun_ln_less_zero: "[| 0 < x; x < 1 |] ==> ( *f* ln) x < 0"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: hypreal_zero_num hypreal_one_num hypreal_less starfun, ultra)
-done
+lemma starfun_ln_less_zero: "!!x. [| 0 < x; x < 1 |] ==> ( *f* ln) x < 0"
+by (transfer, simp)
 
 lemma starfun_ln_Infinitesimal_less_zero:
      "[| x \<in> Infinitesimal; 0 < x |] ==> ( *f* ln) x < 0"
@@ -470,19 +429,19 @@
               ((- 1) ^ ((n - 1) div 2))/(real (fact n))) * x ^ n)  
               \<in> HFinite"
 apply (auto intro!: NSBseq_HFinite_hypreal NSconvergent_NSBseq 
-            simp add: starfunNat_sumr [symmetric] starfunNat hypnat_omega_def
+            simp add: starfunNat_sumr [symmetric] starfun hypnat_omega_def
                       convergent_NSconvergent_iff [symmetric] 
                       summable_convergent_sumr_iff [symmetric])
 apply (simp only: One_nat_def summable_sin)
 done
 
 lemma STAR_sin_zero [simp]: "( *f* sin) 0 = 0"
-by (simp add: starfun hypreal_zero_num)
+by (transfer, simp)
 
 lemma STAR_sin_Infinitesimal [simp]: "x \<in> Infinitesimal ==> ( *f* sin) x @= x"
 apply (case_tac "x = 0")
 apply (cut_tac [2] x = 0 in DERIV_sin)
-apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def hypreal_of_real_zero)
+apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
 apply (drule bspec [where x = x], auto)
 apply (drule approx_mult1 [where c = x])
 apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
@@ -494,32 +453,32 @@
             ((- 1) ^ (n div 2))/(real (fact n)) else  
             0) * x ^ n) \<in> HFinite"
 by (auto intro!: NSBseq_HFinite_hypreal NSconvergent_NSBseq 
-         simp add: starfunNat_sumr [symmetric] starfunNat hypnat_omega_def
+         simp add: starfunNat_sumr [symmetric] starfun hypnat_omega_def
                    convergent_NSconvergent_iff [symmetric] 
                    summable_convergent_sumr_iff [symmetric] summable_cos)
 
 lemma STAR_cos_zero [simp]: "( *f* cos) 0 = 1"
-by (simp add: starfun hypreal_zero_num hypreal_one_num)
+by (simp add: starfun star_n_zero_num star_n_one_num)
 
 lemma STAR_cos_Infinitesimal [simp]: "x \<in> Infinitesimal ==> ( *f* cos) x @= 1"
 apply (case_tac "x = 0")
 apply (cut_tac [2] x = 0 in DERIV_cos)
-apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def hypreal_of_real_zero)
+apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
 apply (drule bspec [where x = x])
-apply (auto simp add: hypreal_of_real_zero hypreal_of_real_one)
+apply auto
 apply (drule approx_mult1 [where c = x])
 apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
-            simp add: mult_assoc hypreal_of_real_one)
+            simp add: mult_assoc)
 apply (rule approx_add_right_cancel [where d = "-1"], auto)
 done
 
 lemma STAR_tan_zero [simp]: "( *f* tan) 0 = 0"
-by (simp add: starfun hypreal_zero_num)
+by (simp add: starfun star_n_zero_num)
 
 lemma STAR_tan_Infinitesimal: "x \<in> Infinitesimal ==> ( *f* tan) x @= x"
 apply (case_tac "x = 0")
 apply (cut_tac [2] x = 0 in DERIV_tan)
-apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def hypreal_of_real_zero)
+apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
 apply (drule bspec [where x = x], auto)
 apply (drule approx_mult1 [where c = x])
 apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
@@ -546,7 +505,7 @@
 lemma STAR_sin_Infinitesimal_divide:
      "[|x \<in> Infinitesimal; x \<noteq> 0 |] ==> ( *f* sin) x/x @= 1"
 apply (cut_tac x = 0 in DERIV_sin)
-apply (simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def hypreal_of_real_zero hypreal_of_real_one)
+apply (simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
 done
 
 (*------------------------------------------------------------------------*) 
@@ -600,35 +559,40 @@
 done
 
 lemma starfunNat_pi_divide_n_Infinitesimal: 
-     "N \<in> HNatInfinite ==> ( *fNat* (%x. pi / real x)) N \<in> Infinitesimal"
+     "N \<in> HNatInfinite ==> ( *f* (%x. pi / real x)) N \<in> Infinitesimal"
 by (auto intro!: Infinitesimal_HFinite_mult2 
-         simp add: starfunNat_mult [symmetric] divide_inverse
-                   starfunNat_inverse [symmetric] starfunNat_real_of_nat)
+         simp add: starfun_mult [symmetric] divide_inverse
+                   starfun_inverse [symmetric] starfunNat_real_of_nat)
 
 lemma STAR_sin_pi_divide_n_approx:
      "N \<in> HNatInfinite ==>  
-      ( *f* sin) (( *fNat* (%x. pi / real x)) N) @=  
+      ( *f* sin) (( *f* (%x. pi / real x)) N) @=  
       hypreal_of_real pi/(hypreal_of_hypnat N)"
-by (auto intro!: STAR_sin_Infinitesimal Infinitesimal_HFinite_mult2 
-         simp add: starfunNat_mult [symmetric] divide_inverse
-                   starfunNat_inverse_real_of_nat_eq)
+apply (simp add: starfunNat_real_of_nat [symmetric])
+apply (rule STAR_sin_Infinitesimal)
+apply (simp add: divide_inverse)
+apply (rule Infinitesimal_HFinite_mult2)
+apply (subst starfun_inverse)
+apply (erule starfunNat_inverse_real_of_nat_Infinitesimal)
+apply simp
+done
 
 lemma NSLIMSEQ_sin_pi: "(%n. real n * sin (pi / real n)) ----NS> pi"
-apply (auto simp add: NSLIMSEQ_def starfunNat_mult [symmetric] starfunNat_real_of_nat)
-apply (rule_tac f1 = sin in starfun_stafunNat_o2 [THEN subst])
-apply (auto simp add: starfunNat_mult [symmetric] starfunNat_real_of_nat divide_inverse)
-apply (rule_tac f1 = inverse in starfun_stafunNat_o2 [THEN subst])
+apply (auto simp add: NSLIMSEQ_def starfun_mult [symmetric] starfunNat_real_of_nat)
+apply (rule_tac f1 = sin in starfun_o2 [THEN subst])
+apply (auto simp add: starfun_mult [symmetric] starfunNat_real_of_nat divide_inverse)
+apply (rule_tac f1 = inverse in starfun_o2 [THEN subst])
 apply (auto dest: STAR_sin_pi_divide_HNatInfinite_approx_pi 
             simp add: starfunNat_real_of_nat mult_commute divide_inverse)
 done
 
 lemma NSLIMSEQ_cos_one: "(%n. cos (pi / real n))----NS> 1"
 apply (simp add: NSLIMSEQ_def, auto)
-apply (rule_tac f1 = cos in starfun_stafunNat_o2 [THEN subst])
+apply (rule_tac f1 = cos in starfun_o2 [THEN subst])
 apply (rule STAR_cos_Infinitesimal)
 apply (auto intro!: Infinitesimal_HFinite_mult2 
-            simp add: starfunNat_mult [symmetric] divide_inverse
-                      starfunNat_inverse [symmetric] starfunNat_real_of_nat)
+            simp add: starfun_mult [symmetric] divide_inverse
+                      starfun_inverse [symmetric] starfunNat_real_of_nat)
 done
 
 lemma NSLIMSEQ_sin_cos_pi:
--- a/src/HOL/Hyperreal/HyperArith.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/HyperArith.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -13,63 +13,10 @@
 
 subsection{*Numerals and Arithmetic*}
 
-(*
-instance hypreal :: number ..
-
-defs (overloaded)
-  hypreal_number_of_def: "(number_of w :: hypreal) == of_int (Rep_Bin w)"
-    --{*the type constraint is essential!*}
-
-instance hypreal :: number_ring
-by (intro_classes, simp add: hypreal_number_of_def) 
-*)
-
-text{*Collapse applications of @{term hypreal_of_real} to @{term number_of}*}
-lemma hypreal_number_of [simp]: "hypreal_of_real (number_of w) = number_of w"
-by (simp add: hypreal_of_real_def)
-
-
-
 use "hypreal_arith.ML"
 
 setup hypreal_arith_setup
 
-lemma hypreal_le_add_order: "[| 0 \<le> x; 0 \<le> y |] ==> (0::hypreal) \<le> x + y"
-by arith
-
-
-subsection{*The Function @{term hypreal_of_real}*}
-
-lemma number_of_less_hypreal_of_real_iff [simp]:
-     "(number_of w < hypreal_of_real z) = (number_of w < z)"
-apply (subst hypreal_of_real_less_iff [symmetric])
-apply (simp (no_asm))
-done
-
-lemma number_of_le_hypreal_of_real_iff [simp]:
-     "(number_of w \<le> hypreal_of_real z) = (number_of w \<le> z)"
-apply (subst hypreal_of_real_le_iff [symmetric])
-apply (simp (no_asm))
-done
-
-lemma hypreal_of_real_eq_number_of_iff [simp]:
-     "(hypreal_of_real z = number_of w) = (z = number_of w)"
-apply (subst hypreal_of_real_eq_iff [symmetric])
-apply (simp (no_asm))
-done
-
-lemma hypreal_of_real_less_number_of_iff [simp]:
-     "(hypreal_of_real z < number_of w) = (z < number_of w)"
-apply (subst hypreal_of_real_less_iff [symmetric])
-apply (simp (no_asm))
-done
-
-lemma hypreal_of_real_le_number_of_iff [simp]:
-     "(hypreal_of_real z \<le> number_of w) = (z \<le> number_of w)"
-apply (subst hypreal_of_real_le_iff [symmetric])
-apply (simp (no_asm))
-done
-
 subsection{*Absolute Value Function for the Hyperreals*}
 
 lemma hrabs_add_less:
@@ -89,22 +36,18 @@
 
 lemma hypreal_of_real_hrabs:
     "abs (hypreal_of_real r) = hypreal_of_real (abs r)"
-apply (unfold hypreal_of_real_def)
-apply (auto simp add: hypreal_hrabs)
-done
+by (rule star_of_abs [symmetric])
 
 
 subsection{*Embedding the Naturals into the Hyperreals*}
 
 constdefs
-
   hypreal_of_nat   :: "nat => hypreal"
    "hypreal_of_nat m  == of_nat m"
 
 lemma SNat_eq: "Nats = {n. \<exists>N. n = hypreal_of_nat N}"
 by (force simp add: hypreal_of_nat_def Nats_def) 
 
-
 lemma hypreal_of_nat_add [simp]:
      "hypreal_of_nat (m + n) = hypreal_of_nat m + hypreal_of_nat n"
 by (simp add: hypreal_of_nat_def)
@@ -131,11 +74,10 @@
 done
 
 lemma hypreal_of_nat:
-     "hypreal_of_nat m = Abs_star(starrel``{%n. real m})"
-apply (fold star_n_def star_of_def)
+     "hypreal_of_nat m = star_n (%n. real m)"
+apply (fold star_of_def)
 apply (induct m)
-apply (simp_all add: hypreal_of_nat_def real_of_nat_def
-                     hypreal_add)
+apply (simp_all add: hypreal_of_nat_def real_of_nat_def star_n_add)
 done
 
 lemma hypreal_of_nat_Suc:
@@ -171,8 +113,6 @@
 
 ML
 {*
-val hypreal_le_add_order = thm"hypreal_le_add_order";
-
 val hypreal_of_nat_def = thm"hypreal_of_nat_def";
 
 val hrabs_add_less = thm "hrabs_add_less";
--- a/src/HOL/Hyperreal/HyperDef.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/HyperDef.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -8,73 +8,29 @@
 header{*Construction of Hyperreals Using Ultrafilters*}
 
 theory HyperDef
-(*imports Filter "../Real/Real"*)
 imports StarClasses "../Real/Real"
 uses ("fuf.ML")  (*Warning: file fuf.ML refers to the name Hyperdef!*)
 begin
 
 types hypreal = "real star"
-(*
-constdefs
 
-  FreeUltrafilterNat   :: "nat set set"    ("\<U>")
-    "FreeUltrafilterNat == (SOME U. freeultrafilter U)"
-
-  hyprel :: "((nat=>real)*(nat=>real)) set"
-    "hyprel == {p. \<exists>X Y. p = ((X::nat=>real),Y) &
-                   {n::nat. X(n) = Y(n)} \<in> FreeUltrafilterNat}"
-
-typedef hypreal = "UNIV//hyprel" 
-    by (auto simp add: quotient_def) 
-
-instance hypreal :: "{ord, zero, one, plus, times, minus, inverse}" ..
-
-defs (overloaded)
-
-  hypreal_zero_def:
-  "0 == Abs_hypreal(hyprel``{%n. 0})"
-
-  hypreal_one_def:
-  "1 == Abs_hypreal(hyprel``{%n. 1})"
+syntax hypreal_of_real :: "real => real star"
+translations "hypreal_of_real" => "star_of :: real => real star"
 
-  hypreal_minus_def:
-  "- P == Abs_hypreal(\<Union>X \<in> Rep_hypreal(P). hyprel``{%n. - (X n)})"
-
-  hypreal_diff_def:
-  "x - y == x + -(y::hypreal)"
-
-  hypreal_inverse_def:
-  "inverse P == Abs_hypreal(\<Union>X \<in> Rep_hypreal(P).
-                    hyprel``{%n. if X n = 0 then 0 else inverse (X n)})"
-
-  hypreal_divide_def:
-  "P / Q::hypreal == P * inverse Q"
-*)
-
-lemma hypreal_zero_def: "0 == Abs_star(starrel``{%n. 0})"
-by (simp only: star_zero_def star_of_def star_n_def)
-
-lemma hypreal_one_def: "1 == Abs_star(starrel``{%n. 1})"
-by (simp only: star_one_def star_of_def star_n_def)
-
-lemma hypreal_diff_def: "x - y == x + -(y::hypreal)"
-by (rule diff_def)
-
-lemma hypreal_divide_def: "P / Q::hypreal == P * inverse Q"
-by (rule divide_inverse [THEN eq_reflection])
+typed_print_translation {*
+  let
+    fun hr_tr' _ (Type ("fun", (Type ("real", []) :: _))) [t] =
+          Syntax.const "hypreal_of_real" $ t
+      | hr_tr' _ _ _ = raise Match;
+  in [("star_of", hr_tr')] end
+*}
 
 constdefs
 
-  hypreal_of_real  :: "real => hypreal"
-(*  "hypreal_of_real r         == Abs_hypreal(hyprel``{%n. r})" *)
-  "hypreal_of_real r == star_of r"
-
   omega   :: hypreal   -- {*an infinite number @{text "= [<1,2,3,...>]"} *}
-(*  "omega == Abs_hypreal(hyprel``{%n. real (Suc n)})" *)
   "omega == star_n (%n. real (Suc n))"
 
   epsilon :: hypreal   -- {*an infinitesimal number @{text "= [<1,1/2,1/3,...>]"} *}
-(*  "epsilon == Abs_hypreal(hyprel``{%n. inverse (real (Suc n))})" *)
   "epsilon == star_n (%n. inverse (real (Suc n)))"
 
 syntax (xsymbols)
@@ -85,26 +41,6 @@
   omega   :: hypreal   ("\<omega>")
   epsilon :: hypreal   ("\<epsilon>")
 
-(*
-defs (overloaded)
-
-  hypreal_add_def:
-  "P + Q == Abs_hypreal(\<Union>X \<in> Rep_hypreal(P). \<Union>Y \<in> Rep_hypreal(Q).
-                hyprel``{%n. X n + Y n})"
-
-  hypreal_mult_def:
-  "P * Q == Abs_hypreal(\<Union>X \<in> Rep_hypreal(P). \<Union>Y \<in> Rep_hypreal(Q).
-                hyprel``{%n. X n * Y n})"
-
-  hypreal_le_def:
-  "P \<le> (Q::hypreal) == \<exists>X Y. X \<in> Rep_hypreal(P) &
-                               Y \<in> Rep_hypreal(Q) &
-                               {n. X n \<le> Y n} \<in> FreeUltrafilterNat"
-
-  hypreal_less_def: "(x < (y::hypreal)) == (x \<le> y & x \<noteq> y)"
-
-  hrabs_def:  "abs (r::hypreal) == (if 0 \<le> r then r else -r)"
-*)
 
 subsection{*The Set of Naturals is not Finite*}
 
@@ -277,104 +213,46 @@
             the Injection from @{typ real} to @{typ hypreal}*}
 
 lemma inj_hypreal_of_real: "inj(hypreal_of_real)"
-apply (rule inj_onI)
-apply (simp add: hypreal_of_real_def split: split_if_asm)
-done
+by (rule inj_onI, simp)
 
 lemma eq_Abs_star:
     "(!!x. z = Abs_star(starrel``{x}) ==> P) ==> P"
 by (fold star_n_def, auto intro: star_cases)
 
-(*
-theorem hypreal_cases [case_names Abs_hypreal, cases type: hypreal]:
-    "(!!x. z = star_n x ==> P) ==> P"
-by (rule eq_Abs_hypreal [of z], blast)
-*)
+lemma Rep_star_star_n_iff [simp]:
+  "(X \<in> Rep_star (star_n Y)) = ({n. Y n = X n} \<in> \<U>)"
+by (simp add: star_n_def)
+
+lemma Rep_star_star_n: "X \<in> Rep_star (star_n X)"
+by simp
 
 subsection{*Hyperreal Addition*}
-(*
-lemma hypreal_add_congruent2: 
-    "congruent2 starrel starrel (%X Y. starrel``{%n. X n + Y n})"
-by (simp add: congruent2_def, auto, ultra)
-*)
-lemma hypreal_add [unfolded star_n_def]: 
+
+lemma star_n_add:
   "star_n X + star_n Y = star_n (%n. X n + Y n)"
 by (simp add: star_add_def Ifun2_of_def star_of_def Ifun_star_n)
 
-lemma hypreal_add_commute: "(z::hypreal) + w = w + z"
-by (rule add_commute)
-
-lemma hypreal_add_assoc: "((z1::hypreal) + z2) + z3 = z1 + (z2 + z3)"
-by (rule add_assoc)
-
-lemma hypreal_add_zero_left: "(0::hypreal) + z = z"
-by (rule comm_monoid_add_class.add_0)
+subsection{*Additive inverse on @{typ hypreal}*}
 
-(*instance hypreal :: comm_monoid_add
-  by intro_classes
-    (assumption | 
-      rule hypreal_add_commute hypreal_add_assoc hypreal_add_zero_left)+*)
-
-lemma hypreal_add_zero_right [simp]: "z + (0::hypreal) = z"
-by (rule OrderedGroup.add_0_right)
-
-
-subsection{*Additive inverse on @{typ hypreal}*}
-(*
-lemma hypreal_minus_congruent: "(%X. starrel``{%n. - (X n)}) respects starrel"
-by (force simp add: congruent_def)
-*)
-lemma hypreal_minus [unfolded star_n_def]: 
+lemma star_n_minus:
    "- star_n X = star_n (%n. -(X n))"
 by (simp add: star_minus_def Ifun_of_def star_of_def Ifun_star_n)
 
-lemma hypreal_diff [unfolded star_n_def]:
+lemma star_n_diff:
      "star_n X - star_n Y = star_n (%n. X n - Y n)"
 by (simp add: star_diff_def Ifun2_of_def star_of_def Ifun_star_n)
 
-lemma hypreal_add_minus [simp]: "z + -z = (0::hypreal)"
-by (rule right_minus)
-
-lemma hypreal_add_minus_left: "-z + z = (0::hypreal)"
-by (rule left_minus)
-
 
 subsection{*Hyperreal Multiplication*}
-(*
-lemma hypreal_mult_congruent2: 
-    "congruent2 starrel starrel (%X Y. starrel``{%n. X n * Y n})"
-by (simp add: congruent2_def, auto, ultra)
-*)
 
-lemma hypreal_mult [unfolded star_n_def]: 
+lemma star_n_mult:
   "star_n X * star_n Y = star_n (%n. X n * Y n)"
 by (simp add: star_mult_def Ifun2_of_def star_of_def Ifun_star_n)
 
-lemma hypreal_mult_commute: "(z::hypreal) * w = w * z"
-by (rule mult_commute)
-
-lemma hypreal_mult_assoc: "((z1::hypreal) * z2) * z3 = z1 * (z2 * z3)"
-by (rule mult_assoc)
-
-lemma hypreal_mult_1: "(1::hypreal) * z = z"
-by (rule mult_1_left)
-
-lemma hypreal_add_mult_distrib:
-     "((z1::hypreal) + z2) * w = (z1 * w) + (z2 * w)"
-by (rule left_distrib)
-
-text{*one and zero are distinct*}
-lemma hypreal_zero_not_eq_one: "0 \<noteq> (1::hypreal)"
-by (rule zero_neq_one)
-
 
 subsection{*Multiplicative Inverse on @{typ hypreal} *}
-(*
-lemma hypreal_inverse_congruent: 
-  "(%X. starrel``{%n. if X n = 0 then 0 else inverse(X n)}) respects starrel"
-by (auto simp add: congruent_def, ultra)
-*)
-lemma hypreal_inverse [unfolded star_n_def]: 
+
+lemma star_n_inverse:
       "inverse (star_n X) = star_n (%n. if X n = (0::real) then 0 else inverse(X n))"
 apply (simp add: star_inverse_def Ifun_of_def star_of_def Ifun_star_n)
 apply (rule_tac f=star_n in arg_cong)
@@ -382,102 +260,23 @@
 apply simp
 done
 
-lemma hypreal_inverse2 [unfolded star_n_def]: 
+lemma star_n_inverse2:
       "inverse (star_n X) = star_n (%n. inverse(X n))"
 by (simp add: star_inverse_def Ifun_of_def star_of_def Ifun_star_n)
 
-lemma hypreal_mult_inverse: 
-     "x \<noteq> 0 ==> x*inverse(x) = (1::hypreal)"
-by (rule right_inverse)
-
-lemma hypreal_mult_inverse_left:
-     "x \<noteq> 0 ==> inverse(x)*x = (1::hypreal)"
-by (rule left_inverse)
-
-(*
-instance hypreal :: field
-proof
-  fix x y z :: hypreal
-  show "- x + x = 0" by (simp add: hypreal_add_minus_left)
-  show "x - y = x + (-y)" by (simp add: hypreal_diff_def)
-  show "(x * y) * z = x * (y * z)" by (rule hypreal_mult_assoc)
-  show "x * y = y * x" by (rule hypreal_mult_commute)
-  show "1 * x = x" by (simp add: hypreal_mult_1)
-  show "(x + y) * z = x * z + y * z" by (simp add: hypreal_add_mult_distrib)
-  show "0 \<noteq> (1::hypreal)" by (rule hypreal_zero_not_eq_one)
-  show "x \<noteq> 0 ==> inverse x * x = 1" by (simp add: hypreal_mult_inverse_left)
-  show "x / y = x * inverse y" by (simp add: hypreal_divide_def)
-qed
-
-
-instance hypreal :: division_by_zero
-proof
-  show "inverse 0 = (0::hypreal)" 
-    by (simp add: hypreal_inverse hypreal_zero_def)
-qed
-*)
 
 subsection{*Properties of The @{text "\<le>"} Relation*}
 
-lemma hypreal_le [unfolded star_n_def]: 
+lemma star_n_le:
       "star_n X \<le> star_n Y =  
        ({n. X n \<le> Y n} \<in> FreeUltrafilterNat)"
 by (simp add: star_le_def Ipred2_of_def star_of_def Ifun_star_n unstar_star_n)
 
-lemma hypreal_le_refl: "w \<le> (w::hypreal)"
-by (rule order_refl)
-
-lemma hypreal_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::hypreal)"
-by (rule order_trans)
-
-lemma hypreal_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::hypreal)"
-by (rule order_antisym)
-
-(* Axiom 'order_less_le' of class 'order': *)
-lemma hypreal_less_le: "((w::hypreal) < z) = (w \<le> z & w \<noteq> z)"
-by (rule order_less_le)
-
-(*
-instance hypreal :: order
-  by intro_classes
-    (assumption |
-      rule hypreal_le_refl hypreal_le_trans hypreal_le_anti_sym hypreal_less_le)+
-*)
-
-(* Axiom 'linorder_linear' of class 'linorder': *)
-lemma hypreal_le_linear: "(z::hypreal) \<le> w | w \<le> z"
-by (rule linorder_linear)
-
-(*
-instance hypreal :: linorder 
-  by intro_classes (rule hypreal_le_linear)
-*)
-
 lemma hypreal_not_refl2: "!!(x::hypreal). x < y ==> x \<noteq> y"
 by (auto)
 
-lemma hypreal_add_left_mono: "x \<le> y ==> z + x \<le> z + (y::hypreal)"
-by (rule add_left_mono)
-
-lemma hypreal_mult_less_mono2: "[| (0::hypreal)<z; x<y |] ==> z*x<z*y"
-by (rule mult_strict_left_mono)
-
-
 subsection{*The Hyperreals Form an Ordered Field*}
 
-(*
-instance hypreal :: ordered_field
-proof
-  fix x y z :: hypreal
-  show "x \<le> y ==> z + x \<le> z + y" 
-    by (rule hypreal_add_left_mono)
-  show "x < y ==> 0 < z ==> z * x < z * y" 
-    by (simp add: hypreal_mult_less_mono2)
-  show "\<bar>x\<bar> = (if x < 0 then -x else x)"
-    by (auto dest: order_le_less_trans simp add: hrabs_def linorder_not_le)
-qed
-*)
-
 lemma hypreal_eq_minus_iff: "((x::hypreal) = y) = (x + - y = 0)"
 by auto
 
@@ -488,88 +287,16 @@
 by auto
 
 
-subsection{*The Embedding @{term hypreal_of_real} Preserves Field and 
-      Order Properties*}
-
-lemma hypreal_of_real_add [simp]: 
-     "hypreal_of_real (w + z) = hypreal_of_real w + hypreal_of_real z"
-by (simp add: hypreal_of_real_def)
-
-lemma hypreal_of_real_minus [simp]:
-     "hypreal_of_real (-r) = - hypreal_of_real  r"
-by (simp add: hypreal_of_real_def)
-
-lemma hypreal_of_real_diff [simp]: 
-     "hypreal_of_real (w - z) = hypreal_of_real w - hypreal_of_real z"
-by (simp add: hypreal_of_real_def)
-
-lemma hypreal_of_real_mult [simp]: 
-     "hypreal_of_real (w * z) = hypreal_of_real w * hypreal_of_real z"
-by (simp add: hypreal_of_real_def)
-
-lemma hypreal_of_real_one [simp]: "hypreal_of_real 1 = (1::hypreal)"
-by (simp add: hypreal_of_real_def)
-
-lemma hypreal_of_real_zero [simp]: "hypreal_of_real 0 = 0"
-by (simp add: hypreal_of_real_def)
-
-lemma hypreal_of_real_le_iff [simp]: 
-     "(hypreal_of_real w \<le> hypreal_of_real z) = (w \<le> z)"
-by (simp add: hypreal_of_real_def)
-
-lemma hypreal_of_real_less_iff [simp]: 
-     "(hypreal_of_real w < hypreal_of_real z) = (w < z)"
-by (simp add: hypreal_of_real_def)
-
-lemma hypreal_of_real_eq_iff [simp]:
-     "(hypreal_of_real w = hypreal_of_real z) = (w = z)"
-by (simp add: hypreal_of_real_def)
-
-text{*As above, for 0*}
-
-declare hypreal_of_real_less_iff [of 0, simplified, simp]
-declare hypreal_of_real_le_iff   [of 0, simplified, simp]
-declare hypreal_of_real_eq_iff   [of 0, simplified, simp]
-
-declare hypreal_of_real_less_iff [of _ 0, simplified, simp]
-declare hypreal_of_real_le_iff   [of _ 0, simplified, simp]
-declare hypreal_of_real_eq_iff   [of _ 0, simplified, simp]
-
-text{*As above, for 1*}
-
-declare hypreal_of_real_less_iff [of 1, simplified, simp]
-declare hypreal_of_real_le_iff   [of 1, simplified, simp]
-declare hypreal_of_real_eq_iff   [of 1, simplified, simp]
-
-declare hypreal_of_real_less_iff [of _ 1, simplified, simp]
-declare hypreal_of_real_le_iff   [of _ 1, simplified, simp]
-declare hypreal_of_real_eq_iff   [of _ 1, simplified, simp]
-
-lemma hypreal_of_real_inverse [simp]:
-     "hypreal_of_real (inverse r) = inverse (hypreal_of_real r)"
-by (simp add: hypreal_of_real_def)
-
-lemma hypreal_of_real_divide [simp]:
-     "hypreal_of_real (w / z) = hypreal_of_real w / hypreal_of_real z"
-by (simp add: hypreal_of_real_def)
-
-lemma hypreal_of_real_of_nat [simp]: "hypreal_of_real (of_nat n) = of_nat n"
-by (simp add: hypreal_of_real_def)
-
-lemma hypreal_of_real_of_int [simp]:  "hypreal_of_real (of_int z) = of_int z"
-by (simp add: hypreal_of_real_def)
-
-
 subsection{*Misc Others*}
 
-lemma hypreal_less [unfolded star_n_def]: 
+lemma star_n_less:
       "star_n X < star_n Y = ({n. X n < Y n} \<in> FreeUltrafilterNat)"
 by (simp add: star_less_def Ipred2_of_def star_of_def Ifun_star_n unstar_star_n)
 
-lemma hypreal_zero_num [unfolded star_n_def]: "0 = star_n (%n. 0)"
+lemma star_n_zero_num: "0 = star_n (%n. 0)"
 by (simp add: star_zero_def star_of_def)
 
-lemma hypreal_one_num [unfolded star_n_def]: "1 = star_n (%n. 1)"
+lemma star_n_one_num: "1 = star_n (%n. 1)"
 by (simp add: star_one_def star_of_def)
 
 lemma hypreal_omega_gt_zero [simp]: "0 < omega"
@@ -577,15 +304,12 @@
 apply (simp add: Ipred2_of_def star_of_def Ifun_star_n unstar_star_n)
 done
 
-lemma hypreal_hrabs [unfolded star_n_def]:
+lemma star_n_abs:
      "abs (star_n X) = star_n (%n. abs (X n))"
 by (simp add: star_abs_def Ifun_of_def star_of_def Ifun_star_n)
 
 subsection{*Existence of Infinite Hyperreal Number*}
-(*
-lemma Rep_hypreal_omega: "Rep_hypreal(omega) \<in> hypreal"
-by (simp add: omega_def)
-*)
+
 text{*Existence of infinite number not corresponding to any real number.
 Use assumption that member @{term FreeUltrafilterNat} is not finite.*}
 
@@ -601,7 +325,7 @@
 
 lemma not_ex_hypreal_of_real_eq_omega: 
       "~ (\<exists>x. hypreal_of_real x = omega)"
-apply (simp add: omega_def hypreal_of_real_def)
+apply (simp add: omega_def)
 apply (simp add: star_of_def star_n_eq_iff)
 apply (auto simp add: real_of_nat_Suc diff_eq_eq [symmetric] 
             lemma_finite_omega_set [THEN FreeUltrafilterNat_finite])
@@ -622,8 +346,7 @@
 by (cut_tac x = x in lemma_epsilon_empty_singleton_disj, auto)
 
 lemma not_ex_hypreal_of_real_eq_epsilon: "~ (\<exists>x. hypreal_of_real x = epsilon)"
-by (auto simp add: epsilon_def hypreal_of_real_def 
-                   star_of_def star_n_eq_iff
+by (auto simp add: epsilon_def star_of_def star_n_eq_iff
                    lemma_finite_epsilon_set [THEN FreeUltrafilterNat_finite])
 
 lemma hypreal_of_real_not_eq_epsilon: "hypreal_of_real x \<noteq> epsilon"
@@ -641,22 +364,8 @@
 
 ML
 {*
-(* val hrabs_def = thm "hrabs_def"; *)
-(* val hypreal_hrabs = thm "hypreal_hrabs"; *)
-
-val hypreal_zero_def = thm "hypreal_zero_def";
-(* val hypreal_one_def = thm "hypreal_one_def"; *)
-(* val hypreal_minus_def = thm "hypreal_minus_def"; *)
-(* val hypreal_diff_def = thm "hypreal_diff_def"; *)
-(* val hypreal_inverse_def = thm "hypreal_inverse_def"; *)
-(* val hypreal_divide_def = thm "hypreal_divide_def"; *)
-val hypreal_of_real_def = thm "hypreal_of_real_def";
 val omega_def = thm "omega_def";
 val epsilon_def = thm "epsilon_def";
-(* val hypreal_add_def = thm "hypreal_add_def"; *)
-(* val hypreal_mult_def = thm "hypreal_mult_def"; *)
-(* val hypreal_less_def = thm "hypreal_less_def"; *)
-(* val hypreal_le_def = thm "hypreal_le_def"; *)
 
 val finite_exhausts = thm "finite_exhausts";
 val finite_not_covers = thm "finite_not_covers";
@@ -686,51 +395,21 @@
 val Rep_hypreal_nonempty = thm "Rep_hypreal_nonempty";
 val inj_hypreal_of_real = thm "inj_hypreal_of_real";
 val eq_Abs_star = thm "eq_Abs_star";
-(* val hypreal_minus_congruent = thm "hypreal_minus_congruent"; *)
-val hypreal_minus = thm "hypreal_minus";
-val hypreal_add = thm "hypreal_add";
-val hypreal_diff = thm "hypreal_diff";
-val hypreal_add_commute = thm "hypreal_add_commute";
-val hypreal_add_assoc = thm "hypreal_add_assoc";
-val hypreal_add_zero_left = thm "hypreal_add_zero_left";
-val hypreal_add_zero_right = thm "hypreal_add_zero_right";
-val hypreal_add_minus = thm "hypreal_add_minus";
-val hypreal_add_minus_left = thm "hypreal_add_minus_left";
-val hypreal_mult = thm "hypreal_mult";
-val hypreal_mult_commute = thm "hypreal_mult_commute";
-val hypreal_mult_assoc = thm "hypreal_mult_assoc";
-val hypreal_mult_1 = thm "hypreal_mult_1";
-val hypreal_zero_not_eq_one = thm "hypreal_zero_not_eq_one";
-(* val hypreal_inverse_congruent = thm "hypreal_inverse_congruent"; *)
-val hypreal_inverse = thm "hypreal_inverse";
-val hypreal_mult_inverse = thm "hypreal_mult_inverse";
-val hypreal_mult_inverse_left = thm "hypreal_mult_inverse_left";
+val star_n_minus = thm "star_n_minus";
+val star_n_add = thm "star_n_add";
+val star_n_diff = thm "star_n_diff";
+val star_n_mult = thm "star_n_mult";
+val star_n_inverse = thm "star_n_inverse";
 val hypreal_mult_left_cancel = thm "hypreal_mult_left_cancel";
 val hypreal_mult_right_cancel = thm "hypreal_mult_right_cancel";
 val hypreal_not_refl2 = thm "hypreal_not_refl2";
-val hypreal_less = thm "hypreal_less";
+val star_n_less = thm "star_n_less";
 val hypreal_eq_minus_iff = thm "hypreal_eq_minus_iff";
-val hypreal_le = thm "hypreal_le";
-val hypreal_le_refl = thm "hypreal_le_refl";
-val hypreal_le_linear = thm "hypreal_le_linear";
-val hypreal_le_trans = thm "hypreal_le_trans";
-val hypreal_le_anti_sym = thm "hypreal_le_anti_sym";
-val hypreal_less_le = thm "hypreal_less_le";
-val hypreal_of_real_add = thm "hypreal_of_real_add";
-val hypreal_of_real_mult = thm "hypreal_of_real_mult";
-val hypreal_of_real_less_iff = thm "hypreal_of_real_less_iff";
-val hypreal_of_real_le_iff = thm "hypreal_of_real_le_iff";
-val hypreal_of_real_eq_iff = thm "hypreal_of_real_eq_iff";
-val hypreal_of_real_minus = thm "hypreal_of_real_minus";
-val hypreal_of_real_one = thm "hypreal_of_real_one";
-val hypreal_of_real_zero = thm "hypreal_of_real_zero";
-val hypreal_of_real_inverse = thm "hypreal_of_real_inverse";
-val hypreal_of_real_divide = thm "hypreal_of_real_divide";
-val hypreal_zero_num = thm "hypreal_zero_num";
-val hypreal_one_num = thm "hypreal_one_num";
+val star_n_le = thm "star_n_le";
+val star_n_zero_num = thm "star_n_zero_num";
+val star_n_one_num = thm "star_n_one_num";
 val hypreal_omega_gt_zero = thm "hypreal_omega_gt_zero";
 
-(* val Rep_hypreal_omega = thm"Rep_hypreal_omega"; *)
 val lemma_omega_empty_singleton_disj = thm"lemma_omega_empty_singleton_disj";
 val lemma_finite_omega_set = thm"lemma_finite_omega_set";
 val not_ex_hypreal_of_real_eq_omega = thm"not_ex_hypreal_of_real_eq_omega";
--- a/src/HOL/Hyperreal/HyperNat.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/HyperNat.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -12,214 +12,48 @@
 begin
 
 types hypnat = "nat star"
-(*
-constdefs
-    hypnatrel :: "((nat=>nat)*(nat=>nat)) set"
-    "hypnatrel == {p. \<exists>X Y. p = ((X::nat=>nat),Y) &
-                       {n::nat. X(n) = Y(n)} \<in> FreeUltrafilterNat}"
 
-typedef hypnat = "UNIV//hypnatrel"
-    by (auto simp add: quotient_def)
+syntax hypnat_of_nat :: "nat => nat star"
+translations "hypnat_of_nat" => "star_of :: nat => nat star"
 
-instance hypnat :: "{ord, zero, one, plus, times, minus}" ..
-*)
 consts whn :: hypnat
 
-
 defs
   (* omega is in fact an infinite hypernatural number = [<1,2,3,...>] *)
-  hypnat_omega_def:  "whn == Abs_star(starrel``{%n::nat. n})"
-
-lemma hypnat_zero_def:  "0 == Abs_star(starrel``{%n::nat. 0})"
-by (simp only: star_zero_def star_of_def star_n_def)
-
-lemma hypnat_one_def:   "1 == Abs_star(starrel``{%n::nat. 1})"
-by (simp only: star_one_def star_of_def star_n_def)
-
-  (** hypernatural arithmetic **)
-(*
-  hypnat_zero_def:  "0 == Abs_hypnat(hypnatrel``{%n::nat. 0})"
-  hypnat_one_def:   "1 == Abs_hypnat(hypnatrel``{%n::nat. 1})"
-*)
-(*
-  hypnat_add_def:
-  "P + Q == Abs_hypnat(\<Union>X \<in> Rep_hypnat(P). \<Union>Y \<in> Rep_hypnat(Q).
-                hypnatrel``{%n::nat. X n + Y n})"
-
-  hypnat_mult_def:
-  "P * Q == Abs_hypnat(\<Union>X \<in> Rep_hypnat(P). \<Union>Y \<in> Rep_hypnat(Q).
-                hypnatrel``{%n::nat. X n * Y n})"
-
-  hypnat_minus_def:
-  "P - Q == Abs_hypnat(\<Union>X \<in> Rep_hypnat(P). \<Union>Y \<in> Rep_hypnat(Q).
-                hypnatrel``{%n::nat. X n - Y n})"
-*)
-
-(*
-subsection{*Properties of @{term hypnatrel}*}
-
-text{*Proving that @{term hypnatrel} is an equivalence relation*}
-
-lemma hypnatrel_iff:
-     "((X,Y) \<in> hypnatrel) = ({n. X n = Y n}: FreeUltrafilterNat)"
-apply (simp add: hypnatrel_def)
-done
-
-lemma hypnatrel_refl: "(x,x) \<in> hypnatrel"
-by (simp add: hypnatrel_def)
-
-lemma hypnatrel_sym: "(x,y) \<in> hypnatrel ==> (y,x) \<in> hypnatrel"
-by (auto simp add: hypnatrel_def eq_commute)
-
-lemma hypnatrel_trans [rule_format (no_asm)]:
-     "(x,y) \<in> hypnatrel --> (y,z) \<in> hypnatrel --> (x,z) \<in> hypnatrel"
-by (auto simp add: hypnatrel_def, ultra)
-
-lemma equiv_hypnatrel:
-     "equiv UNIV hypnatrel"
-apply (simp add: equiv_def refl_def sym_def trans_def hypnatrel_refl)
-apply (blast intro: hypnatrel_sym hypnatrel_trans)
-done
-*)
-(* (hypnatrel `` {x} = hypnatrel `` {y}) = ((x,y) \<in> hypnatrel) *)
-(*
-lemmas equiv_hypnatrel_iff =
-    eq_equiv_class_iff [OF equiv_hypnatrel UNIV_I UNIV_I, simp]
-
-lemma hypnatrel_in_hypnat [simp]: "hypnatrel``{x}:hypnat"
-by (simp add: hypnat_def hypnatrel_def quotient_def, blast)
-
-declare Abs_hypnat_inject [simp] Abs_hypnat_inverse [simp]
-declare equiv_hypnatrel [THEN eq_equiv_class_iff, simp]
-declare hypnatrel_iff [iff]
+  hypnat_omega_def:  "whn == star_n (%n::nat. n)"
 
-lemma lemma_hypnatrel_refl: "x \<in> hypnatrel `` {x}"
-by (simp add: hypnatrel_def)
-
-declare lemma_hypnatrel_refl [simp]
-
-lemma hypnat_empty_not_mem: "{} \<notin> hypnat"
-apply (simp add: hypnat_def)
-apply (auto elim!: quotientE equalityCE)
-done
-
-declare hypnat_empty_not_mem [simp]
-
-lemma Rep_hypnat_nonempty: "Rep_hypnat x \<noteq> {}"
-by (cut_tac x = x in Rep_hypnat, auto)
-
-declare Rep_hypnat_nonempty [simp]
-
-
-lemma eq_Abs_hypnat:
-    "(!!x. z = Abs_hypnat(hypnatrel``{x}) ==> P) ==> P"
-apply (rule_tac x1=z in Rep_hypnat [unfolded hypnat_def, THEN quotientE])
-apply (drule_tac f = Abs_hypnat in arg_cong)
-apply (force simp add: Rep_hypnat_inverse)
-done
-
-theorem hypnat_cases [case_names Abs_hypnat, cases type: hypnat]:
-    "(!!x. z = Abs_hypnat(hypnatrel``{x}) ==> P) ==> P"
-by (rule eq_Abs_hypnat [of z], blast)
-*)
-subsection{*Hypernat Addition*}
-(*
-lemma hypnat_add_congruent2:
-     "(%X Y. hypnatrel``{%n. X n + Y n}) respects2 hypnatrel"
-by (simp add: congruent2_def, auto, ultra)
-*)
-lemma hypnat_add:
-  "Abs_star(starrel``{%n. X n}) + Abs_star(starrel``{%n. Y n}) =
-   Abs_star(starrel``{%n. X n + Y n})"
-by (rule hypreal_add)
-
-lemma hypnat_add_commute: "(z::hypnat) + w = w + z"
-by (rule add_commute)
-
-lemma hypnat_add_assoc: "((z1::hypnat) + z2) + z3 = z1 + (z2 + z3)"
-by (rule add_assoc)
-
-lemma hypnat_add_zero_left: "(0::hypnat) + z = z"
-by (rule comm_monoid_add_class.add_0)
-
-(*
-instance hypnat :: comm_monoid_add
-  by intro_classes
-    (assumption |
-      rule hypnat_add_commute hypnat_add_assoc hypnat_add_zero_left)+
-*)
-
-subsection{*Subtraction inverse on @{typ hypreal}*}
-
-(*
-lemma hypnat_minus_congruent2:
-    "(%X Y. starrel``{%n. X n - Y n}) respects2 starrel"
-by (simp add: congruent2_def, auto, ultra)
-*)
-lemma hypnat_minus:
-  "Abs_star(starrel``{%n. X n}) - Abs_star(starrel``{%n. Y n}) =
-   Abs_star(starrel``{%n. X n - Y n})"
-by (rule hypreal_diff)
-
-lemma hypnat_minus_zero: "!!z. z - z = (0::hypnat)"
+lemma hypnat_minus_zero [simp]: "!!z. z - z = (0::hypnat)"
 by transfer (rule diff_self_eq_0)
 
-lemma hypnat_diff_0_eq_0: "!!n. (0::hypnat) - n = 0"
+lemma hypnat_diff_0_eq_0 [simp]: "!!n. (0::hypnat) - n = 0"
 by transfer (rule diff_0_eq_0)
 
-declare hypnat_minus_zero [simp] hypnat_diff_0_eq_0 [simp]
-
-lemma hypnat_add_is_0: "!!m n. (m+n = (0::hypnat)) = (m=0 & n=0)"
+lemma hypnat_add_is_0 [iff]: "!!m n. (m+n = (0::hypnat)) = (m=0 & n=0)"
 by transfer (rule add_is_0)
 
-declare hypnat_add_is_0 [iff]
-
 lemma hypnat_diff_diff_left: "!!i j k. (i::hypnat) - j - k = i - (j+k)"
 by transfer (rule diff_diff_left)
 
 lemma hypnat_diff_commute: "!!i j k. (i::hypnat) - j - k = i-k-j"
 by transfer (rule diff_commute)
 
-lemma hypnat_diff_add_inverse: "!!m n. ((n::hypnat) + m) - n = m"
+lemma hypnat_diff_add_inverse [simp]: "!!m n. ((n::hypnat) + m) - n = m"
 by transfer (rule diff_add_inverse)
-declare hypnat_diff_add_inverse [simp]
 
-lemma hypnat_diff_add_inverse2:  "!!m n. ((m::hypnat) + n) - n = m"
+lemma hypnat_diff_add_inverse2 [simp]:  "!!m n. ((m::hypnat) + n) - n = m"
 by transfer (rule diff_add_inverse2)
-declare hypnat_diff_add_inverse2 [simp]
 
-lemma hypnat_diff_cancel: "!!k m n. ((k::hypnat) + m) - (k+n) = m - n"
+lemma hypnat_diff_cancel [simp]: "!!k m n. ((k::hypnat) + m) - (k+n) = m - n"
 by transfer (rule diff_cancel)
-declare hypnat_diff_cancel [simp]
 
-lemma hypnat_diff_cancel2: "!!k m n. ((m::hypnat) + k) - (n+k) = m - n"
+lemma hypnat_diff_cancel2 [simp]: "!!k m n. ((m::hypnat) + k) - (n+k) = m - n"
 by transfer (rule diff_cancel2)
-declare hypnat_diff_cancel2 [simp]
 
-lemma hypnat_diff_add_0: "!!m n. (n::hypnat) - (n+m) = (0::hypnat)"
+lemma hypnat_diff_add_0 [simp]: "!!m n. (n::hypnat) - (n+m) = (0::hypnat)"
 by transfer (rule diff_add_0)
-declare hypnat_diff_add_0 [simp]
 
 
 subsection{*Hyperreal Multiplication*}
-(*
-lemma hypnat_mult_congruent2:
-    "(%X Y. starrel``{%n. X n * Y n}) respects2 starrel"
-by (simp add: congruent2_def, auto, ultra)
-*)
-lemma hypnat_mult:
-  "Abs_star(starrel``{%n. X n}) * Abs_star(starrel``{%n. Y n}) =
-   Abs_star(starrel``{%n. X n * Y n})"
-by (rule hypreal_mult)
-
-lemma hypnat_mult_commute: "(z::hypnat) * w = w * z"
-by (rule mult_commute)
-
-lemma hypnat_mult_assoc: "((z1::hypnat) * z2) * z3 = z1 * (z2 * z3)"
-by (rule mult_assoc)
-
-lemma hypnat_mult_1: "(1::hypnat) * z = z"
-by (rule mult_1_left)
 
 lemma hypnat_diff_mult_distrib: "!!k m n. ((m::hypnat) - n) * k = (m * k) - (n * k)"
 by transfer (rule diff_mult_distrib)
@@ -227,87 +61,8 @@
 lemma hypnat_diff_mult_distrib2: "!!k m n. (k::hypnat) * (m - n) = (k * m) - (k * n)"
 by transfer (rule diff_mult_distrib2)
 
-lemma hypnat_add_mult_distrib: "((z1::hypnat) + z2) * w = (z1 * w) + (z2 * w)"
-by (rule left_distrib)
-
-lemma hypnat_add_mult_distrib2: "(w::hypnat) * (z1 + z2) = (w * z1) + (w * z2)"
-by (rule right_distrib)
-
-text{*one and zero are distinct*}
-lemma hypnat_zero_not_eq_one: "(0::hypnat) \<noteq> (1::hypnat)"
-by (rule zero_neq_one)
-declare hypnat_zero_not_eq_one [THEN not_sym, simp]
-
-(*
-text{*The hypernaturals form a @{text comm_semiring_1_cancel}: *}
-instance hypnat :: comm_semiring_1_cancel
-proof
-  fix i j k :: hypnat
-  show "(i * j) * k = i * (j * k)" by (rule hypnat_mult_assoc)
-  show "i * j = j * i" by (rule hypnat_mult_commute)
-  show "1 * i = i" by (rule hypnat_mult_1)
-  show "(i + j) * k = i * k + j * k" by (simp add: hypnat_add_mult_distrib)
-  show "0 \<noteq> (1::hypnat)" by (rule hypnat_zero_not_eq_one)
-  assume "k+i = k+j"
-  hence "(k+i) - k = (k+j) - k" by simp
-  thus "i=j" by simp
-qed
-*)
-
 subsection{*Properties of The @{text "\<le>"} Relation*}
 
-lemma hypnat_le:
-      "(Abs_star(starrel``{%n. X n}) \<le> Abs_star(starrel``{%n. Y n})) =
-       ({n. X n \<le> Y n} \<in> FreeUltrafilterNat)"
-by (rule hypreal_le)
-
-lemma hypnat_le_refl: "w \<le> (w::hypnat)"
-by (rule order_refl)
-
-lemma hypnat_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::hypnat)"
-by (rule order_trans)
-
-lemma hypnat_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::hypnat)"
-by (rule order_antisym)
-
-(* Axiom 'order_less_le' of class 'order': *)
-lemma hypnat_less_le: "((w::hypnat) < z) = (w \<le> z & w \<noteq> z)"
-by (rule order_less_le)
-
-(*
-instance hypnat :: order
-  by intro_classes
-    (assumption |
-      rule hypnat_le_refl hypnat_le_trans hypnat_le_anti_sym hypnat_less_le)+
-*)
-(* Axiom 'linorder_linear' of class 'linorder': *)
-lemma hypnat_le_linear: "(z::hypnat) \<le> w | w \<le> z"
-by (rule linorder_linear)
-(*
-instance hypnat :: linorder
-  by intro_classes (rule hypnat_le_linear)
-*)
-lemma hypnat_add_left_mono: "x \<le> y ==> z + x \<le> z + (y::hypnat)"
-by (rule add_left_mono)
-
-lemma hypnat_mult_less_mono2: "[| (0::hypnat)<z; x<y |] ==> z*x<z*y"
-by (rule mult_strict_left_mono)
-
-
-subsection{*The Hypernaturals Form an Ordered @{text comm_semiring_1_cancel} *}
-(*
-instance hypnat :: ordered_semidom
-proof
-  fix x y z :: hypnat
-  show "0 < (1::hypnat)"
-    by (simp add: hypnat_zero_def hypnat_one_def linorder_not_le [symmetric],
-        simp add: hypnat_le)
-  show "x \<le> y ==> z + x \<le> z + y"
-    by (rule hypnat_add_left_mono)
-  show "x < y ==> 0 < z ==> z * x < z * y"
-    by (simp add: hypnat_mult_less_mono2)
-qed
-*)
 lemma hypnat_le_zero_cancel [iff]: "!!n. (n \<le> (0::hypnat)) = (n = 0)"
 by transfer (rule le_0_eq)
 
@@ -321,11 +76,6 @@
 
 subsection{*Theorems for Ordering*}
 
-lemma hypnat_less:
-      "(Abs_star(starrel``{%n. X n}) < Abs_star(starrel``{%n. Y n})) =
-       ({n. X n < Y n} \<in> FreeUltrafilterNat)"
-by (rule hypreal_less)
-
 lemma hypnat_not_less0 [iff]: "!!n. ~ n < (0::hypnat)"
 by transfer (rule not_less0)
 
@@ -389,42 +139,42 @@
 
 constdefs
 
-  hypnat_of_nat   :: "nat => hypnat"
-  "hypnat_of_nat m  == of_nat m"
-
   (* the set of infinite hypernatural numbers *)
   HNatInfinite :: "hypnat set"
   "HNatInfinite == {n. n \<notin> Nats}"
 
 
+lemma hypnat_of_nat_def: "hypnat_of_nat m == of_nat m"
+by (transfer star_of_nat_def) simp
+
 lemma hypnat_of_nat_add:
       "hypnat_of_nat ((z::nat) + w) = hypnat_of_nat z + hypnat_of_nat w"
-by (simp add: hypnat_of_nat_def)
+by simp
 
 lemma hypnat_of_nat_mult:
       "hypnat_of_nat (z * w) = hypnat_of_nat z * hypnat_of_nat w"
-by (simp add: hypnat_of_nat_def)
+by simp
 
-lemma hypnat_of_nat_less_iff [simp]:
+lemma hypnat_of_nat_less_iff:
       "(hypnat_of_nat z < hypnat_of_nat w) = (z < w)"
-by (simp add: hypnat_of_nat_def)
+by simp
 
-lemma hypnat_of_nat_le_iff [simp]:
+lemma hypnat_of_nat_le_iff:
       "(hypnat_of_nat z \<le> hypnat_of_nat w) = (z \<le> w)"
-by (simp add: hypnat_of_nat_def)
+by simp
 
-lemma hypnat_of_nat_eq_iff [simp]:
+lemma hypnat_of_nat_eq_iff:
       "(hypnat_of_nat z = hypnat_of_nat w) = (z = w)"
-by (simp add: hypnat_of_nat_def)
+by simp
 
 lemma hypnat_of_nat_one [simp]: "hypnat_of_nat (Suc 0) = (1::hypnat)"
-by (simp add: hypnat_of_nat_def)
+by simp
 
-lemma hypnat_of_nat_zero [simp]: "hypnat_of_nat 0 = 0"
-by (simp add: hypnat_of_nat_def)
+lemma hypnat_of_nat_zero: "hypnat_of_nat 0 = 0"
+by simp
 
-lemma hypnat_of_nat_zero_iff [simp]: "(hypnat_of_nat n = 0) = (n = 0)"
-by (simp add: hypnat_of_nat_def)
+lemma hypnat_of_nat_zero_iff: "(hypnat_of_nat n = 0) = (n = 0)"
+by simp
 
 lemma hypnat_of_nat_Suc [simp]:
      "hypnat_of_nat (Suc n) = hypnat_of_nat n + (1::hypnat)"
@@ -432,17 +182,11 @@
 
 lemma hypnat_of_nat_minus:
       "hypnat_of_nat ((j::nat) - k) = hypnat_of_nat j - hypnat_of_nat k"
-by (simp add: hypnat_of_nat_def split: nat_diff_split hypnat_diff_split)
+by simp
 
 
 subsection{*Existence of an infinite hypernatural number*}
 
-lemma hypnat_omega: "starrel``{%n::nat. n} \<in> star"
-by auto
-
-lemma Rep_star_omega: "Rep_star(whn) \<in> star"
-by (simp add: hypnat_omega_def)
-
 text{*Existence of infinite number not corresponding to any natural number
 follows because member @{term FreeUltrafilterNat} is not finite.
 See @{text HyperDef.thy} for similar argument.*}
@@ -471,11 +215,10 @@
 lemma Compl_Collect_le: "- {n::nat. N \<le> n} = {n. n < N}"
 by (simp add: Collect_neg_eq [symmetric] linorder_not_le) 
 
-
 lemma hypnat_of_nat_eq:
-     "hypnat_of_nat m  = Abs_star(starrel``{%n::nat. m})"
+     "hypnat_of_nat m  = star_n (%n::nat. m)"
 apply (induct m) 
-apply (simp_all add: hypnat_zero_def hypnat_one_def hypnat_add) 
+apply (simp_all add: star_n_zero_num star_n_one_num star_n_add)
 done
 
 lemma SHNat_eq: "Nats = {n. \<exists>N. n = hypnat_of_nat N}"
@@ -483,7 +226,7 @@
 
 lemma hypnat_omega_gt_SHNat:
      "n \<in> Nats ==> n < whn"
-by (auto simp add: hypnat_of_nat_eq hypnat_less hypnat_omega_def SHNat_eq)
+by (auto simp add: hypnat_of_nat_eq star_n_less hypnat_omega_def SHNat_eq)
 
 (* Infinite hypernatural not in embedded Nats *)
 lemma SHNAT_omega_not_mem [simp]: "whn \<notin> Nats"
@@ -532,10 +275,10 @@
 
 lemma HNatInfinite_iff: "HNatInfinite = {N. \<forall>n \<in> Nats. n < N}"
 apply (auto simp add: HNatInfinite_def SHNat_eq hypnat_of_nat_eq)
-apply (rule_tac z = x in eq_Abs_star)
+apply (rule_tac x = x in star_cases)
 apply (auto elim: HNatInfinite_FreeUltrafilterNat_lemma 
-            simp add: hypnat_less FreeUltrafilterNat_Compl_iff1 
-                      Collect_neg_eq [symmetric])
+            simp add: star_n_less FreeUltrafilterNat_Compl_iff1 
+                      star_n_eq_iff Collect_neg_eq [symmetric])
 done
 
 
@@ -545,17 +288,17 @@
 lemma HNatInfinite_FreeUltrafilterNat:
      "x \<in> HNatInfinite 
       ==> \<exists>X \<in> Rep_star x. \<forall>u. {n. u < X n}:  FreeUltrafilterNat"
-apply (rule_tac z=x in eq_Abs_star)
+apply (cases x)
 apply (auto simp add: HNatInfinite_iff SHNat_eq hypnat_of_nat_eq)
-apply (rule bexI [OF _ lemma_starrel_refl], clarify) 
-apply (auto simp add: hypnat_of_nat_def hypnat_less)
+apply (rule bexI [OF _ Rep_star_star_n], clarify) 
+apply (auto simp add: hypnat_of_nat_def star_n_less)
 done
 
 lemma FreeUltrafilterNat_HNatInfinite:
      "\<exists>X \<in> Rep_star x. \<forall>u. {n. u < X n}:  FreeUltrafilterNat
       ==> x \<in> HNatInfinite"
-apply (rule_tac z=x in eq_Abs_star)
-apply (auto simp add: hypnat_less HNatInfinite_iff SHNat_eq hypnat_of_nat_eq)
+apply (cases x)
+apply (auto simp add: star_n_less HNatInfinite_iff SHNat_eq hypnat_of_nat_eq)
 apply (drule spec, ultra, auto) 
 done
 
@@ -628,72 +371,51 @@
 
 constdefs
   hypreal_of_hypnat :: "hypnat => hypreal"
-   "hypreal_of_hypnat N  == 
-      Abs_star(\<Union>X \<in> Rep_star(N). starrel``{%n::nat. real (X n)})"
+   "hypreal_of_hypnat == *f* real"
 
 
 lemma HNat_hypreal_of_nat [simp]: "hypreal_of_nat N \<in> Nats"
 by (simp add: hypreal_of_nat_def) 
 
-(*WARNING: FRAGILE!*)
-lemma lemma_starrel_FUFN:
-     "(Ya \<in> starrel ``{%n. f(n)}) = ({n. f n = Ya n} \<in> FreeUltrafilterNat)"
-by force
-
 lemma hypreal_of_hypnat:
-      "hypreal_of_hypnat (Abs_star(starrel``{%n. X n})) =
-       Abs_star(starrel `` {%n. real (X n)})"
-apply (simp add: hypreal_of_hypnat_def)
-apply (rule_tac f = Abs_star in arg_cong)
-apply (auto elim: FreeUltrafilterNat_Int [THEN FreeUltrafilterNat_subset] 
-       simp add: lemma_starrel_FUFN)
-done
+      "hypreal_of_hypnat (star_n X) = star_n (%n. real (X n))"
+by (simp add: hypreal_of_hypnat_def starfun)
 
 lemma hypreal_of_hypnat_inject [simp]:
-     "(hypreal_of_hypnat m = hypreal_of_hypnat n) = (m=n)"
-apply (rule_tac z=m in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (auto simp add: hypreal_of_hypnat)
-done
+     "!!m n. (hypreal_of_hypnat m = hypreal_of_hypnat n) = (m=n)"
+by (unfold hypreal_of_hypnat_def, transfer, simp)
 
 lemma hypreal_of_hypnat_zero: "hypreal_of_hypnat 0 = 0"
-by (simp add: hypnat_zero_def hypreal_zero_def hypreal_of_hypnat)
+by (simp add: star_n_zero_num hypreal_of_hypnat)
 
 lemma hypreal_of_hypnat_one: "hypreal_of_hypnat (1::hypnat) = 1"
-by (simp add: hypnat_one_def hypreal_one_def hypreal_of_hypnat)
+by (simp add: star_n_one_num hypreal_of_hypnat)
 
 lemma hypreal_of_hypnat_add [simp]:
-     "hypreal_of_hypnat (m + n) = hypreal_of_hypnat m + hypreal_of_hypnat n"
-apply (rule_tac z=m in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (simp add: hypreal_of_hypnat hypreal_add hypnat_add real_of_nat_add)
-done
+     "!!m n. hypreal_of_hypnat (m + n) = hypreal_of_hypnat m + hypreal_of_hypnat n"
+by (unfold hypreal_of_hypnat_def, transfer, rule real_of_nat_add)
 
 lemma hypreal_of_hypnat_mult [simp]:
-     "hypreal_of_hypnat (m * n) = hypreal_of_hypnat m * hypreal_of_hypnat n"
-apply (rule_tac z=m in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (simp add: hypreal_of_hypnat hypreal_mult hypnat_mult real_of_nat_mult)
-done
+     "!!m n. hypreal_of_hypnat (m * n) = hypreal_of_hypnat m * hypreal_of_hypnat n"
+by (unfold hypreal_of_hypnat_def, transfer, rule real_of_nat_mult)
 
 lemma hypreal_of_hypnat_less_iff [simp]:
-     "(hypreal_of_hypnat n < hypreal_of_hypnat m) = (n < m)"
-apply (rule_tac z=m in eq_Abs_star, rule_tac z=n in eq_Abs_star)
-apply (simp add: hypreal_of_hypnat hypreal_less hypnat_less)
-done
+     "!!m n. (hypreal_of_hypnat n < hypreal_of_hypnat m) = (n < m)"
+by (unfold hypreal_of_hypnat_def, transfer, simp)
 
 lemma hypreal_of_hypnat_eq_zero_iff: "(hypreal_of_hypnat N = 0) = (N = 0)"
 by (simp add: hypreal_of_hypnat_zero [symmetric])
 declare hypreal_of_hypnat_eq_zero_iff [simp]
 
-lemma hypreal_of_hypnat_ge_zero [simp]: "0 \<le> hypreal_of_hypnat n"
-apply (rule_tac z=n in eq_Abs_star)
-apply (simp add: hypreal_of_hypnat hypreal_zero_num hypreal_le)
-done
+lemma hypreal_of_hypnat_ge_zero [simp]: "!!n. 0 \<le> hypreal_of_hypnat n"
+by (unfold hypreal_of_hypnat_def, transfer, simp)
 
 lemma HNatInfinite_inverse_Infinitesimal [simp]:
      "n \<in> HNatInfinite ==> inverse (hypreal_of_hypnat n) \<in> Infinitesimal"
-apply (rule_tac z=n in eq_Abs_star)
-apply (auto simp add: hypreal_of_hypnat hypreal_inverse 
+apply (cases n)
+apply (auto simp add: hypreal_of_hypnat star_n_inverse 
       HNatInfinite_FreeUltrafilterNat_iff Infinitesimal_FreeUltrafilterNat_iff2)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl, auto)
+apply (rule bexI [OF _ Rep_star_star_n], auto)
 apply (drule_tac x = "m + 1" in spec, ultra)
 done
 
@@ -709,22 +431,11 @@
 val hypnat_of_nat_def = thm"hypnat_of_nat_def";
 val HNatInfinite_def = thm"HNatInfinite_def";
 val hypreal_of_hypnat_def = thm"hypreal_of_hypnat_def";
-val hypnat_zero_def = thm"hypnat_zero_def";
-val hypnat_one_def = thm"hypnat_one_def";
 val hypnat_omega_def = thm"hypnat_omega_def";
 
 val starrel_iff = thm "starrel_iff";
-(* val starrel_in_hypnat = thm "starrel_in_hypnat"; *)
 val lemma_starrel_refl = thm "lemma_starrel_refl";
-(* val hypnat_empty_not_mem = thm "hypnat_empty_not_mem"; *)
-(* val Rep_star_nonempty = thm "Rep_star_nonempty"; *)
 val eq_Abs_star = thm "eq_Abs_star";
-val hypnat_add = thm "hypnat_add";
-val hypnat_add_commute = thm "hypnat_add_commute";
-val hypnat_add_assoc = thm "hypnat_add_assoc";
-val hypnat_add_zero_left = thm "hypnat_add_zero_left";
-(* val hypnat_minus_congruent2 = thm "hypnat_minus_congruent2"; *)
-val hypnat_minus = thm "hypnat_minus";
 val hypnat_minus_zero = thm "hypnat_minus_zero";
 val hypnat_diff_0_eq_0 = thm "hypnat_diff_0_eq_0";
 val hypnat_add_is_0 = thm "hypnat_add_is_0";
@@ -735,26 +446,9 @@
 val hypnat_diff_cancel = thm "hypnat_diff_cancel";
 val hypnat_diff_cancel2 = thm "hypnat_diff_cancel2";
 val hypnat_diff_add_0 = thm "hypnat_diff_add_0";
-(* val hypnat_mult_congruent2 = thm "hypnat_mult_congruent2"; *)
-val hypnat_mult = thm "hypnat_mult";
-val hypnat_mult_commute = thm "hypnat_mult_commute";
-val hypnat_mult_assoc = thm "hypnat_mult_assoc";
-val hypnat_mult_1 = thm "hypnat_mult_1";
 val hypnat_diff_mult_distrib = thm "hypnat_diff_mult_distrib";
 val hypnat_diff_mult_distrib2 = thm "hypnat_diff_mult_distrib2";
-val hypnat_add_mult_distrib = thm "hypnat_add_mult_distrib";
-val hypnat_add_mult_distrib2 = thm "hypnat_add_mult_distrib2";
-val hypnat_zero_not_eq_one = thm "hypnat_zero_not_eq_one";
-val hypnat_le = thm "hypnat_le";
-val hypnat_le_refl = thm "hypnat_le_refl";
-val hypnat_le_trans = thm "hypnat_le_trans";
-val hypnat_le_anti_sym = thm "hypnat_le_anti_sym";
-val hypnat_less_le = thm "hypnat_less_le";
-val hypnat_le_linear = thm "hypnat_le_linear";
-val hypnat_add_left_mono = thm "hypnat_add_left_mono";
-val hypnat_mult_less_mono2 = thm "hypnat_mult_less_mono2";
 val hypnat_mult_is_0 = thm "hypnat_mult_is_0";
-val hypnat_less = thm "hypnat_less";
 val hypnat_not_less0 = thm "hypnat_not_less0";
 val hypnat_less_one = thm "hypnat_less_one";
 val hypnat_add_diff_inverse = thm "hypnat_add_diff_inverse";
@@ -777,8 +471,6 @@
 val hypnat_of_nat_zero = thm "hypnat_of_nat_zero";
 val hypnat_of_nat_zero_iff = thm "hypnat_of_nat_zero_iff";
 val hypnat_of_nat_Suc = thm "hypnat_of_nat_Suc";
-val hypnat_omega = thm "hypnat_omega";
-val Rep_star_omega = thm "Rep_star_omega";
 val SHNAT_omega_not_mem = thm "SHNAT_omega_not_mem";
 val cofinite_mem_FreeUltrafilterNat = thm "cofinite_mem_FreeUltrafilterNat";
 val hypnat_omega_gt_SHNat = thm "hypnat_omega_gt_SHNat";
--- a/src/HOL/Hyperreal/HyperPow.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/HyperPow.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -10,22 +10,7 @@
 imports HyperArith HyperNat
 begin
 
-(* instance hypreal :: power .. *)
-
-consts hpowr :: "[hypreal,nat] => hypreal"  
-(*
-primrec
-   hpowr_0:   "r ^ 0       = (1::hypreal)"
-   hpowr_Suc: "r ^ (Suc n) = (r::hypreal) * (r ^ n)"
-
-instance hypreal :: recpower
-proof
-  fix z :: hypreal
-  fix n :: nat
-  show "z^0 = 1" by simp
-  show "z^(Suc n) = z * (z^n)" by simp
-qed
-*)
+(* consts hpowr :: "[hypreal,nat] => hypreal" *)
 
 lemma hpowr_0 [simp]:   "r ^ 0       = (1::hypreal)"
 by (rule power_0)
@@ -34,15 +19,13 @@
 by (rule power_Suc)
 
 consts
-  "pow"  :: "[hypreal,hypnat] => hypreal"     (infixr 80)
+  "pow"  :: "[hypreal,hypnat] => hypreal"     (infixr "pow" 80)
 
 defs
 
   (* hypernatural powers of hyperreals *)
   hyperpow_def:
-  "(R::hypreal) pow (N::hypnat) ==
-      Abs_star(\<Union>X \<in> Rep_star(R). \<Union>Y \<in> Rep_star(N).
-                        starrel``{%n::nat. (X n) ^ (Y n)})"
+  "(R::hypreal) pow (N::hypnat) == Ifun2_of (op ^) R N"
 
 lemma hrealpow_two: "(r::hypreal) ^ Suc (Suc 0) = r * r"
 by simp
@@ -52,11 +35,11 @@
 
 lemma hrealpow_two_le_add_order [simp]:
      "(0::hypreal) \<le> u ^ Suc (Suc 0) + v ^ Suc (Suc 0)"
-by (simp only: hrealpow_two_le hypreal_le_add_order)
+by (simp only: hrealpow_two_le add_nonneg_nonneg)
 
 lemma hrealpow_two_le_add_order2 [simp]:
      "(0::hypreal) \<le> u ^ Suc (Suc 0) + v ^ Suc (Suc 0) + w ^ Suc (Suc 0)"
-by (simp only: hrealpow_two_le hypreal_le_add_order)
+by (simp only: hrealpow_two_le add_nonneg_nonneg)
 
 lemma hypreal_add_nonneg_eq_0_iff:
      "[| 0 \<le> x; 0 \<le> y |] ==> (x+y = 0) = (x = 0 & y = (0::hypreal))"
@@ -66,7 +49,7 @@
 text{*FIXME: DELETE THESE*}
 lemma hypreal_three_squares_add_zero_iff:
      "(x*x + y*y + z*z = 0) = (x = 0 & y = 0 & z = (0::hypreal))"
-apply (simp only: zero_le_square hypreal_le_add_order hypreal_add_nonneg_eq_0_iff, auto)
+apply (simp only: zero_le_square add_nonneg_nonneg hypreal_add_nonneg_eq_0_iff, auto)
 done
 
 lemma hrealpow_three_squares_add_zero_iff [simp]:
@@ -90,9 +73,9 @@
 done
 
 lemma hrealpow:
-    "Abs_star(starrel``{%n. X n}) ^ m = Abs_star(starrel``{%n. (X n::real) ^ m})"
+    "star_n X ^ m = star_n (%n. (X n::real) ^ m)"
 apply (induct_tac "m")
-apply (auto simp add: hypreal_one_def hypreal_mult power_0)
+apply (auto simp add: star_n_one_num star_n_mult power_0)
 done
 
 lemma hrealpow_sum_square_expand:
@@ -103,14 +86,10 @@
 
 subsection{*Literal Arithmetic Involving Powers and Type @{typ hypreal}*}
 
-lemma hypreal_of_real_power [simp]:
-     "hypreal_of_real (x ^ n) = hypreal_of_real x ^ n"
-by (induct_tac "n", simp_all add: nat_mult_distrib)
-
 lemma power_hypreal_of_real_number_of:
      "(number_of v :: hypreal) ^ n = hypreal_of_real ((number_of v) ^ n)"
-by (simp only: hypreal_number_of [symmetric] hypreal_of_real_power)
-
+by simp
+(* why is this a simp rule? - BH *)
 declare power_hypreal_of_real_number_of [of _ "number_of w", standard, simp]
 
 lemma hrealpow_HFinite: "x \<in> HFinite ==> x ^ n \<in> HFinite"
@@ -121,98 +100,50 @@
 
 subsection{*Powers with Hypernatural Exponents*}
 
-lemma hyperpow_congruent: "(%X Y. starrel``{%n. (X n ^ Y n)}) respects starrel"
-by (auto simp add: congruent_def intro!: ext, fuf+)
-
-lemma hyperpow:
-  "Abs_star(starrel``{%n. X n}) pow Abs_star(starrel``{%n. Y n}) =
-   Abs_star(starrel``{%n. X n ^ Y n})"
-apply (unfold hyperpow_def)
-apply (rule_tac f = Abs_star in arg_cong)
-apply (auto intro!: lemma_starrel_refl bexI 
-           simp add: starrel_in_hypreal [THEN Abs_star_inverse] equiv_starrel 
-                     hyperpow_congruent, fuf)
-done
+lemma hyperpow: "star_n X pow star_n Y = star_n (%n. X n ^ Y n)"
+by (simp add: hyperpow_def Ifun2_of_def star_of_def Ifun_star_n)
 
-lemma hyperpow_zero: "(0::hypreal) pow (n + (1::hypnat)) = 0"
-apply (unfold hypnat_one_def)
-apply (simp (no_asm) add: hypreal_zero_def)
-apply (rule_tac z = n in eq_Abs_star)
-apply (auto simp add: hyperpow hypnat_add)
-done
-declare hyperpow_zero [simp]
+lemma hyperpow_zero [simp]: "!!n. (0::hypreal) pow (n + (1::hypnat)) = 0"
+by (unfold hyperpow_def, transfer, simp)
 
-lemma hyperpow_not_zero [rule_format (no_asm)]:
-     "r \<noteq> (0::hypreal) --> r pow n \<noteq> 0"
-apply (simp (no_asm) add: hypreal_zero_def, rule_tac z=n in eq_Abs_star, rule_tac z=r in eq_Abs_star)
-apply (auto simp add: hyperpow)
-apply (drule FreeUltrafilterNat_Compl_mem, ultra)
-done
+lemma hyperpow_not_zero: "!!r n. r \<noteq> (0::hypreal) ==> r pow n \<noteq> 0"
+by (unfold hyperpow_def, transfer, simp)
 
 lemma hyperpow_inverse:
-     "r \<noteq> (0::hypreal) --> inverse(r pow n) = (inverse r) pow n"
-apply (simp (no_asm) add: hypreal_zero_def, rule_tac z=n in eq_Abs_star, rule_tac z=r in eq_Abs_star)
-apply (auto dest!: FreeUltrafilterNat_Compl_mem simp add: hypreal_inverse hyperpow)
-apply (rule FreeUltrafilterNat_subset)
-apply (auto dest: realpow_not_zero intro: power_inverse)
-done
+     "!!r n. r \<noteq> (0::hypreal) ==> inverse(r pow n) = (inverse r) pow n"
+by (unfold hyperpow_def, transfer, rule power_inverse)
+
+lemma hyperpow_hrabs: "!!r n. abs r pow n = abs (r pow n)"
+by (unfold hyperpow_def, transfer, rule power_abs [symmetric])
 
-lemma hyperpow_hrabs: "abs r pow n = abs (r pow n)"
-apply (rule_tac z=n in eq_Abs_star, rule_tac z=r in eq_Abs_star)
-apply (auto simp add: hypreal_hrabs hyperpow power_abs [symmetric])
-done
+lemma hyperpow_add: "!!r n m. r pow (n + m) = (r pow n) * (r pow m)"
+by (unfold hyperpow_def, transfer, rule power_add)
 
-lemma hyperpow_add: "r pow (n + m) = (r pow n) * (r pow m)"
-apply (rule_tac z=n in eq_Abs_star, rule_tac z=m in eq_Abs_star, rule_tac z=r in eq_Abs_star)
-apply (auto simp add: hyperpow hypnat_add hypreal_mult power_add)
-done
-
-lemma hyperpow_one [simp]: "r pow (1::hypnat) = r"
-apply (unfold hypnat_one_def, rule_tac z=r in eq_Abs_star)
-apply (auto simp add: hyperpow)
-done
+lemma hyperpow_one [simp]: "!!r. r pow (1::hypnat) = r"
+by (unfold hyperpow_def, transfer, simp)
 
 lemma hyperpow_two:
-     "r pow ((1::hypnat) + (1::hypnat)) = r * r"
-apply (unfold hypnat_one_def, rule_tac z=r in eq_Abs_star)
-apply (auto simp add: hyperpow hypnat_add hypreal_mult)
-done
+     "!!r. r pow ((1::hypnat) + (1::hypnat)) = r * r"
+by (unfold hyperpow_def, transfer, simp)
 
-lemma hyperpow_gt_zero: "(0::hypreal) < r ==> 0 < r pow n"
-apply (simp add: hypreal_zero_def, rule_tac z=n in eq_Abs_star, rule_tac z=r in eq_Abs_star)
-apply (auto elim!: FreeUltrafilterNat_subset zero_less_power
-                   simp add: hyperpow hypreal_less hypreal_le)
-done
+lemma hyperpow_gt_zero: "!!r n. (0::hypreal) < r ==> 0 < r pow n"
+by (unfold hyperpow_def, transfer, rule zero_less_power)
 
-lemma hyperpow_ge_zero: "(0::hypreal) \<le> r ==> 0 \<le> r pow n"
-apply (simp add: hypreal_zero_def, rule_tac z=n in eq_Abs_star, rule_tac z=r in eq_Abs_star)
-apply (auto elim!: FreeUltrafilterNat_subset zero_le_power 
-            simp add: hyperpow hypreal_le)
-done
+lemma hyperpow_ge_zero: "!!r n. (0::hypreal) \<le> r ==> 0 \<le> r pow n"
+by (unfold hyperpow_def, transfer, rule zero_le_power)
 
-lemma hyperpow_le: "[|(0::hypreal) < x; x \<le> y|] ==> x pow n \<le> y pow n"
-apply (simp add: hypreal_zero_def, rule_tac z=n in eq_Abs_star, rule_tac z=x in eq_Abs_star, rule_tac z=y in eq_Abs_star)
-apply (auto simp add: hyperpow hypreal_le hypreal_less)
-apply (erule FreeUltrafilterNat_Int [THEN FreeUltrafilterNat_subset], assumption)
-apply (auto intro: power_mono)
-done
-
-lemma hyperpow_eq_one [simp]: "1 pow n = (1::hypreal)"
-apply (rule_tac z=n in eq_Abs_star)
-apply (auto simp add: hypreal_one_def hyperpow)
-done
+lemma hyperpow_le:
+  "!!x y n. [|(0::hypreal) < x; x \<le> y|] ==> x pow n \<le> y pow n"
+by (unfold hyperpow_def, transfer, rule power_mono [OF _ order_less_imp_le])
 
-lemma hrabs_hyperpow_minus_one [simp]: "abs(-1 pow n) = (1::hypreal)"
-apply (subgoal_tac "abs ((- (1::hypreal)) pow n) = (1::hypreal) ")
-apply simp
-apply (rule_tac z=n in eq_Abs_star)
-apply (auto simp add: hypreal_one_def hyperpow hypreal_minus hypreal_hrabs)
-done
+lemma hyperpow_eq_one [simp]: "!!n. 1 pow n = (1::hypreal)"
+by (unfold hyperpow_def, transfer, simp)
 
-lemma hyperpow_mult: "(r * s) pow n = (r pow n) * (s pow n)"
-apply (rule_tac z=n in eq_Abs_star, rule_tac z=r in eq_Abs_star, rule_tac z=s in eq_Abs_star)
-apply (auto simp add: hyperpow hypreal_mult power_mult_distrib)
-done
+lemma hrabs_hyperpow_minus_one [simp]: "!!n. abs(-1 pow n) = (1::hypreal)"
+by (unfold hyperpow_def, transfer, simp)
+
+lemma hyperpow_mult: "!!r s n. (r * s) pow n = (r pow n) * (s pow n)"
+by (unfold hyperpow_def, transfer, rule power_mult_distrib)
 
 lemma hyperpow_two_le [simp]: "0 \<le> r pow (1 + 1)"
 by (auto simp add: hyperpow_two zero_le_mult_iff)
@@ -244,23 +175,12 @@
 done
 
 lemma hyperpow_minus_one2 [simp]:
-     "-1 pow ((1 + 1)*n) = (1::hypreal)"
-apply (subgoal_tac " (- ((1::hypreal))) pow ((1 + 1)*n) = (1::hypreal) ")
-apply simp
-apply (simp only: hypreal_one_def, rule_tac z=n in eq_Abs_star)
-apply (auto simp add: nat_mult_2 [symmetric] hyperpow hypnat_add hypreal_minus
-                      hypnat_mult left_distrib)
-done
+     "!!n. -1 pow ((1 + 1)*n) = (1::hypreal)"
+by (unfold hyperpow_def, transfer, simp)
 
 lemma hyperpow_less_le:
-     "[|(0::hypreal) \<le> r; r \<le> 1; n < N|] ==> r pow N \<le> r pow n"
-apply (rule_tac z=n in eq_Abs_star, rule_tac z=N in eq_Abs_star, rule_tac z=r in eq_Abs_star)
-apply (auto simp add: hyperpow hypreal_le hypreal_less hypnat_less 
-            hypreal_zero_def hypreal_one_def)
-apply (erule FreeUltrafilterNat_Int [THEN FreeUltrafilterNat_subset])
-apply (erule FreeUltrafilterNat_Int, assumption)
-apply (auto intro: power_decreasing)
-done
+     "!!r n N. [|(0::hypreal) \<le> r; r \<le> 1; n < N|] ==> r pow N \<le> r pow n"
+by (unfold hyperpow_def, transfer, rule power_decreasing [OF order_less_imp_le])
 
 lemma hyperpow_SHNat_le:
      "[| 0 \<le> r;  r \<le> (1::hypreal);  N \<in> HNatInfinite |]
@@ -269,11 +189,11 @@
 
 lemma hyperpow_realpow:
       "(hypreal_of_real r) pow (hypnat_of_nat n) = hypreal_of_real (r ^ n)"
-by (simp add: hypreal_of_real_def star_of_def star_n_def hypnat_of_nat_eq hyperpow)
+by (simp add: star_of_def hypnat_of_nat_eq hyperpow)
 
 lemma hyperpow_SReal [simp]:
      "(hypreal_of_real r) pow (hypnat_of_nat n) \<in> Reals"
-by (simp del: hypreal_of_real_power add: hyperpow_realpow SReal_def)
+by (simp del: star_of_power add: hyperpow_realpow SReal_def)
 
 
 lemma hyperpow_zero_HNatInfinite [simp]:
@@ -307,16 +227,13 @@
 
 lemma hrealpow_hyperpow_Infinitesimal_iff:
      "(x ^ n \<in> Infinitesimal) = (x pow (hypnat_of_nat n) \<in> Infinitesimal)"
-apply (rule_tac z=x in eq_Abs_star)
+apply (cases x)
 apply (simp add: hrealpow hyperpow hypnat_of_nat_eq)
 done
 
 lemma Infinitesimal_hrealpow:
      "[| x \<in> Infinitesimal; 0 < n |] ==> x ^ n \<in> Infinitesimal"
-by (force intro!: Infinitesimal_hyperpow
-          simp add: hrealpow_hyperpow_Infinitesimal_iff 
-                    hypnat_of_nat_less_iff [symmetric] hypnat_of_nat_zero
-          simp del: hypnat_of_nat_less_iff)
+by (simp add: hrealpow_hyperpow_Infinitesimal_iff Infinitesimal_hyperpow)
 
 ML
 {*
@@ -332,10 +249,8 @@
 val two_hrealpow_gt = thm "two_hrealpow_gt";
 val hrealpow = thm "hrealpow";
 val hrealpow_sum_square_expand = thm "hrealpow_sum_square_expand";
-val hypreal_of_real_power = thm "hypreal_of_real_power";
 val power_hypreal_of_real_number_of = thm "power_hypreal_of_real_number_of";
 val hrealpow_HFinite = thm "hrealpow_HFinite";
-val hyperpow_congruent = thm "hyperpow_congruent";
 val hyperpow = thm "hyperpow";
 val hyperpow_zero = thm "hyperpow_zero";
 val hyperpow_not_zero = thm "hyperpow_not_zero";
--- a/src/HOL/Hyperreal/Integration.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/Integration.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -322,7 +322,8 @@
 apply arith
 apply (drule add_strict_mono, assumption)
 apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] 
-                mult_less_cancel_right, arith)
+                mult_less_cancel_right)
+apply arith
 done
 
 lemma Integral_zero [simp]: "Integral(a,a) f 0"
--- a/src/HOL/Hyperreal/Lim.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/Lim.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -226,12 +226,12 @@
       "f -- x --> L ==> f -- x --NS> L"
 apply (simp add: LIM_def NSLIM_def approx_def)
 apply (simp add: Infinitesimal_FreeUltrafilterNat_iff, safe)
-apply (rule_tac z = xa in eq_Abs_star)
-apply (auto simp add: real_add_minus_iff starfun hypreal_minus hypreal_of_real_def star_of_def star_n_def hypreal_add)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl, clarify) 
+apply (rule_tac x = xa in star_cases)
+apply (auto simp add: real_add_minus_iff starfun star_n_minus star_of_def star_n_add star_n_eq_iff)
+apply (rule bexI [OF _ Rep_star_star_n], clarify) 
 apply (drule_tac x = u in spec, clarify)
 apply (drule_tac x = s in spec, clarify)
-apply (subgoal_tac "\<forall>n::nat. (xa n) \<noteq> x & \<bar>(xa n) + - x\<bar> < s --> \<bar>f (xa n) + - L\<bar> < u")
+apply (subgoal_tac "\<forall>n::nat. (Xa n) \<noteq> x & \<bar>(Xa n) + - x\<bar> < s --> \<bar>f (Xa n) + - L\<bar> < u")
 prefer 2 apply blast
 apply (drule FreeUltrafilterNat_all, ultra)
 done
@@ -270,10 +270,10 @@
 apply (rule ccontr, simp)
 apply (simp add: linorder_not_less)
 apply (drule lemma_skolemize_LIM2, safe)
-apply (drule_tac x = "Abs_star (starrel``{X})" in spec)
-apply (auto simp add: starfun hypreal_minus hypreal_of_real_def star_of_def star_n_def hypreal_add)
+apply (drule_tac x = "star_n X" in spec)
+apply (auto simp add: starfun star_n_minus star_of_def star_n_add star_n_eq_iff)
 apply (drule lemma_simp [THEN real_seq_to_hypreal_Infinitesimal])
-apply (simp add: Infinitesimal_FreeUltrafilterNat_iff hypreal_of_real_def star_of_def star_n_def hypreal_minus hypreal_add, blast)
+apply (simp add: Infinitesimal_FreeUltrafilterNat_iff star_of_def star_n_minus star_n_add star_n_eq_iff, blast)
 apply (drule spec, drule mp, assumption)
 apply (drule FreeUltrafilterNat_all, ultra)
 done
@@ -451,9 +451,9 @@
 apply (rule mem_infmal_iff [THEN iffD2, THEN Infinitesimal_add_approx_self [THEN approx_sym]])
 apply (rule_tac [4] approx_minus_iff2 [THEN iffD1])
  prefer 3 apply (simp add: add_commute)
-apply (rule_tac [2] z = x in eq_Abs_star)
-apply (rule_tac [4] z = x in eq_Abs_star)
-apply (auto simp add: starfun hypreal_of_real_def star_of_def star_n_def hypreal_minus hypreal_add add_assoc approx_refl hypreal_zero_def)
+apply (rule_tac [2] x = x in star_cases)
+apply (rule_tac [4] x = x in star_cases)
+apply (auto simp add: starfun star_of_def star_n_minus star_n_add add_assoc approx_refl star_n_zero_num)
 done
 
 lemma NSLIM_isCont_iff: "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)"
@@ -579,13 +579,13 @@
 lemma isUCont_isNSUCont: "isUCont f ==> isNSUCont f"
 apply (simp add: isNSUCont_def isUCont_def approx_def)
 apply (simp add: Infinitesimal_FreeUltrafilterNat_iff, safe)
-apply (rule_tac z = x in eq_Abs_star)
-apply (rule_tac z = y in eq_Abs_star)
-apply (auto simp add: starfun hypreal_minus hypreal_add)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl, safe)
+apply (rule_tac x = x in star_cases)
+apply (rule_tac x = y in star_cases)
+apply (auto simp add: starfun star_n_minus star_n_add)
+apply (rule bexI [OF _ Rep_star_star_n], safe)
 apply (drule_tac x = u in spec, clarify)
 apply (drule_tac x = s in spec, clarify)
-apply (subgoal_tac "\<forall>n::nat. abs ((xa n) + - (xb n)) < s --> abs (f (xa n) + - f (xb n)) < u")
+apply (subgoal_tac "\<forall>n::nat. abs ((Xa n) + - (Xb n)) < s --> abs (f (Xa n) + - f (Xb n)) < u")
 prefer 2 apply blast
 apply (erule_tac V = "\<forall>x y. \<bar>x + - y\<bar> < s --> \<bar>f x + - f y\<bar> < u" in thin_rl)
 apply (drule FreeUltrafilterNat_all, ultra)
@@ -620,11 +620,11 @@
 apply (rule ccontr, simp)
 apply (simp add: linorder_not_less)
 apply (drule lemma_skolemize_LIM2u, safe)
-apply (drule_tac x = "Abs_star (starrel``{X})" in spec)
-apply (drule_tac x = "Abs_star (starrel``{Y})" in spec)
-apply (simp add: starfun hypreal_minus hypreal_add, auto)
+apply (drule_tac x = "star_n X" in spec)
+apply (drule_tac x = "star_n Y" in spec)
+apply (simp add: starfun star_n_minus star_n_add, auto)
 apply (drule lemma_simpu [THEN real_seq_to_hypreal_Infinitesimal2])
-apply (simp add: Infinitesimal_FreeUltrafilterNat_iff hypreal_minus hypreal_add, blast)
+apply (simp add: Infinitesimal_FreeUltrafilterNat_iff star_n_minus star_n_add, blast)
 apply (drule_tac x = r in spec, clarify)
 apply (drule FreeUltrafilterNat_all, ultra)
 done
--- a/src/HOL/Hyperreal/NSA.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/NSA.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -8,7 +8,7 @@
 header{*Infinite Numbers, Infinitesimals, Infinitely Close Relation*}
 
 theory NSA
-imports HyperArith RComplete
+imports HyperArith "../Real/RComplete"
 begin
 
 constdefs
@@ -66,16 +66,15 @@
 
 lemma SReal_inverse: "(x::hypreal) \<in> Reals ==> inverse x \<in> Reals"
 apply (simp add: SReal_def)
-apply (blast intro: hypreal_of_real_inverse [symmetric])
+apply (blast intro: star_of_inverse [symmetric])
 done
 
 lemma SReal_divide: "[| (x::hypreal) \<in> Reals;  y \<in> Reals |] ==> x/y \<in> Reals"
-apply (simp (no_asm_simp) add: SReal_mult SReal_inverse hypreal_divide_def)
-done
+by (simp (no_asm_simp) add: SReal_mult SReal_inverse divide_inverse)
 
 lemma SReal_minus: "(x::hypreal) \<in> Reals ==> -x \<in> Reals"
 apply (simp add: SReal_def)
-apply (blast intro: hypreal_of_real_minus [symmetric])
+apply (blast intro: star_of_minus [symmetric])
 done
 
 lemma SReal_minus_iff [simp]: "(-x \<in> Reals) = ((x::hypreal) \<in> Reals)"
@@ -91,15 +90,16 @@
 done
 
 lemma SReal_hrabs: "(x::hypreal) \<in> Reals ==> abs x \<in> Reals"
-apply (simp add: SReal_def)
-apply (auto simp add: hypreal_of_real_hrabs)
+apply (auto simp add: SReal_def)
+apply (rule_tac x="abs r" in exI)
+apply simp
 done
 
 lemma SReal_hypreal_of_real [simp]: "hypreal_of_real x \<in> Reals"
 by (simp add: SReal_def)
 
 lemma SReal_number_of [simp]: "(number_of w ::hypreal) \<in> Reals"
-apply (simp only: hypreal_number_of [symmetric])
+apply (simp only: star_of_number_of [symmetric])
 apply (rule SReal_hypreal_of_real)
 done
 
@@ -116,7 +116,7 @@
 done
 
 lemma SReal_divide_number_of: "r \<in> Reals ==> r/(number_of w::hypreal) \<in> Reals"
-apply (unfold hypreal_divide_def)
+apply (simp only: divide_inverse)
 apply (blast intro!: SReal_number_of SReal_mult SReal_inverse)
 done
 
@@ -247,8 +247,8 @@
 lemma SReal_subset_HFinite: "Reals \<subseteq> HFinite"
 apply (auto simp add: SReal_def HFinite_def)
 apply (rule_tac x = "1 + abs (hypreal_of_real r) " in exI)
-apply (auto simp add: hypreal_of_real_hrabs)
-apply (rule_tac x = "1 + abs r" in exI, simp)
+apply (rule conjI, rule_tac x = "1 + abs r" in exI)
+apply simp_all
 done
 
 lemma HFinite_hypreal_of_real [simp]: "hypreal_of_real x \<in> HFinite"
@@ -278,7 +278,7 @@
 lemma HFinite_bounded: "[|x \<in> HFinite; y \<le> x; 0 \<le> y |] ==> y \<in> HFinite"
 apply (case_tac "x \<le> 0")
 apply (drule_tac y = x in order_trans)
-apply (drule_tac [2] hypreal_le_anti_sym)
+apply (drule_tac [2] order_antisym)
 apply (auto simp add: linorder_not_le)
 apply (auto intro: order_le_less_trans simp add: abs_if HFinite_def)
 done
@@ -309,7 +309,7 @@
 
 lemma Infinitesimal_diff:
      "[| x \<in> Infinitesimal;  y \<in> Infinitesimal |] ==> x-y \<in> Infinitesimal"
-by (simp add: hypreal_diff_def Infinitesimal_add)
+by (simp add: diff_def Infinitesimal_add)
 
 lemma Infinitesimal_mult:
      "[| x \<in> Infinitesimal; y \<in> Infinitesimal |] ==> (x * y) \<in> Infinitesimal"
@@ -331,7 +331,7 @@
 
 lemma Infinitesimal_HFinite_mult2:
      "[| x \<in> Infinitesimal; y \<in> HFinite |] ==> (y * x) \<in> Infinitesimal"
-by (auto dest: Infinitesimal_HFinite_mult simp add: hypreal_mult_commute)
+by (auto dest: Infinitesimal_HFinite_mult simp add: mult_commute)
 
 (*** rather long proof ***)
 lemma HInfinite_inverse_Infinitesimal:
@@ -359,11 +359,11 @@
 lemma HInfinite_add_ge_zero:
       "[|x \<in> HInfinite; 0 \<le> y; 0 \<le> x|] ==> (x + y): HInfinite"
 by (auto intro!: hypreal_add_zero_less_le_mono 
-       simp add: abs_if hypreal_add_commute hypreal_le_add_order HInfinite_def)
+       simp add: abs_if add_commute add_nonneg_nonneg HInfinite_def)
 
 lemma HInfinite_add_ge_zero2:
      "[|x \<in> HInfinite; 0 \<le> y; 0 \<le> x|] ==> (y + x): HInfinite"
-by (auto intro!: HInfinite_add_ge_zero simp add: hypreal_add_commute)
+by (auto intro!: HInfinite_add_ge_zero simp add: add_commute)
 
 lemma HInfinite_add_gt_zero:
      "[|x \<in> HInfinite; 0 < y; 0 < x|] ==> (x + y): HInfinite"
@@ -472,13 +472,13 @@
 by (simp add: approx_def)
 
 lemma approx_minus_iff2: " (x @= y) = (-y + x @= 0)"
-by (simp add: approx_def hypreal_add_commute)
+by (simp add: approx_def add_commute)
 
 lemma approx_refl [iff]: "x @= x"
 by (simp add: approx_def Infinitesimal_def)
 
 lemma hypreal_minus_distrib1: "-(y + -(x::hypreal)) = x + -y"
-by (simp add: hypreal_add_commute)
+by (simp add: add_commute)
 
 lemma approx_sym: "x @= y ==> y @= x"
 apply (simp add: approx_def)
@@ -624,7 +624,7 @@
 *}
 
 lemma Infinitesimal_approx_minus: "(x-y \<in> Infinitesimal) = (x @= y)"
-by (auto simp add: hypreal_diff_def approx_minus_iff [symmetric] mem_infmal_iff)
+by (auto simp add: diff_def approx_minus_iff [symmetric] mem_infmal_iff)
 
 lemma approx_monad_iff: "(x @= y) = (monad(x)=monad(y))"
 apply (simp add: monad_def)
@@ -648,7 +648,7 @@
 lemma approx_minus: "a @= b ==> -a @= -b"
 apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])
 apply (drule approx_minus_iff [THEN iffD1])
-apply (simp (no_asm) add: hypreal_add_commute)
+apply (simp (no_asm) add: add_commute)
 done
 
 lemma approx_minus2: "-a @= -b ==> a @= b"
@@ -666,7 +666,7 @@
          del: minus_mult_left [symmetric])
 
 lemma approx_mult2: "[|a @= b; c: HFinite|] ==> c*a @= c*b"
-by (simp add: approx_mult1 hypreal_mult_commute)
+by (simp add: approx_mult1 mult_commute)
 
 lemma approx_mult_subst: "[|u @= v*x; x @= y; v \<in> HFinite|] ==> u @= v*y"
 by (blast intro: approx_mult2 approx_trans)
@@ -694,17 +694,17 @@
 lemma Infinitesimal_add_approx: "[| y \<in> Infinitesimal; x + y = z |] ==> x @= z"
 apply (rule bex_Infinitesimal_iff [THEN iffD1])
 apply (drule Infinitesimal_minus_iff [THEN iffD2])
-apply (auto simp add: hypreal_add_assoc [symmetric])
+apply (auto simp add: add_assoc [symmetric])
 done
 
 lemma Infinitesimal_add_approx_self: "y \<in> Infinitesimal ==> x @= x + y"
 apply (rule bex_Infinitesimal_iff [THEN iffD1])
 apply (drule Infinitesimal_minus_iff [THEN iffD2])
-apply (auto simp add: hypreal_add_assoc [symmetric])
+apply (auto simp add: add_assoc [symmetric])
 done
 
 lemma Infinitesimal_add_approx_self2: "y \<in> Infinitesimal ==> x @= y + x"
-by (auto dest: Infinitesimal_add_approx_self simp add: hypreal_add_commute)
+by (auto dest: Infinitesimal_add_approx_self simp add: add_commute)
 
 lemma Infinitesimal_add_minus_approx_self: "y \<in> Infinitesimal ==> x @= x + -y"
 by (blast intro!: Infinitesimal_add_approx_self Infinitesimal_minus_iff [THEN iffD2])
@@ -718,7 +718,7 @@
      "[| y \<in> Infinitesimal; x @= z + y|] ==> x @= z"
 apply (drule_tac x = z in Infinitesimal_add_approx_self2 [THEN approx_sym])
 apply (erule approx_trans3 [THEN approx_sym])
-apply (simp add: hypreal_add_commute)
+apply (simp add: add_commute)
 apply (erule approx_sym)
 done
 
@@ -729,7 +729,7 @@
 
 lemma approx_add_right_cancel: "b + d @= c + d ==> b @= c"
 apply (rule approx_add_left_cancel)
-apply (simp add: hypreal_add_commute)
+apply (simp add: add_commute)
 done
 
 lemma approx_add_mono1: "b @= c ==> d + b @= d + c"
@@ -738,19 +738,19 @@
 done
 
 lemma approx_add_mono2: "b @= c ==> b + a @= c + a"
-by (simp add: hypreal_add_commute approx_add_mono1)
+by (simp add: add_commute approx_add_mono1)
 
 lemma approx_add_left_iff [simp]: "(a + b @= a + c) = (b @= c)"
 by (fast elim: approx_add_left_cancel approx_add_mono1)
 
 lemma approx_add_right_iff [simp]: "(b + a @= c + a) = (b @= c)"
-by (simp add: hypreal_add_commute)
+by (simp add: add_commute)
 
 lemma approx_HFinite: "[| x \<in> HFinite; x @= y |] ==> y \<in> HFinite"
 apply (drule bex_Infinitesimal_iff2 [THEN iffD2], safe)
 apply (drule Infinitesimal_subset_HFinite [THEN subsetD, THEN HFinite_minus_iff [THEN iffD2]])
 apply (drule HFinite_add)
-apply (auto simp add: hypreal_add_assoc)
+apply (auto simp add: add_assoc)
 done
 
 lemma approx_hypreal_of_real_HFinite: "x @= hypreal_of_real D ==> x \<in> HFinite"
@@ -773,7 +773,7 @@
 lemma approx_SReal_mult_cancel_zero:
      "[| a \<in> Reals; a \<noteq> 0; a*x @= 0 |] ==> x @= 0"
 apply (drule SReal_inverse [THEN SReal_subset_HFinite [THEN subsetD]])
-apply (auto dest: approx_mult2 simp add: hypreal_mult_assoc [symmetric])
+apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric])
 done
 
 lemma approx_mult_SReal1: "[| a \<in> Reals; x @= 0 |] ==> x*a @= 0"
@@ -789,7 +789,7 @@
 lemma approx_SReal_mult_cancel:
      "[| a \<in> Reals; a \<noteq> 0; a* w @= a*z |] ==> w @= z"
 apply (drule SReal_inverse [THEN SReal_subset_HFinite [THEN subsetD]])
-apply (auto dest: approx_mult2 simp add: hypreal_mult_assoc [symmetric])
+apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric])
 done
 
 lemma approx_SReal_mult_cancel_iff1 [simp]:
@@ -884,7 +884,7 @@
 lemma Infinitesimal_ratio:
      "[| y \<noteq> 0;  y \<in> Infinitesimal;  x/y \<in> HFinite |] ==> x \<in> Infinitesimal"
 apply (drule Infinitesimal_HFinite_mult2, assumption)
-apply (simp add: hypreal_divide_def hypreal_mult_assoc)
+apply (simp add: divide_inverse mult_assoc)
 done
 
 lemma Infinitesimal_SReal_divide: 
@@ -952,7 +952,7 @@
      "[| isLub R S x; isLub R S y |] ==> x = (y::hypreal)"
 apply (frule isLub_isUb)
 apply (frule_tac x = y in isLub_isUb)
-apply (blast intro!: hypreal_le_anti_sym dest!: isLub_le_isUb)
+apply (blast intro!: order_antisym dest!: isLub_le_isUb)
 done
 
 lemma lemma_st_part_ub:
@@ -980,7 +980,7 @@
 apply safe
 apply (drule_tac c = "-t" in add_left_mono)
 apply (drule_tac [2] c = t in add_left_mono)
-apply (auto simp add: hypreal_add_assoc [symmetric])
+apply (auto simp add: add_assoc [symmetric])
 done
 
 lemma lemma_st_part_le1:
@@ -1011,7 +1011,7 @@
 
 lemma lemma_minus_le_zero: "t \<le> t + -r ==> r \<le> (0::hypreal)"
 apply (drule_tac c = "-t" in add_left_mono)
-apply (auto simp add: hypreal_add_assoc [symmetric])
+apply (auto simp add: add_assoc [symmetric])
 done
 
 lemma lemma_st_part_le2:
@@ -1179,7 +1179,7 @@
      "x \<in> HFinite - Infinitesimal ==> inverse x \<in> HFinite - Infinitesimal"
 apply (auto intro: Infinitesimal_inverse_HFinite)
 apply (drule Infinitesimal_HFinite_mult2, assumption)
-apply (simp add: not_Infinitesimal_not_zero hypreal_mult_inverse)
+apply (simp add: not_Infinitesimal_not_zero right_inverse)
 done
 
 lemma approx_inverse:
@@ -1191,7 +1191,7 @@
 apply (drule HFinite_inverse2)+
 apply (drule approx_mult2, assumption, auto)
 apply (drule_tac c = "inverse x" in approx_mult1, assumption)
-apply (auto intro: approx_sym simp add: hypreal_mult_assoc)
+apply (auto intro: approx_sym simp add: mult_assoc)
 done
 
 (*Used for NSLIM_inverse, NSLIMSEQ_inverse*)
@@ -1206,7 +1206,7 @@
 lemma inverse_add_Infinitesimal_approx2:
      "[| x \<in> HFinite - Infinitesimal;
          h \<in> Infinitesimal |] ==> inverse(h + x) @= inverse x"
-apply (rule hypreal_add_commute [THEN subst])
+apply (rule add_commute [THEN subst])
 apply (blast intro: inverse_add_Infinitesimal_approx)
 done
 
@@ -1222,7 +1222,7 @@
 apply (auto intro: Infinitesimal_mult)
 apply (rule ccontr, frule Infinitesimal_inverse_HFinite)
 apply (frule not_Infinitesimal_not_zero)
-apply (auto dest: Infinitesimal_HFinite_mult simp add: hypreal_mult_assoc)
+apply (auto dest: Infinitesimal_HFinite_mult simp add: mult_assoc)
 done
 declare Infinitesimal_square_iff [symmetric, simp]
 
@@ -1239,7 +1239,7 @@
 apply safe
 apply (frule HFinite_inverse, assumption)
 apply (drule not_Infinitesimal_not_zero)
-apply (auto dest: approx_mult2 simp add: hypreal_mult_assoc [symmetric])
+apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric])
 done
 
 lemma approx_HFinite_mult_cancel_iff1:
@@ -1256,7 +1256,7 @@
 lemma HInfinite_HFinite_add:
      "[| x \<in> HInfinite; y \<in> HFinite |] ==> x + y \<in> HInfinite"
 apply (rule_tac y = "-y" in HInfinite_HFinite_add_cancel)
-apply (auto simp add: hypreal_add_assoc HFinite_minus_iff)
+apply (auto simp add: add_assoc HFinite_minus_iff)
 done
 
 lemma HInfinite_ge_HInfinite:
@@ -1276,13 +1276,13 @@
 apply (frule HFinite_Infinitesimal_not_zero)
 apply (drule HFinite_not_Infinitesimal_inverse)
 apply (safe, drule HFinite_mult)
-apply (auto simp add: hypreal_mult_assoc HFinite_HInfinite_iff)
+apply (auto simp add: mult_assoc HFinite_HInfinite_iff)
 done
 
 lemma HInfinite_HFinite_not_Infinitesimal_mult2:
      "[| x \<in> HInfinite; y \<in> HFinite - Infinitesimal |]
       ==> y * x \<in> HInfinite"
-by (auto simp add: hypreal_mult_commute HInfinite_HFinite_not_Infinitesimal_mult)
+by (auto simp add: mult_commute HInfinite_HFinite_not_Infinitesimal_mult)
 
 lemma HInfinite_gt_SReal: "[| x \<in> HInfinite; 0 < x; y \<in> Reals |] ==> y < x"
 by (auto dest!: bspec simp add: HInfinite_def abs_if order_less_imp_le)
@@ -1443,13 +1443,14 @@
 apply (drule_tac x = "hypreal_of_real r" in approx_hrabs_add_Infinitesimal)
 apply (drule approx_sym [THEN bex_Infinitesimal_iff2 [THEN iffD2]])
 apply (auto intro!: Infinitesimal_add_hypreal_of_real_less
+            simp del: star_of_abs
             simp add: hypreal_of_real_hrabs)
 done
 
 lemma Infinitesimal_add_hrabs_hypreal_of_real_less2:
      "[| x \<in> Infinitesimal;  abs(hypreal_of_real r) < hypreal_of_real y |]
       ==> abs (x + hypreal_of_real r) < hypreal_of_real y"
-apply (rule hypreal_add_commute [THEN subst])
+apply (rule add_commute [THEN subst])
 apply (erule Infinitesimal_add_hrabs_hypreal_of_real_less, assumption)
 done
 
@@ -1466,8 +1467,8 @@
      "[| u \<in> Infinitesimal; v \<in> Infinitesimal;
          hypreal_of_real x + u \<le> hypreal_of_real y + v |]
       ==> x \<le> y"
-by (blast intro!: hypreal_of_real_le_iff [THEN iffD1] 
-                  hypreal_of_real_le_add_Infininitesimal_cancel)
+by (blast intro: star_of_le [THEN iffD1] 
+          intro!: hypreal_of_real_le_add_Infininitesimal_cancel)
 
 lemma hypreal_of_real_less_Infinitesimal_le_zero:
     "[| hypreal_of_real x < e; e \<in> Infinitesimal |] ==> hypreal_of_real x \<le> 0"
@@ -1498,13 +1499,13 @@
 lemma Infinitesimal_square_cancel2 [simp]:
      "x*x + y*y \<in> Infinitesimal ==> y*y \<in> Infinitesimal"
 apply (rule Infinitesimal_square_cancel)
-apply (rule hypreal_add_commute [THEN subst])
+apply (rule add_commute [THEN subst])
 apply (simp (no_asm))
 done
 
 lemma HFinite_square_cancel2 [simp]: "x*x + y*y \<in> HFinite ==> y*y \<in> HFinite"
 apply (rule HFinite_square_cancel)
-apply (rule hypreal_add_commute [THEN subst])
+apply (rule add_commute [THEN subst])
 apply (simp (no_asm))
 done
 
@@ -1625,7 +1626,7 @@
 
 lemma st_Infinitesimal_add_SReal2:
      "[| x \<in> Reals; e \<in> Infinitesimal |] ==> st(e + x) = x"
-apply (rule hypreal_add_commute [THEN subst])
+apply (rule add_commute [THEN subst])
 apply (blast intro!: st_Infinitesimal_add_SReal)
 done
 
@@ -1667,7 +1668,7 @@
 qed
 
 lemma st_diff: "[| x \<in> HFinite; y \<in> HFinite |] ==> st (x-y) = st(x) - st(y)"
-apply (simp add: hypreal_diff_def)
+apply (simp add: diff_def)
 apply (frule_tac y1 = y in st_minus [symmetric])
 apply (drule_tac x1 = y in HFinite_minus_iff [THEN iffD2])
 apply (simp (no_asm_simp) add: st_add)
@@ -1692,11 +1693,11 @@
 apply (erule_tac V = "y = st y + ea" in thin_rl)
 apply (simp add: left_distrib right_distrib)
 apply (drule st_SReal)+
-apply (simp (no_asm_use) add: hypreal_add_assoc)
+apply (simp (no_asm_use) add: add_assoc)
 apply (rule st_Infinitesimal_add_SReal)
 apply (blast intro!: SReal_mult)
 apply (drule SReal_subset_HFinite [THEN subsetD])+
-apply (rule hypreal_add_assoc [THEN subst])
+apply (rule add_assoc [THEN subst])
 apply (blast intro!: lemma_st_mult)
 done
 
@@ -1716,13 +1717,13 @@
       ==> st(inverse x) = inverse (st x)"
 apply (rule_tac c1 = "st x" in hypreal_mult_left_cancel [THEN iffD1])
 apply (auto simp add: st_mult [symmetric] st_not_Infinitesimal HFinite_inverse)
-apply (subst hypreal_mult_inverse, auto)
+apply (subst right_inverse, auto)
 done
 
 lemma st_divide [simp]:
      "[| x \<in> HFinite; y \<in> HFinite; st y \<noteq> 0 |]
       ==> st(x/y) = (st x) / (st y)"
-by (simp add: hypreal_divide_def st_mult st_not_Infinitesimal HFinite_inverse st_inverse)
+by (simp add: divide_inverse st_mult st_not_Infinitesimal HFinite_inverse st_inverse)
 
 lemma st_idempotent [simp]: "x \<in> HFinite ==> st(st(x)) = st(x)"
 by (blast intro: st_HFinite st_approx_self approx_st_eq)
@@ -1748,7 +1749,7 @@
 apply (frule HFinite_st_Infinitesimal_add, safe)
 apply (rule Infinitesimal_add_st_le_cancel)
 apply (rule_tac [3] x = ea and y = e in Infinitesimal_diff)
-apply (auto simp add: hypreal_add_assoc [symmetric])
+apply (auto simp add: add_assoc [symmetric])
 done
 
 lemma st_zero_le: "[| 0 \<le> x;  x \<in> HFinite |] ==> 0 \<le> st x"
@@ -1781,29 +1782,30 @@
 lemma HFinite_FreeUltrafilterNat:
     "x \<in> HFinite 
      ==> \<exists>X \<in> Rep_star x. \<exists>u. {n. abs (X n) < u} \<in> FreeUltrafilterNat"
-apply (rule_tac z = x in eq_Abs_star)
+apply (cases x)
 apply (auto simp add: HFinite_def abs_less_iff minus_less_iff [of x] 
-              star_of_def star_n_def
-              hypreal_less SReal_iff hypreal_minus hypreal_of_real_def)
-apply (rule_tac x=x in bexI) 
-apply (rule_tac x=y in exI, auto, ultra)
+              star_of_def
+              star_n_less SReal_iff star_n_minus)
+apply (rule_tac x=X in bexI)
+apply (rule_tac x=y in exI, ultra)
+apply simp
 done
 
 lemma FreeUltrafilterNat_HFinite:
      "\<exists>X \<in> Rep_star x.
        \<exists>u. {n. abs (X n) < u} \<in> FreeUltrafilterNat
        ==>  x \<in> HFinite"
-apply (rule_tac z = x in eq_Abs_star)
+apply (cases x)
 apply (auto simp add: HFinite_def abs_less_iff minus_less_iff [of x])
 apply (rule_tac x = "hypreal_of_real u" in bexI)
-apply (auto simp add: hypreal_less SReal_iff hypreal_minus hypreal_of_real_def star_of_def star_n_def, ultra+)
+apply (auto simp add: star_n_less SReal_iff star_n_minus star_of_def)
+apply ultra+
 done
 
 lemma HFinite_FreeUltrafilterNat_iff:
      "(x \<in> HFinite) = (\<exists>X \<in> Rep_star x.
            \<exists>u. {n. abs (X n) < u} \<in> FreeUltrafilterNat)"
-apply (blast intro!: HFinite_FreeUltrafilterNat FreeUltrafilterNat_HFinite)
-done
+by (blast intro!: HFinite_FreeUltrafilterNat FreeUltrafilterNat_HFinite)
 
 
 subsection{*Alternative Definitions for @{term HInfinite} using Free Ultrafilter*}
@@ -1817,8 +1819,7 @@
 lemma lemma_Int_eq1:
      "{n. abs (xa n) \<le> (u::real)} Int {n. u \<le> abs (xa n)}
           = {n. abs(xa n) = u}"
-apply auto
-done
+by auto
 
 lemma lemma_FreeUltrafilterNat_one:
      "{n. abs (xa n) = u} \<le> {n. abs (xa n) < u + (1::real)}"
@@ -1888,11 +1889,11 @@
            \<forall>u. 0 < u --> {n. abs (X n) < u} \<in> FreeUltrafilterNat"
 apply (simp add: Infinitesimal_def)
 apply (auto simp add: abs_less_iff minus_less_iff [of x])
-apply (rule_tac z = x in eq_Abs_star)
-apply (auto, rule bexI [OF _ lemma_starrel_refl], safe)
-apply (drule hypreal_of_real_less_iff [THEN iffD2])
+apply (cases x)
+apply (auto, rule bexI [OF _ Rep_star_star_n], safe)
+apply (drule star_of_less [THEN iffD2])
 apply (drule_tac x = "hypreal_of_real u" in bspec, auto)
-apply (auto simp add: hypreal_less hypreal_minus hypreal_of_real_def star_of_def star_n_def, ultra)
+apply (auto simp add: star_n_less star_n_minus star_of_def, ultra)
 done
 
 lemma FreeUltrafilterNat_Infinitesimal:
@@ -1900,18 +1901,17 @@
             \<forall>u. 0 < u --> {n. abs (X n) < u} \<in> FreeUltrafilterNat
       ==> x \<in> Infinitesimal"
 apply (simp add: Infinitesimal_def)
-apply (rule_tac z = x in eq_Abs_star)
+apply (cases x)
 apply (auto simp add: abs_less_iff abs_interval_iff minus_less_iff [of x])
 apply (auto simp add: SReal_iff)
 apply (drule_tac [!] x=y in spec) 
-apply (auto simp add: hypreal_less hypreal_minus hypreal_of_real_def star_of_def star_n_def, ultra+)
+apply (auto simp add: star_n_less star_n_minus star_of_def, ultra+)
 done
 
 lemma Infinitesimal_FreeUltrafilterNat_iff:
      "(x \<in> Infinitesimal) = (\<exists>X \<in> Rep_star x.
            \<forall>u. 0 < u --> {n. abs (X n) < u} \<in> FreeUltrafilterNat)"
-apply (blast intro!: Infinitesimal_FreeUltrafilterNat FreeUltrafilterNat_Infinitesimal)
-done
+by (blast intro!: Infinitesimal_FreeUltrafilterNat FreeUltrafilterNat_Infinitesimal)
 
 (*------------------------------------------------------------------------
          Infinitesimals as smaller than 1/n for all n::nat (> 0)
@@ -1934,11 +1934,11 @@
 apply safe
 apply (drule_tac x = "inverse (hypreal_of_real (real (Suc n))) " in bspec)
 apply (simp (no_asm_use) add: SReal_inverse)
-apply (rule real_of_nat_Suc_gt_zero [THEN positive_imp_inverse_positive, THEN hypreal_of_real_less_iff [THEN iffD2], THEN [2] impE])
+apply (rule real_of_nat_Suc_gt_zero [THEN positive_imp_inverse_positive, THEN star_of_less [THEN iffD2], THEN [2] impE])
 prefer 2 apply assumption
 apply (simp add: real_of_nat_Suc_gt_zero hypreal_of_nat_eq)
 apply (auto dest!: reals_Archimedean simp add: SReal_iff)
-apply (drule hypreal_of_real_less_iff [THEN iffD2])
+apply (drule star_of_less [THEN iffD2])
 apply (simp add: real_of_nat_Suc_gt_zero hypreal_of_nat_eq)
 apply (blast intro: order_less_trans)
 done
@@ -2009,9 +2009,6 @@
 
 text{*@{term omega} is a member of @{term HInfinite}*}
 
-lemma hypreal_omega: "starrel``{%n::nat. real (Suc n)} \<in> star"
-by auto
-
 lemma FreeUltrafilterNat_omega: "{n. u < real n} \<in> FreeUltrafilterNat"
 apply (cut_tac u = u in rabs_real_of_nat_le_real_FreeUltrafilterNat)
 apply (auto dest: FreeUltrafilterNat_Compl_mem simp add: Compl_real_le_eq)
@@ -2123,13 +2120,13 @@
  -----------------------------------------------------*)
 lemma real_seq_to_hypreal_Infinitesimal:
      "\<forall>n. abs(X n + -x) < inverse(real(Suc n))
-     ==> Abs_star(starrel``{X}) + -hypreal_of_real x \<in> Infinitesimal"
-apply (auto intro!: bexI dest: FreeUltrafilterNat_inverse_real_of_posnat FreeUltrafilterNat_all FreeUltrafilterNat_Int intro: order_less_trans FreeUltrafilterNat_subset simp add: hypreal_minus hypreal_of_real_def star_of_def star_n_def hypreal_add Infinitesimal_FreeUltrafilterNat_iff hypreal_inverse)
+     ==> star_n X + -hypreal_of_real x \<in> Infinitesimal"
+apply (auto intro!: bexI dest: FreeUltrafilterNat_inverse_real_of_posnat FreeUltrafilterNat_all FreeUltrafilterNat_Int intro: order_less_trans FreeUltrafilterNat_subset simp add: star_n_minus star_of_def star_n_add Infinitesimal_FreeUltrafilterNat_iff star_n_inverse)
 done
 
 lemma real_seq_to_hypreal_approx:
      "\<forall>n. abs(X n + -x) < inverse(real(Suc n))
-      ==> Abs_star(starrel``{X}) @= hypreal_of_real x"
+      ==> star_n X @= hypreal_of_real x"
 apply (subst approx_minus_iff)
 apply (rule mem_infmal_iff [THEN subst])
 apply (erule real_seq_to_hypreal_Infinitesimal)
@@ -2137,20 +2134,19 @@
 
 lemma real_seq_to_hypreal_approx2:
      "\<forall>n. abs(x + -X n) < inverse(real(Suc n))
-               ==> Abs_star(starrel``{X}) @= hypreal_of_real x"
+               ==> star_n X @= hypreal_of_real x"
 apply (simp add: abs_minus_add_cancel real_seq_to_hypreal_approx)
 done
 
 lemma real_seq_to_hypreal_Infinitesimal2:
      "\<forall>n. abs(X n + -Y n) < inverse(real(Suc n))
-      ==> Abs_star(starrel``{X}) +
-          -Abs_star(starrel``{Y}) \<in> Infinitesimal"
+      ==> star_n X + -star_n Y \<in> Infinitesimal"
 by (auto intro!: bexI
 	 dest: FreeUltrafilterNat_inverse_real_of_posnat 
 	       FreeUltrafilterNat_all FreeUltrafilterNat_Int
 	 intro: order_less_trans FreeUltrafilterNat_subset 
-	 simp add: Infinitesimal_FreeUltrafilterNat_iff hypreal_minus 
-                   hypreal_add hypreal_inverse)
+	 simp add: Infinitesimal_FreeUltrafilterNat_iff star_n_minus 
+                   star_n_add star_n_inverse)
 
 
 ML
@@ -2335,7 +2331,6 @@
 val finite_rabs_real_of_nat_le_real = thm "finite_rabs_real_of_nat_le_real";
 val rabs_real_of_nat_le_real_FreeUltrafilterNat = thm "rabs_real_of_nat_le_real_FreeUltrafilterNat";
 val FreeUltrafilterNat_nat_gt_real = thm "FreeUltrafilterNat_nat_gt_real";
-val hypreal_omega = thm "hypreal_omega";
 val FreeUltrafilterNat_omega = thm "FreeUltrafilterNat_omega";
 val HInfinite_omega = thm "HInfinite_omega";
 val Infinitesimal_epsilon = thm "Infinitesimal_epsilon";
--- a/src/HOL/Hyperreal/NatStar.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/NatStar.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -11,523 +11,211 @@
 imports HyperPow
 begin
 
-text{*Extends sets of nats, and functions from the nats to nats or reals*}
-
-
-constdefs
-
-  (* internal sets and nonstandard extensions -- see Star.thy as well *)
-
-    starsetNat :: "nat set => hypnat set"          ("*sNat* _" [80] 80)
-    "*sNat* A ==
-       {x. \<forall>X\<in>Rep_star(x). {n::nat. X n \<in> A}: FreeUltrafilterNat}"
-
-   starsetNat_n :: "(nat => nat set) => hypnat set"      ("*sNatn* _" [80] 80)
-    "*sNatn* As ==
-       {x. \<forall>X\<in>Rep_star(x). {n::nat. X n \<in> (As n)}: FreeUltrafilterNat}"
-
-   InternalNatSets :: "hypnat set set"
-    "InternalNatSets == {X. \<exists>As. X = *sNatn* As}"
-
-    (* star transform of functions f:Nat --> Real *)
-
-    starfunNat :: "(nat => real) => hypnat => hypreal"
-                  ("*fNat* _" [80] 80)
-    "*fNat* f  == (%x. Abs_star(\<Union>X\<in>Rep_star(x). starrel``{%n. f (X n)}))"
-
-   starfunNat_n :: "(nat => (nat => real)) => hypnat => hypreal"
-                   ("*fNatn* _" [80] 80)
-    "*fNatn* F ==
-      (%x. Abs_star(\<Union>X\<in>Rep_star(x). starrel``{%n. (F n)(X n)}))"
-
-   InternalNatFuns :: "(hypnat => hypreal) set"
-    "InternalNatFuns == {X. \<exists>F. X = *fNatn* F}"
-
-    (* star transform of functions f:Nat --> Nat *)
-
-   starfunNat2 :: "(nat => nat) => hypnat => hypnat"
-                   ("*fNat2* _" [80] 80)
-    "*fNat2* f == %x. Abs_star(\<Union>X\<in>Rep_star(x). starrel``{%n. f (X n)})"
-
-   starfunNat2_n :: "(nat => (nat => nat)) => hypnat => hypnat"
-                     ("*fNat2n* _" [80] 80)
-    "*fNat2n* F == 
-        %x. Abs_star(\<Union>X\<in>Rep_star(x). starrel``{%n. (F n)(X n)})"
-
-   InternalNatFuns2 :: "(hypnat => hypnat) set"
-    "InternalNatFuns2 == {X. \<exists>F. X = *fNat2n* F}"
-
-
-lemma NatStar_real_set: "*sNat*(UNIV::nat set) = (UNIV::hypnat set)"
-by (simp add: starsetNat_def)
-
-lemma NatStar_empty_set [simp]: "*sNat* {} = {}"
-by (simp add: starsetNat_def)
-
-lemma NatStar_Un: "*sNat* (A Un B) = *sNat* A Un *sNat* B"
-apply (auto simp add: starsetNat_def)
- prefer 2 apply (blast intro: FreeUltrafilterNat_subset)
- prefer 2 apply (blast intro: FreeUltrafilterNat_subset)
-apply (drule FreeUltrafilterNat_Compl_mem)
-apply (drule bspec, assumption)
-apply (rule_tac z = x in eq_Abs_star, auto, ultra)
+lemma starset_n_Un: "*sn* (%n. (A n) Un (B n)) = *sn* A Un *sn* B"
+apply (simp add: starset_n_def expand_set_eq all_star_eq)
+apply (simp add: Iset_star_n fuf_disj)
 done
 
-lemma starsetNat_n_Un: "*sNatn* (%n. (A n) Un (B n)) = *sNatn* A Un *sNatn* B"
-apply (auto simp add: starsetNat_n_def)
-apply (drule_tac x = Xa in bspec)
-apply (rule_tac [2] z = x in eq_Abs_star)
-apply (auto dest!: bspec, ultra+)
+lemma InternalSets_Un:
+     "[| X \<in> InternalSets; Y \<in> InternalSets |]
+      ==> (X Un Y) \<in> InternalSets"
+by (auto simp add: InternalSets_def starset_n_Un [symmetric])
+
+lemma starset_n_Int:
+      "*sn* (%n. (A n) Int (B n)) = *sn* A Int *sn* B"
+apply (simp add: starset_n_def expand_set_eq all_star_eq)
+apply (simp add: Iset_star_n fuf_conj)
 done
 
-lemma InternalNatSets_Un:
-     "[| X \<in> InternalNatSets; Y \<in> InternalNatSets |]
-      ==> (X Un Y) \<in> InternalNatSets"
-by (auto simp add: InternalNatSets_def starsetNat_n_Un [symmetric])
+lemma InternalSets_Int:
+     "[| X \<in> InternalSets; Y \<in> InternalSets |]
+      ==> (X Int Y) \<in> InternalSets"
+by (auto simp add: InternalSets_def starset_n_Int [symmetric])
 
-lemma NatStar_Int: "*sNat* (A Int B) = *sNat* A Int *sNat* B"
-apply (auto simp add: starsetNat_def)
-prefer 3 apply (blast intro: FreeUltrafilterNat_Int FreeUltrafilterNat_subset)
-apply (blast intro: FreeUltrafilterNat_subset)+
+lemma starset_n_Compl: "*sn* ((%n. - A n)) = -( *sn* A)"
+apply (simp add: starset_n_def expand_set_eq all_star_eq)
+apply (simp add: Iset_star_n fuf_not)
 done
 
-lemma starsetNat_n_Int:
-      "*sNatn* (%n. (A n) Int (B n)) = *sNatn* A Int *sNatn* B"
-apply (auto simp add: starsetNat_n_def)
-apply (auto dest!: bspec, ultra+)
-done
-
-lemma InternalNatSets_Int:
-     "[| X \<in> InternalNatSets; Y \<in> InternalNatSets |]
-      ==> (X Int Y) \<in> InternalNatSets"
-by (auto simp add: InternalNatSets_def starsetNat_n_Int [symmetric])
+lemma InternalSets_Compl: "X \<in> InternalSets ==> -X \<in> InternalSets"
+by (auto simp add: InternalSets_def starset_n_Compl [symmetric])
 
-lemma NatStar_Compl: "*sNat* (-A) = -( *sNat* A)"
-apply (auto simp add: starsetNat_def)
-apply (rule_tac z = x in eq_Abs_star)
-apply (rule_tac [2] z = x in eq_Abs_star)
-apply (auto dest!: bspec, ultra+)
-done
-
-lemma starsetNat_n_Compl: "*sNatn* ((%n. - A n)) = -( *sNatn* A)"
-apply (auto simp add: starsetNat_n_def)
-apply (rule_tac z = x in eq_Abs_star)
-apply (rule_tac [2] z = x in eq_Abs_star)
-apply (auto dest!: bspec, ultra+)
-done
-
-lemma InternalNatSets_Compl: "X \<in> InternalNatSets ==> -X \<in> InternalNatSets"
-by (auto simp add: InternalNatSets_def starsetNat_n_Compl [symmetric])
-
-lemma starsetNat_n_diff: "*sNatn* (%n. (A n) - (B n)) = *sNatn* A - *sNatn* B"
-apply (auto simp add: starsetNat_n_def)
-apply (rule_tac [2] z = x in eq_Abs_star)
-apply (rule_tac [3] z = x in eq_Abs_star)
-apply (auto dest!: bspec, ultra+)
+lemma starset_n_diff: "*sn* (%n. (A n) - (B n)) = *sn* A - *sn* B"
+apply (simp add: starset_n_def expand_set_eq all_star_eq)
+apply (simp add: Iset_star_n fuf_conj fuf_not)
 done
 
-lemma InternalNatSets_diff:
-     "[| X \<in> InternalNatSets; Y \<in> InternalNatSets |]
-      ==> (X - Y) \<in> InternalNatSets"
-by (auto simp add: InternalNatSets_def starsetNat_n_diff [symmetric])
-
-lemma NatStar_subset: "A \<le> B ==> *sNat* A \<le> *sNat* B"
-apply (simp add: starsetNat_def)
-apply (blast intro: FreeUltrafilterNat_subset)
-done
+lemma InternalSets_diff:
+     "[| X \<in> InternalSets; Y \<in> InternalSets |]
+      ==> (X - Y) \<in> InternalSets"
+by (auto simp add: InternalSets_def starset_n_diff [symmetric])
 
-lemma NatStar_mem: "a \<in> A ==> hypnat_of_nat a \<in> *sNat* A"
-by (auto intro: FreeUltrafilterNat_subset 
-         simp add: starsetNat_def hypnat_of_nat_eq)
-
-lemma NatStar_hypreal_of_real_image_subset: "hypnat_of_nat ` A \<le> *sNat* A"
-apply (auto simp add: starsetNat_def hypnat_of_nat_eq)
-apply (blast intro: FreeUltrafilterNat_subset)
-done
-
-lemma NatStar_SHNat_subset: "Nats \<le> *sNat* (UNIV:: nat set)"
-by (simp add: starsetNat_def SHNat_eq hypnat_of_nat_eq)
+lemma NatStar_SHNat_subset: "Nats \<le> *s* (UNIV:: nat set)"
+by simp
 
 lemma NatStar_hypreal_of_real_Int:
-     "*sNat* X Int Nats = hypnat_of_nat ` X"
-apply (auto simp add: starsetNat_def hypnat_of_nat_eq SHNat_eq)
-apply (simp add: hypnat_of_nat_eq [symmetric])
-apply (rule imageI, rule ccontr)
-apply (drule bspec)
-apply (rule lemma_starrel_refl)
-prefer 2 apply (blast intro: FreeUltrafilterNat_subset, auto)
-done
+     "*s* X Int Nats = hypnat_of_nat ` X"
+by (auto simp add: SHNat_eq STAR_mem_iff)
 
-lemma starsetNat_starsetNat_n_eq: "*sNat* X = *sNatn* (%n. X)"
-by (simp add: starsetNat_n_def starsetNat_def)
+lemma starset_starset_n_eq: "*s* X = *sn* (%n. X)"
+by (simp add: starset_n_starset)
 
-lemma InternalNatSets_starsetNat_n [simp]: "( *sNat* X) \<in> InternalNatSets"
-by (auto simp add: InternalNatSets_def starsetNat_starsetNat_n_eq)
+lemma InternalSets_starset_n [simp]: "( *s* X) \<in> InternalSets"
+by (auto simp add: InternalSets_def starset_starset_n_eq)
 
-lemma InternalNatSets_UNIV_diff:
-     "X \<in> InternalNatSets ==> UNIV - X \<in> InternalNatSets"
+lemma InternalSets_UNIV_diff:
+     "X \<in> InternalSets ==> UNIV - X \<in> InternalSets"
 apply (subgoal_tac "UNIV - X = - X")
-by (auto intro: InternalNatSets_Compl)
-
-text{*Nonstandard extension of a set (defined using a constant
-   sequence) as a special case of an internal set*}
-lemma starsetNat_n_starsetNat: "\<forall>n. (As n = A) ==> *sNatn* As = *sNat* A"
-by (simp add: starsetNat_n_def starsetNat_def)
+by (auto intro: InternalSets_Compl)
 
 
 subsection{*Nonstandard Extensions of Functions*}
 
-text{* Nonstandard extension of a function (defined using a constant
-   sequence) as a special case of an internal function*}
-
-lemma starfunNat_n_starfunNat: "\<forall>n. (F n = f) ==> *fNatn* F = *fNat* f"
-by (simp add: starfunNat_n_def starfunNat_def)
-
-lemma starfunNat2_n_starfunNat2: "\<forall>n. (F n = f) ==> *fNat2n* F = *fNat2* f"
-by (simp add: starfunNat2_n_def starfunNat2_def)
-
-lemma starfunNat_congruent: "(%X. starrel``{%n. f (X n)}) respects starrel"
-apply (simp add: congruent_def, safe)
-apply (ultra+)
-done
-
-(* f::nat=>real *)
-lemma starfunNat:
-      "( *fNat* f) (Abs_star(starrel``{%n. X n})) =
-       Abs_star(starrel `` {%n. f (X n)})"
-apply (simp add: starfunNat_def)
-apply (rule arg_cong [where f = Abs_star])
-apply (auto, ultra)
-done
-
-(* f::nat=>nat *)
-lemma starfunNat2:
-      "( *fNat2* f) (Abs_star(starrel``{%n. X n})) =
-       Abs_star(starrel `` {%n. f (X n)})"
-apply (simp add: starfunNat2_def)
-apply (rule arg_cong [where f = Abs_star])
-apply (simp add: starrel_in_hypreal [THEN Abs_star_inverse]
-         UN_equiv_class [OF equiv_starrel starfunNat_congruent])
-done
-
-subsubsection{*Multiplication: @{text "( *f) x ( *g) = *(f x g)"}*}
-
-lemma starfunNat_mult:
-     "( *fNat* f) z * ( *fNat* g) z = ( *fNat* (%x. f x * g x)) z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunNat hypreal_mult)
-done
-
-lemma starfunNat2_mult:
-     "( *fNat2* f) z * ( *fNat2* g) z = ( *fNat2* (%x. f x * g x)) z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunNat2 hypnat_mult)
-done
-
-subsubsection{*Addition: @{text "( *f) + ( *g) = *(f + g)"}*}
-
-lemma starfunNat_add:
-     "( *fNat* f) z + ( *fNat* g) z = ( *fNat* (%x. f x + g x)) z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunNat hypreal_add)
-done
-
-lemma starfunNat2_add:
-     "( *fNat2* f) z + ( *fNat2* g) z = ( *fNat2* (%x. f x + g x)) z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunNat2 hypnat_add)
-done
-
-lemma starfunNat2_minus:
-     "( *fNat2* f) z - ( *fNat2* g) z = ( *fNat2* (%x. f x - g x)) z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunNat2 hypnat_minus)
-done
-
-subsubsection{*Composition: @{text "( *f) o ( *g) = *(f o g)"}*}
-
-(***** ( *f::nat=>real) o ( *g::nat=>nat) = *(f o g) *****)
-
-lemma starfunNatNat2_o:
-     "( *fNat* f) o ( *fNat2* g) = ( *fNat* (f o g))"
-apply (rule ext)
-apply (rule_tac z = x in eq_Abs_star)
-apply (simp add: starfunNat2 starfunNat)
-done
-
-lemma starfunNatNat2_o2:
-     "(%x. ( *fNat* f) (( *fNat2* g) x)) = ( *fNat* (%x. f(g x)))"
-apply (insert starfunNatNat2_o [of f g]) 
-apply (simp add: o_def) 
-done
-
-(***** ( *f::nat=>nat) o ( *g::nat=>nat) = *(f o g) *****)
-lemma starfunNat2_o: "( *fNat2* f) o ( *fNat2* g) = ( *fNat2* (f o g))"
-apply (rule ext)
-apply (rule_tac z = x in eq_Abs_star)
-apply (simp add: starfunNat2)
-done
-
-(***** ( *f::real=>real) o ( *g::nat=>real) = *(f o g) *****)
-
-lemma starfun_stafunNat_o: "( *f* f) o ( *fNat* g) = ( *fNat* (f o g))"
-apply (rule ext)
-apply (rule_tac z = x in eq_Abs_star)
-apply (simp add: starfunNat starfun)
-done
-
-lemma starfun_stafunNat_o2:
-     "(%x. ( *f* f) (( *fNat* g) x)) = ( *fNat* (%x. f (g x)))"
-apply (insert starfun_stafunNat_o [of f g]) 
-apply (simp add: o_def) 
-done
-
-
-text{*NS extension of constant function*}
-
-lemma starfunNat_const_fun [simp]: "( *fNat* (%x. k)) z = hypreal_of_real k"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunNat hypreal_of_real_def star_of_def star_n_def)
-done
-
-lemma starfunNat2_const_fun [simp]: "( *fNat2* (%x. k)) z = hypnat_of_nat  k"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunNat2 hypnat_of_nat_eq)
-done
-
-lemma starfunNat_minus: "- ( *fNat* f) x = ( *fNat* (%x. - f x)) x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfunNat hypreal_minus)
-done
-
-lemma starfunNat_inverse:
-     "inverse (( *fNat* f) x) = ( *fNat* (%x. inverse (f x))) x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfunNat hypreal_inverse)
-done
-
-text{* Extended 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)*}
-
-lemma starfunNat_eq [simp]:
-     "( *fNat* f) (hypnat_of_nat a) = hypreal_of_real (f a)"
-by (simp add: starfunNat hypnat_of_nat_eq hypreal_of_real_def star_of_def star_n_def)
-
-lemma starfunNat2_eq [simp]:
-     "( *fNat2* f) (hypnat_of_nat a) = hypnat_of_nat (f a)"
-by (simp add: starfunNat2 hypnat_of_nat_eq)
-
-lemma starfunNat_approx:
-     "( *fNat* f) (hypnat_of_nat a) @= hypreal_of_real (f a)"
-by auto
-
-
 text{* Example of transfer of a property from reals to hyperreals
     --- used for limit comparison of sequences*}
 
 lemma starfun_le_mono:
      "\<forall>n. N \<le> n --> f n \<le> g n
-      ==> \<forall>n. hypnat_of_nat N \<le> n --> ( *fNat* f) n \<le> ( *fNat* g) n"
-apply safe
-apply (rule_tac z = n in eq_Abs_star)
-apply (auto simp add: starfunNat hypnat_of_nat_eq hypreal_le hypreal_less hypnat_le hypnat_less, ultra, auto)
-done
+      ==> \<forall>n. hypnat_of_nat N \<le> n --> ( *f* f) n \<le> ( *f* g) n"
+by transfer
 
 (*****----- and another -----*****)
 lemma starfun_less_mono:
      "\<forall>n. N \<le> n --> f n < g n
-      ==> \<forall>n. hypnat_of_nat N \<le> n --> ( *fNat* f) n < ( *fNat* g) n"
-apply safe
-apply (rule_tac z = n in eq_Abs_star)
-apply (auto simp add: starfunNat hypnat_of_nat_eq hypreal_le hypreal_less hypnat_le hypnat_less, ultra, auto)
-done
+      ==> \<forall>n. hypnat_of_nat N \<le> n --> ( *f* f) n < ( *f* g) n"
+by transfer
 
 text{*Nonstandard extension when we increment the argument by one*}
 
-lemma starfunNat_shift_one:
-     "( *fNat* (%n. f (Suc n))) N = ( *fNat* f) (N + (1::hypnat))"
-apply (rule_tac z=N in eq_Abs_star)
-apply (simp add: starfunNat hypnat_one_def hypnat_add)
-done
+lemma starfun_shift_one:
+     "!!N. ( *f* (%n. f (Suc n))) N = ( *f* f) (N + (1::hypnat))"
+by (transfer, simp)
 
 text{*Nonstandard extension with absolute value*}
 
-lemma starfunNat_rabs: "( *fNat* (%n. abs (f n))) N = abs(( *fNat* f) N)"
-apply (rule_tac z=N in eq_Abs_star)
-apply (simp add: starfunNat hypreal_hrabs)
-done
+lemma starfun_abs: "!!N. ( *f* (%n. abs (f n))) N = abs(( *f* f) N)"
+by (transfer, rule refl)
 
 text{*The hyperpow function as a nonstandard extension of realpow*}
 
-lemma starfunNat_pow: "( *fNat* (%n. r ^ n)) N = (hypreal_of_real r) pow N"
-apply (rule_tac z=N in eq_Abs_star)
-apply (simp add: hyperpow hypreal_of_real_def star_of_def star_n_def starfunNat)
-done
+lemma starfun_pow: "!!N. ( *f* (%n. r ^ n)) N = (hypreal_of_real r) pow N"
+by (unfold hyperpow_def, transfer, rule refl)
 
-lemma starfunNat_pow2:
-     "( *fNat* (%n. (X n) ^ m)) N = ( *fNat* X) N pow hypnat_of_nat m"
-apply (rule_tac z=N in eq_Abs_star)
-apply (simp add: hyperpow hypnat_of_nat_eq starfunNat)
-done
+lemma starfun_pow2:
+     "!!N. ( *f* (%n. (X n) ^ m)) N = ( *f* X) N pow hypnat_of_nat m"
+by (unfold hyperpow_def, transfer, rule refl)
 
-lemma starfun_pow: "( *f* (%r. r ^ n)) R = (R) pow hypnat_of_nat n"
-apply (rule_tac z = R in eq_Abs_star)
-apply (simp add: hyperpow starfun hypnat_of_nat_eq)
-done
+lemma starfun_pow3: "!!R. ( *f* (%r. r ^ n)) R = (R) pow hypnat_of_nat n"
+by (unfold hyperpow_def, transfer, rule refl)
 
 text{*The @{term hypreal_of_hypnat} function as a nonstandard extension of
   @{term real_of_nat} *}
 
-lemma starfunNat_real_of_nat: "( *fNat* real) = hypreal_of_hypnat"
+lemma starfunNat_real_of_nat: "( *f* real) = hypreal_of_hypnat"
+apply (unfold hypreal_of_hypnat_def)
 apply (rule ext)
 apply (rule_tac z = x in eq_Abs_star)
-apply (simp add: hypreal_of_hypnat starfunNat)
+apply (simp add: hypreal_of_hypnat starfun)
 done
 
-lemma starfunNat_inverse_real_of_nat_eq:
+lemma starfun_inverse_real_of_nat_eq:
      "N \<in> HNatInfinite
-   ==> ( *fNat* (%x::nat. inverse(real x))) N = inverse(hypreal_of_hypnat N)"
-apply (rule_tac f1 = inverse in starfun_stafunNat_o2 [THEN subst])
+   ==> ( *f* (%x::nat. inverse(real x))) N = inverse(hypreal_of_hypnat N)"
+apply (rule_tac f1 = inverse in starfun_o2 [THEN subst])
 apply (subgoal_tac "hypreal_of_hypnat N ~= 0")
 apply (simp_all add: HNatInfinite_not_eq_zero starfunNat_real_of_nat starfun_inverse_inverse)
 done
 
-text{*Internal functions - some redundancy with *fNat* now*}
-
-lemma starfunNat_n_congruent:
-      "(%X. starrel``{%n. f n (X n)}) respects starrel"
-by (auto simp add: congruent_def, ultra+)
+text{*Internal functions - some redundancy with *f* now*}
 
-lemma starfunNat_n:
-     "( *fNatn* f) (Abs_star(starrel``{%n. X n})) =
-      Abs_star(starrel `` {%n. f n (X n)})"
-apply (simp add: starfunNat_n_def)
-apply (rule_tac f = Abs_star in arg_cong, auto, ultra)
-done
+lemma starfun_n: "( *fn* f) (star_n X) = star_n (%n. f n (X n))"
+by (simp add: starfun_n_def Ifun_star_n)
 
 text{*Multiplication: @{text "( *fn) x ( *gn) = *(fn x gn)"}*}
 
-lemma starfunNat_n_mult:
-     "( *fNatn* f) z * ( *fNatn* g) z = ( *fNatn* (% i x. f i x * g i x)) z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunNat_n hypreal_mult)
+lemma starfun_n_mult:
+     "( *fn* f) z * ( *fn* g) z = ( *fn* (% i x. f i x * g i x)) z"
+apply (cases z)
+apply (simp add: starfun_n star_n_mult)
 done
 
 text{*Addition: @{text "( *fn) + ( *gn) = *(fn + gn)"}*}
 
-lemma starfunNat_n_add:
-     "( *fNatn* f) z + ( *fNatn* g) z = ( *fNatn* (%i x. f i x + g i x)) z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunNat_n hypreal_add)
+lemma starfun_n_add:
+     "( *fn* f) z + ( *fn* g) z = ( *fn* (%i x. f i x + g i x)) z"
+apply (cases z)
+apply (simp add: starfun_n star_n_add)
 done
 
 text{*Subtraction: @{text "( *fn) - ( *gn) = *(fn + - gn)"}*}
 
-lemma starfunNat_n_add_minus:
-     "( *fNatn* f) z + -( *fNatn* g) z = ( *fNatn* (%i x. f i x + -g i x)) z"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunNat_n hypreal_minus hypreal_add)
+lemma starfun_n_add_minus:
+     "( *fn* f) z + -( *fn* g) z = ( *fn* (%i x. f i x + -g i x)) z"
+apply (cases z)
+apply (simp add: starfun_n star_n_minus star_n_add)
 done
 
 
 text{*Composition: @{text "( *fn) o ( *gn) = *(fn o gn)"}*}
 
-lemma starfunNat_n_const_fun [simp]:
-     "( *fNatn* (%i x. k)) z = hypreal_of_real  k"
-apply (rule_tac z=z in eq_Abs_star)
-apply (simp add: starfunNat_n hypreal_of_real_def star_of_def star_n_def)
+lemma starfun_n_const_fun [simp]:
+     "( *fn* (%i x. k)) z = star_of k"
+apply (cases z)
+apply (simp add: starfun_n star_of_def)
 done
 
-lemma starfunNat_n_minus: "- ( *fNatn* f) x = ( *fNatn* (%i x. - (f i) x)) x"
-apply (rule_tac z=x in eq_Abs_star)
-apply (simp add: starfunNat_n hypreal_minus)
+lemma starfun_n_minus: "- ( *fn* f) x = ( *fn* (%i x. - (f i) x)) x"
+apply (cases x)
+apply (simp add: starfun_n star_n_minus)
 done
 
-lemma starfunNat_n_eq [simp]:
-     "( *fNatn* f) (hypnat_of_nat n) = Abs_star(starrel `` {%i. f i n})"
-by (simp add: starfunNat_n hypnat_of_nat_eq)
+lemma starfun_n_eq [simp]:
+     "( *fn* f) (star_of n) = star_n (%i. f i n)"
+by (simp add: starfun_n star_of_def)
 
-lemma starfun_eq_iff: "(( *fNat* f) = ( *fNat* g)) = (f = g)"
-apply auto
-apply (rule ext, rule ccontr)
-apply (drule_tac x = "hypnat_of_nat (x) " in fun_cong)
-apply (simp add: starfunNat hypnat_of_nat_eq)
-done
-
+lemma starfun_eq_iff: "(( *f* f) = ( *f* g)) = (f = g)"
+by (transfer, rule refl)
 
 lemma starfunNat_inverse_real_of_nat_Infinitesimal [simp]:
-     "N \<in> HNatInfinite ==> ( *fNat* (%x. inverse (real x))) N \<in> Infinitesimal"
-apply (rule_tac f1 = inverse in starfun_stafunNat_o2 [THEN subst])
+     "N \<in> HNatInfinite ==> ( *f* (%x. inverse (real x))) N \<in> Infinitesimal"
+apply (rule_tac f1 = inverse in starfun_o2 [THEN subst])
 apply (subgoal_tac "hypreal_of_hypnat N ~= 0")
 apply (simp_all add: HNatInfinite_not_eq_zero starfunNat_real_of_nat)
 done
 
 ML
 {*
-val starsetNat_def = thm "starsetNat_def";
-
-val NatStar_real_set = thm "NatStar_real_set";
-val NatStar_empty_set = thm "NatStar_empty_set";
-val NatStar_Un = thm "NatStar_Un";
-val starsetNat_n_Un = thm "starsetNat_n_Un";
-val InternalNatSets_Un = thm "InternalNatSets_Un";
-val NatStar_Int = thm "NatStar_Int";
-val starsetNat_n_Int = thm "starsetNat_n_Int";
-val InternalNatSets_Int = thm "InternalNatSets_Int";
-val NatStar_Compl = thm "NatStar_Compl";
-val starsetNat_n_Compl = thm "starsetNat_n_Compl";
-val InternalNatSets_Compl = thm "InternalNatSets_Compl";
-val starsetNat_n_diff = thm "starsetNat_n_diff";
-val InternalNatSets_diff = thm "InternalNatSets_diff";
-val NatStar_subset = thm "NatStar_subset";
-val NatStar_mem = thm "NatStar_mem";
-val NatStar_hypreal_of_real_image_subset = thm "NatStar_hypreal_of_real_image_subset";
+val starset_n_Un = thm "starset_n_Un";
+val InternalSets_Un = thm "InternalSets_Un";
+val starset_n_Int = thm "starset_n_Int";
+val InternalSets_Int = thm "InternalSets_Int";
+val starset_n_Compl = thm "starset_n_Compl";
+val InternalSets_Compl = thm "InternalSets_Compl";
+val starset_n_diff = thm "starset_n_diff";
+val InternalSets_diff = thm "InternalSets_diff";
 val NatStar_SHNat_subset = thm "NatStar_SHNat_subset";
 val NatStar_hypreal_of_real_Int = thm "NatStar_hypreal_of_real_Int";
-val starsetNat_starsetNat_n_eq = thm "starsetNat_starsetNat_n_eq";
-val InternalNatSets_starsetNat_n = thm "InternalNatSets_starsetNat_n";
-val InternalNatSets_UNIV_diff = thm "InternalNatSets_UNIV_diff";
-val starsetNat_n_starsetNat = thm "starsetNat_n_starsetNat";
-val starfunNat_n_starfunNat = thm "starfunNat_n_starfunNat";
-val starfunNat2_n_starfunNat2 = thm "starfunNat2_n_starfunNat2";
-val starfunNat_congruent = thm "starfunNat_congruent";
-val starfunNat = thm "starfunNat";
-val starfunNat2 = thm "starfunNat2";
-val starfunNat_mult = thm "starfunNat_mult";
-val starfunNat2_mult = thm "starfunNat2_mult";
-val starfunNat_add = thm "starfunNat_add";
-val starfunNat2_add = thm "starfunNat2_add";
-val starfunNat2_minus = thm "starfunNat2_minus";
-val starfunNatNat2_o = thm "starfunNatNat2_o";
-val starfunNatNat2_o2 = thm "starfunNatNat2_o2";
-val starfunNat2_o = thm "starfunNat2_o";
-val starfun_stafunNat_o = thm "starfun_stafunNat_o";
-val starfun_stafunNat_o2 = thm "starfun_stafunNat_o2";
-val starfunNat_const_fun = thm "starfunNat_const_fun";
-val starfunNat2_const_fun = thm "starfunNat2_const_fun";
-val starfunNat_minus = thm "starfunNat_minus";
-val starfunNat_inverse = thm "starfunNat_inverse";
-val starfunNat_eq = thm "starfunNat_eq";
-val starfunNat2_eq = thm "starfunNat2_eq";
-val starfunNat_approx = thm "starfunNat_approx";
+val starset_starset_n_eq = thm "starset_starset_n_eq";
+val InternalSets_starset_n = thm "InternalSets_starset_n";
+val InternalSets_UNIV_diff = thm "InternalSets_UNIV_diff";
+val starset_n_starset = thm "starset_n_starset";
+val starfun_const_fun = thm "starfun_const_fun";
 val starfun_le_mono = thm "starfun_le_mono";
 val starfun_less_mono = thm "starfun_less_mono";
-val starfunNat_shift_one = thm "starfunNat_shift_one";
-val starfunNat_rabs = thm "starfunNat_rabs";
-val starfunNat_pow = thm "starfunNat_pow";
-val starfunNat_pow2 = thm "starfunNat_pow2";
+val starfun_shift_one = thm "starfun_shift_one";
+val starfun_abs = thm "starfun_abs";
 val starfun_pow = thm "starfun_pow";
+val starfun_pow2 = thm "starfun_pow2";
+val starfun_pow3 = thm "starfun_pow3";
 val starfunNat_real_of_nat = thm "starfunNat_real_of_nat";
-val starfunNat_inverse_real_of_nat_eq = thm "starfunNat_inverse_real_of_nat_eq";
-val starfunNat_n_congruent = thm "starfunNat_n_congruent";
-val starfunNat_n = thm "starfunNat_n";
-val starfunNat_n_mult = thm "starfunNat_n_mult";
-val starfunNat_n_add = thm "starfunNat_n_add";
-val starfunNat_n_add_minus = thm "starfunNat_n_add_minus";
-val starfunNat_n_const_fun = thm "starfunNat_n_const_fun";
-val starfunNat_n_minus = thm "starfunNat_n_minus";
-val starfunNat_n_eq = thm "starfunNat_n_eq";
+val starfun_inverse_real_of_nat_eq = thm "starfun_inverse_real_of_nat_eq";
+val starfun_n = thm "starfun_n";
+val starfun_n_mult = thm "starfun_n_mult";
+val starfun_n_add = thm "starfun_n_add";
+val starfun_n_add_minus = thm "starfun_n_add_minus";
+val starfun_n_const_fun = thm "starfun_n_const_fun";
+val starfun_n_minus = thm "starfun_n_minus";
+val starfun_n_eq = thm "starfun_n_eq";
 val starfun_eq_iff = thm "starfun_eq_iff";
 val starfunNat_inverse_real_of_nat_Infinitesimal = thm "starfunNat_inverse_real_of_nat_Infinitesimal";
 *}
@@ -536,65 +224,47 @@
 
 subsection{*Nonstandard Characterization of Induction*}
 
-constdefs
-
-  starPNat :: "(nat => bool) => hypnat => bool"  ("*pNat* _" [80] 80)
-  "*pNat* P == (%x. \<exists>X. (X \<in> Rep_star(x) &
-                      {n. P(X n)} \<in> FreeUltrafilterNat))"
+syntax
+  starP :: "('a => bool) => 'a star => bool" ("*p* _" [80] 80)
+  starP2 :: "('a => 'b => bool) => 'a star => 'b star => bool"
+               ("*p2* _" [80] 80)
 
+translations
+  "starP" == "Ipred_of"
+  "starP2" == "Ipred2_of"
 
-  starPNat2 :: "(nat => nat => bool) => hypnat =>hypnat => bool"
-               ("*pNat2* _" [80] 80)
-  "*pNat2* P == (%x y. \<exists>X. \<exists>Y. (X \<in> Rep_star(x) & Y \<in> Rep_star(y) &
-                      {n. P(X n) (Y n)} \<in> FreeUltrafilterNat))"
-
+constdefs
   hSuc :: "hypnat => hypnat"
   "hSuc n == n + 1"
 
+lemma starP: "(( *p* P) (star_n X)) = ({n. P (X n)} \<in> FreeUltrafilterNat)"
+by (simp add: Ipred_of_def star_of_def Ifun_star_n unstar_star_n)
 
-lemma starPNat:
-     "(( *pNat* P) (Abs_star(starrel``{%n. X n}))) =
-      ({n. P (X n)} \<in> FreeUltrafilterNat)"
-by (auto simp add: starPNat_def, ultra)
-
-lemma starPNat_hypnat_of_nat [simp]: "( *pNat* P) (hypnat_of_nat n) = P n"
-by (auto simp add: starPNat hypnat_of_nat_eq)
+lemma starP_star_of [simp]: "( *p* P) (star_of n) = P n"
+by (transfer, rule refl)
 
 lemma hypnat_induct_obj:
-    "(( *pNat* P) 0 &
-            (\<forall>n. ( *pNat* P)(n) --> ( *pNat* P)(n + 1)))
-       --> ( *pNat* P)(n)"
-apply (rule_tac z=n in eq_Abs_star)
-apply (auto simp add: hypnat_zero_def hypnat_one_def starPNat, ultra)
-apply (erule nat_induct)
-apply (drule_tac x = "hypnat_of_nat n" in spec)
-apply (rule ccontr)
-apply (auto simp add: starPNat hypnat_of_nat_eq hypnat_add)
-done
+    "!!n. (( *p* P) (0::hypnat) &
+            (\<forall>n. ( *p* P)(n) --> ( *p* P)(n + 1)))
+       --> ( *p* P)(n)"
+by (transfer, induct_tac n, auto)
 
 lemma hypnat_induct:
-  "[| ( *pNat* P) 0;
-      !!n. ( *pNat* P)(n) ==> ( *pNat* P)(n + 1)|]
-     ==> ( *pNat* P)(n)"
-by (insert hypnat_induct_obj [of P n], auto)
-
-lemma starPNat2:
-"(( *pNat2* P) (Abs_star(starrel``{%n. X n}))
-             (Abs_star(starrel``{%n. Y n}))) =
-      ({n. P (X n) (Y n)} \<in> FreeUltrafilterNat)"
-by (auto simp add: starPNat2_def, ultra)
+  "!!n. [| ( *p* P) (0::hypnat);
+      !!n. ( *p* P)(n) ==> ( *p* P)(n + 1)|]
+     ==> ( *p* P)(n)"
+by (transfer, induct_tac n, auto)
 
-lemma starPNat2_eq_iff: "( *pNat2* (op =)) = (op =)"
-apply (simp add: starPNat2_def)
-apply (rule ext)
-apply (rule ext)
-apply (rule_tac z = x in eq_Abs_star)
-apply (rule_tac z = y in eq_Abs_star)
-apply (auto, ultra)
-done
+lemma starP2:
+"(( *p2* P) (star_n X) (star_n Y)) =
+      ({n. P (X n) (Y n)} \<in> FreeUltrafilterNat)"
+by (simp add: Ipred2_of_def star_of_def Ifun_star_n unstar_star_n)
 
-lemma starPNat2_eq_iff2: "( *pNat2* (%x y. x = y)) X Y = (X = Y)"
-by (simp add: starPNat2_eq_iff)
+lemma starP2_eq_iff: "( *p2* (op =)) = (op =)"
+by (transfer, rule refl)
+
+lemma starP2_eq_iff2: "( *p2* (%x y. x = y)) X Y = (X = Y)"
+by (simp add: starP2_eq_iff)
 
 lemma lemma_hyp: "(\<exists>h. P(h::hypnat)) = (\<exists>x. P(Abs_star(starrel `` {x})))"
 apply auto
@@ -607,34 +277,39 @@
 lemmas zero_not_hSuc = hSuc_not_zero [THEN not_sym, standard, iff]
 
 lemma hSuc_hSuc_eq [iff]: "(hSuc m = hSuc n) = (m = n)"
-by (simp add: hSuc_def hypnat_one_def)
+by (simp add: hSuc_def star_n_one_num)
 
 lemma nonempty_nat_set_Least_mem: "c \<in> (S :: nat set) ==> (LEAST n. n  \<in> S)  \<in> S"
 by (erule LeastI)
 
+lemma nonempty_set_star_has_least:
+    "!!S::nat set star. Iset S \<noteq> {} ==> \<exists>n \<in> Iset S. \<forall>m \<in> Iset S. n \<le> m"
+apply (transfer empty_def)
+apply (rule_tac x="LEAST n. n \<in> S" in bexI)
+apply (simp add: Least_le)
+apply (rule LeastI_ex, auto)
+done
+
 lemma nonempty_InternalNatSet_has_least:
-    "[| S \<in> InternalNatSets; S \<noteq> {} |] ==> \<exists>n \<in> S. \<forall>m \<in> S. n \<le> m"
-apply (auto simp add: InternalNatSets_def starsetNat_n_def lemma_hyp)
-apply (rule_tac z = x in eq_Abs_star)
-apply (auto dest!: bspec simp add: hypnat_le)
-apply (drule_tac x = "%m. LEAST n. n \<in> As m" in spec, auto)
-apply (ultra, force dest: nonempty_nat_set_Least_mem)
-apply (drule_tac x = x in bspec, auto)
-apply (ultra, auto intro: Least_le)
+    "[| (S::hypnat set) \<in> InternalSets; S \<noteq> {} |] ==> \<exists>n \<in> S. \<forall>m \<in> S. n \<le> m"
+apply (clarsimp simp add: InternalSets_def starset_n_def)
+apply (erule nonempty_set_star_has_least)
 done
 
 text{* Goldblatt page 129 Thm 11.3.2*}
+lemma internal_induct_lemma:
+     "!!X::nat set star. [| (0::hypnat) \<in> Iset X; \<forall>n. n \<in> Iset X --> n + 1 \<in> Iset X |]
+      ==> Iset X = (UNIV:: hypnat set)"
+apply (transfer UNIV_def)
+apply (rule equalityI [OF subset_UNIV subsetI])
+apply (induct_tac x, auto)
+done
+
 lemma internal_induct:
-     "[| X \<in> InternalNatSets; 0 \<in> X; \<forall>n. n \<in> X --> n + 1 \<in> X |]
+     "[| X \<in> InternalSets; (0::hypnat) \<in> X; \<forall>n. n \<in> X --> n + 1 \<in> X |]
       ==> X = (UNIV:: hypnat set)"
-apply (rule ccontr)
-apply (frule InternalNatSets_Compl)
-apply (drule_tac S = "- X" in nonempty_InternalNatSet_has_least, auto)
-apply (subgoal_tac "1 \<le> n")
-apply (drule_tac x = "n - 1" in bspec, safe)
-apply (drule_tac x = "n - 1" in spec)
-apply (drule_tac [2] c = 1 and a = n in add_right_mono)
-apply (auto simp add: linorder_not_less [symmetric] iff del: hypnat_neq0_conv)
+apply (clarsimp simp add: InternalSets_def starset_n_def)
+apply (erule (1) internal_induct_lemma)
 done
 
 
--- a/src/HOL/Hyperreal/SEQ.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/SEQ.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -18,7 +18,7 @@
 
   NSLIMSEQ :: "[nat=>real,real] => bool"    ("((_)/ ----NS> (_))" [60, 60] 60)
     --{*Nonstandard definition of convergence of sequence*}
-  "X ----NS> L == (\<forall>N \<in> HNatInfinite. ( *fNat* X) N \<approx> hypreal_of_real L)"
+  "X ----NS> L == (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> hypreal_of_real L)"
 
   lim :: "(nat => real) => real"
     --{*Standard definition of limit using choice operator*}
@@ -42,7 +42,7 @@
 
   NSBseq :: "(nat=>real) => bool"
     --{*Nonstandard definition for bounded sequence*}
-  "NSBseq X == (\<forall>N \<in> HNatInfinite. ( *fNat* X) N : HFinite)"
+  "NSBseq X == (\<forall>N \<in> HNatInfinite. ( *f* X) N : HFinite)"
 
   monoseq :: "(nat=>real)=>bool"
     --{*Definition for monotonicity*}
@@ -59,7 +59,7 @@
   NSCauchy :: "(nat => real) => bool"
     --{*Nonstandard definition*}
   "NSCauchy X == (\<forall>M \<in> HNatInfinite. \<forall>N \<in> HNatInfinite.
-                      ( *fNat* X) M \<approx> ( *fNat* X) N)"
+                      ( *f* X) M \<approx> ( *f* X) N)"
 
 
 
@@ -68,9 +68,10 @@
    the whn'nth term of the hypersequence is a member of Infinitesimal*}
 
 lemma SEQ_Infinitesimal:
-      "( *fNat* (%n::nat. inverse(real(Suc n)))) whn : Infinitesimal"
-apply (simp add: hypnat_omega_def Infinitesimal_FreeUltrafilterNat_iff starfunNat)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl)
+      "( *f* (%n::nat. inverse(real(Suc n)))) whn : Infinitesimal"
+apply (simp add: hypnat_omega_def Infinitesimal_FreeUltrafilterNat_iff starfun)
+apply (simp add: star_n_inverse2)
+apply (rule bexI [OF _ Rep_star_star_n])
 apply (simp add: real_of_nat_Suc_gt_zero FreeUltrafilterNat_inverse_real_of_posnat)
 done
 
@@ -82,7 +83,7 @@
 by (simp add: LIMSEQ_def)
 
 lemma NSLIMSEQ_iff:
-    "(X ----NS> L) = (\<forall>N \<in> HNatInfinite. ( *fNat* X) N \<approx> hypreal_of_real L)"
+    "(X ----NS> L) = (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> hypreal_of_real L)"
 by (simp add: NSLIMSEQ_def)
 
 
@@ -92,12 +93,11 @@
       "X ----> L ==> X ----NS> L"
 apply (simp add: LIMSEQ_def NSLIMSEQ_def)
 apply (auto simp add: HNatInfinite_FreeUltrafilterNat_iff)
-apply (rule_tac z = N in eq_Abs_star)
+apply (rule_tac x = N in star_cases)
 apply (rule approx_minus_iff [THEN iffD2])
-apply (auto simp add: starfunNat mem_infmal_iff [symmetric] hypreal_of_real_def
-              star_of_def star_n_def
-              hypreal_minus hypreal_add Infinitesimal_FreeUltrafilterNat_iff)
-apply (rule bexI [OF _ lemma_starrel_refl], safe)
+apply (auto simp add: starfun mem_infmal_iff [symmetric] star_of_def
+              star_n_minus star_n_add Infinitesimal_FreeUltrafilterNat_iff)
+apply (rule bexI [OF _ Rep_star_star_n], safe)
 apply (drule_tac x = u in spec, safe)
 apply (drule_tac x = no in spec, fuf)
 apply (blast dest: less_imp_le)
@@ -145,9 +145,9 @@
 
 text{* thus, the sequence defines an infinite hypernatural! *}
 lemma HNatInfinite_NSLIMSEQ: "\<forall>n. n \<le> f n
-          ==> Abs_star (starrel `` {f}) : HNatInfinite"
+          ==> star_n f : HNatInfinite"
 apply (auto simp add: HNatInfinite_FreeUltrafilterNat_iff)
-apply (rule bexI [OF _ lemma_starrel_refl], safe)
+apply (rule bexI [OF _ Rep_star_star_n], safe)
 apply (erule FreeUltrafilterNat_NSLIMSEQ)
 done
 
@@ -161,9 +161,9 @@
 by auto
 
 lemma lemmaLIM3: "[| 0 < r; \<forall>n. r \<le> \<bar>X (f n) + - L\<bar>;
-           ( *fNat* X) (Abs_star (starrel `` {f})) +
+           ( *f* X) (star_n f) +
            - hypreal_of_real  L \<approx> 0 |] ==> False"
-apply (auto simp add: starfunNat mem_infmal_iff [symmetric] hypreal_of_real_def star_of_def star_n_def hypreal_minus hypreal_add Infinitesimal_FreeUltrafilterNat_iff)
+apply (auto simp add: starfun mem_infmal_iff [symmetric] star_of_def star_n_minus star_n_add Infinitesimal_FreeUltrafilterNat_iff)
 apply (rename_tac "Y")
 apply (drule_tac x = r in spec, safe)
 apply (drule FreeUltrafilterNat_Int, assumption)
@@ -179,7 +179,7 @@
 apply (rule ccontr, simp, safe)
 txt{* skolemization step *}
 apply (drule choice, safe)
-apply (drule_tac x = "Abs_star (starrel``{f}) " in bspec)
+apply (drule_tac x = "star_n f" in bspec)
 apply (drule_tac [2] approx_minus_iff [THEN iffD1])
 apply (simp_all add: linorder_not_less)
 apply (blast intro: HNatInfinite_NSLIMSEQ)
@@ -201,7 +201,7 @@
 
 lemma NSLIMSEQ_add:
       "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n + Y n) ----NS> a + b"
-by (auto intro: approx_add simp add: NSLIMSEQ_def starfunNat_add [symmetric])
+by (auto intro: approx_add simp add: NSLIMSEQ_def starfun_add [symmetric])
 
 lemma LIMSEQ_add: "[| X ----> a; Y ----> b |] ==> (%n. X n + Y n) ----> a + b"
 by (simp add: LIMSEQ_NSLIMSEQ_iff NSLIMSEQ_add)
@@ -218,13 +218,13 @@
 lemma NSLIMSEQ_mult:
       "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n * Y n) ----NS> a * b"
 by (auto intro!: approx_mult_HFinite 
-        simp add: NSLIMSEQ_def starfunNat_mult [symmetric])
+        simp add: NSLIMSEQ_def starfun_mult [symmetric])
 
 lemma LIMSEQ_mult: "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b"
 by (simp add: LIMSEQ_NSLIMSEQ_iff NSLIMSEQ_mult)
 
 lemma NSLIMSEQ_minus: "X ----NS> a ==> (%n. -(X n)) ----NS> -a"
-by (auto simp add: NSLIMSEQ_def starfunNat_minus [symmetric])
+by (auto simp add: NSLIMSEQ_def starfun_minus [symmetric])
 
 lemma LIMSEQ_minus: "X ----> a ==> (%n. -(X n)) ----> -a"
 by (simp add: LIMSEQ_NSLIMSEQ_iff NSLIMSEQ_minus)
@@ -266,7 +266,7 @@
 text{*Proof is like that of @{text NSLIM_inverse}.*}
 lemma NSLIMSEQ_inverse:
      "[| X ----NS> a;  a ~= 0 |] ==> (%n. inverse(X n)) ----NS> inverse(a)"
-by (simp add: NSLIMSEQ_def starfunNat_inverse [symmetric] 
+by (simp add: NSLIMSEQ_def starfun_inverse [symmetric] 
               hypreal_of_real_approx_inverse)
 
 
@@ -491,10 +491,10 @@
 lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. \<bar>X n\<bar> < real(Suc N))"
 by (simp add: Bseq_def lemma_NBseq_def2)
 
-lemma NSBseqD: "[| NSBseq X;  N: HNatInfinite |] ==> ( *fNat* X) N : HFinite"
+lemma NSBseqD: "[| NSBseq X;  N: HNatInfinite |] ==> ( *f* X) N : HFinite"
 by (simp add: NSBseq_def)
 
-lemma NSBseqI: "\<forall>N \<in> HNatInfinite. ( *fNat* X) N : HFinite ==> NSBseq X"
+lemma NSBseqI: "\<forall>N \<in> HNatInfinite. ( *f* X) N : HFinite ==> NSBseq X"
 by (simp add: NSBseq_def)
 
 text{*The standard definition implies the nonstandard definition*}
@@ -504,10 +504,10 @@
 
 lemma Bseq_NSBseq: "Bseq X ==> NSBseq X"
 apply (simp add: Bseq_def NSBseq_def, safe)
-apply (rule_tac z = N in eq_Abs_star)
-apply (auto simp add: starfunNat HFinite_FreeUltrafilterNat_iff 
+apply (rule_tac x = N in star_cases)
+apply (auto simp add: starfun HFinite_FreeUltrafilterNat_iff 
                       HNatInfinite_FreeUltrafilterNat_iff)
-apply (rule bexI [OF _ lemma_starrel_refl]) 
+apply (rule bexI [OF _ Rep_star_star_n]) 
 apply (drule_tac f = Xa in lemma_Bseq)
 apply (rule_tac x = "K+1" in exI)
 apply (drule_tac P="%n. ?f n \<le> K" in FreeUltrafilterNat_all, ultra)
@@ -538,9 +538,9 @@
 
 lemma real_seq_to_hypreal_HInfinite:
      "\<forall>N. real(Suc N) < \<bar>X (f N)\<bar>
-      ==>  Abs_star(starrel``{X o f}) : HInfinite"
+      ==>  star_n (X o f) : HInfinite"
 apply (auto simp add: HInfinite_FreeUltrafilterNat_iff o_def)
-apply (rule bexI [OF _ lemma_starrel_refl], clarify)  
+apply (rule bexI [OF _ Rep_star_star_n], clarify)  
 apply (cut_tac u = u in FreeUltrafilterNat_nat_gt_real)
 apply (drule FreeUltrafilterNat_all)
 apply (erule FreeUltrafilterNat_Int [THEN FreeUltrafilterNat_subset])
@@ -567,9 +567,9 @@
 
 lemma HNatInfinite_skolem_f:
      "\<forall>N. real(Suc N) < \<bar>X (f N)\<bar>
-      ==> Abs_star(starrel``{f}) : HNatInfinite"
+      ==> star_n f : HNatInfinite"
 apply (auto simp add: HNatInfinite_FreeUltrafilterNat_iff)
-apply (rule bexI [OF _ lemma_starrel_refl], safe)
+apply (rule bexI [OF _ Rep_star_star_n], safe)
 apply (rule ccontr, drule FreeUltrafilterNat_Compl_mem)
 apply (rule lemma_finite_NSBseq2 [THEN FreeUltrafilterNat_finite, THEN notE]) 
 apply (subgoal_tac "{n. f n \<le> u & real (Suc n) < \<bar>X (f n)\<bar>} =
@@ -585,7 +585,7 @@
 apply (drule lemmaNSBseq2, safe)
 apply (frule_tac X = X and f = f in real_seq_to_hypreal_HInfinite)
 apply (drule HNatInfinite_skolem_f [THEN [2] bspec])
-apply (auto simp add: starfunNat o_def HFinite_HInfinite_iff)
+apply (auto simp add: starfun o_def HFinite_HInfinite_iff)
 done
 
 text{* Equivalence of nonstandard and standard definitions
@@ -762,7 +762,7 @@
 subsubsection{*Standard Implies Nonstandard*}
 
 lemma lemmaCauchy1:
-     "Abs_star (starrel `` {x}) : HNatInfinite
+     "star_n x : HNatInfinite
       ==> {n. M \<le> x n} : FreeUltrafilterNat"
 apply (auto simp add: HNatInfinite_FreeUltrafilterNat_iff)
 apply (drule_tac x = M in spec, ultra)
@@ -776,16 +776,16 @@
 
 lemma Cauchy_NSCauchy: "Cauchy X ==> NSCauchy X"
 apply (simp add: Cauchy_def NSCauchy_def, safe)
-apply (rule_tac z = M in eq_Abs_star)
-apply (rule_tac z = N in eq_Abs_star)
+apply (rule_tac x = M in star_cases)
+apply (rule_tac x = N in star_cases)
 apply (rule approx_minus_iff [THEN iffD2])
 apply (rule mem_infmal_iff [THEN iffD1])
-apply (auto simp add: starfunNat hypreal_minus hypreal_add Infinitesimal_FreeUltrafilterNat_iff)
+apply (auto simp add: starfun star_n_minus star_n_add Infinitesimal_FreeUltrafilterNat_iff)
 apply (rule bexI, auto)
 apply (drule spec, auto)
 apply (drule_tac M = M in lemmaCauchy1)
 apply (drule_tac M = M in lemmaCauchy1)
-apply (rule_tac x1 = xa in lemmaCauchy2 [THEN [2] FreeUltrafilterNat_subset])
+apply (rule_tac x1 = Xaa in lemmaCauchy2 [THEN [2] FreeUltrafilterNat_subset])
 apply (rule FreeUltrafilterNat_Int)
 apply (auto intro: FreeUltrafilterNat_Int)
 done
@@ -797,11 +797,11 @@
 apply (rule ccontr, simp)
 apply (auto dest!: choice HNatInfinite_NSLIMSEQ simp add: all_conj_distrib)  
 apply (drule bspec, assumption)
-apply (drule_tac x = "Abs_star (starrel `` {fa}) " in bspec); 
-apply (auto simp add: starfunNat)
+apply (drule_tac x = "star_n fa" in bspec); 
+apply (auto simp add: starfun)
 apply (drule approx_minus_iff [THEN iffD1])
 apply (drule mem_infmal_iff [THEN iffD2])
-apply (auto simp add: hypreal_minus hypreal_add Infinitesimal_FreeUltrafilterNat_iff)
+apply (auto simp add: star_n_minus star_n_add Infinitesimal_FreeUltrafilterNat_iff)
 apply (rename_tac "Y")
 apply (drule_tac x = e in spec, auto)
 apply (drule FreeUltrafilterNat_Int, assumption)
@@ -971,7 +971,7 @@
 
 lemma NSLIMSEQ_Suc: "f ----NS> l ==> (%n. f(Suc n)) ----NS> l"
 apply (frule NSconvergentI [THEN NSCauchy_NSconvergent_iff [THEN iffD2]])
-apply (auto simp add: NSCauchy_def NSLIMSEQ_def starfunNat_shift_one)
+apply (auto simp add: NSCauchy_def NSLIMSEQ_def starfun_shift_one)
 apply (drule bspec, assumption)
 apply (drule bspec, assumption)
 apply (drule Nats_1 [THEN [2] HNatInfinite_SHNat_add])
@@ -984,7 +984,7 @@
 
 lemma NSLIMSEQ_imp_Suc: "(%n. f(Suc n)) ----NS> l ==> f ----NS> l"
 apply (frule NSconvergentI [THEN NSCauchy_NSconvergent_iff [THEN iffD2]])
-apply (auto simp add: NSCauchy_def NSLIMSEQ_def starfunNat_shift_one)
+apply (auto simp add: NSCauchy_def NSLIMSEQ_def starfun_shift_one)
 apply (drule bspec, assumption)
 apply (drule bspec, assumption)
 apply (frule Nats_1 [THEN [2] HNatInfinite_SHNat_diff])
@@ -1016,7 +1016,7 @@
 lemma NSLIMSEQ_imp_rabs: "f ----NS> l ==> (%n. \<bar>f n\<bar>) ----NS> \<bar>l\<bar>"
 apply (simp add: NSLIMSEQ_def)
 apply (auto intro: approx_hrabs 
-            simp add: starfunNat_rabs hypreal_of_real_hrabs [symmetric])
+            simp add: starfun_abs hypreal_of_real_hrabs [symmetric])
 done
 
 text{* standard version *}
@@ -1123,7 +1123,7 @@
 apply (simp add: NSLIMSEQ_def)
 apply (auto dest!: convergent_realpow simp add: convergent_NSconvergent_iff)
 apply (frule NSconvergentD)
-apply (auto simp add: NSLIMSEQ_def NSCauchy_NSconvergent_iff [symmetric] NSCauchy_def starfunNat_pow)
+apply (auto simp add: NSLIMSEQ_def NSCauchy_NSconvergent_iff [symmetric] NSCauchy_def starfun_pow)
 apply (frule HNatInfinite_add_one)
 apply (drule bspec, assumption)
 apply (drule bspec, assumption)
@@ -1131,7 +1131,7 @@
 apply (simp add: hyperpow_add)
 apply (drule approx_mult_subst_SReal, assumption)
 apply (drule approx_trans3, assumption)
-apply (auto simp del: hypreal_of_real_mult simp add: hypreal_of_real_mult [symmetric])
+apply (auto simp del: star_of_mult simp add: star_of_mult [symmetric])
 done
 
 text{* standard version *}
@@ -1164,7 +1164,7 @@
 subsection{*Hyperreals and Sequences*}
 
 text{*A bounded sequence is a finite hyperreal*}
-lemma NSBseq_HFinite_hypreal: "NSBseq X ==> Abs_star(starrel``{X}) : HFinite"
+lemma NSBseq_HFinite_hypreal: "NSBseq X ==> star_n X : HFinite"
 by (auto intro!: bexI lemma_starrel_refl 
             intro: FreeUltrafilterNat_all [THEN FreeUltrafilterNat_subset]
             simp add: HFinite_FreeUltrafilterNat_iff Bseq_NSBseq_iff [symmetric]
@@ -1172,11 +1172,11 @@
 
 text{*A sequence converging to zero defines an infinitesimal*}
 lemma NSLIMSEQ_zero_Infinitesimal_hypreal:
-      "X ----NS> 0 ==> Abs_star(starrel``{X}) : Infinitesimal"
+      "X ----NS> 0 ==> star_n X : Infinitesimal"
 apply (simp add: NSLIMSEQ_def)
 apply (drule_tac x = whn in bspec)
 apply (simp add: HNatInfinite_whn)
-apply (auto simp add: hypnat_omega_def mem_infmal_iff [symmetric] starfunNat)
+apply (auto simp add: hypnat_omega_def mem_infmal_iff [symmetric] starfun)
 done
 
 (***---------------------------------------------------------------
--- a/src/HOL/Hyperreal/Star.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/Star.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -10,14 +10,17 @@
 imports NSA
 begin
 
+(* nonstandard extension of sets *)
+syntax starset :: "'a set => 'a star set" ("*s* _" [80] 80)
+translations "starset" == "Iset_of"
+
+syntax starfun :: "('a => 'b) => 'a star => 'b star" ("*f* _" [80] 80)
+translations "starfun" == "Ifun_of"
+
 constdefs
-    (* nonstandard extension of sets *)
-    starset :: "'a set => 'a star set"          ("*s* _" [80] 80)
-    "*s* A  == {x. \<forall>X \<in> Rep_star(x). {n::nat. X n  \<in> A}: FreeUltrafilterNat}"
-
     (* internal sets *)
     starset_n :: "(nat => 'a set) => 'a star set"        ("*sn* _" [80] 80)
-    "*sn* As  == {x. \<forall>X \<in> Rep_star(x). {n::nat. X n : (As n)}: FreeUltrafilterNat}"
+    "*sn* As == Iset (star_n As)"
 
     InternalSets :: "'a star set set"
     "InternalSets == {X. \<exists>As. X = *sn* As}"
@@ -26,14 +29,10 @@
     is_starext  :: "['a star => 'a star, 'a => 'a] => bool"
     "is_starext F f == (\<forall>x y. \<exists>X \<in> Rep_star(x). \<exists>Y \<in> Rep_star(y).
                         ((y = (F x)) = ({n. Y n = f(X n)} : FreeUltrafilterNat)))"
-
-    starfun :: "('a => 'b) => 'a star => 'b star"       ("*f* _" [80] 80)
-    "*f* f  == (%x. Abs_star(\<Union>X \<in> Rep_star(x). starrel``{%n. f(X n)}))"
-
     (* internal functions *)
     starfun_n :: "(nat => ('a => 'b)) => 'a star => 'b star"
                  ("*fn* _" [80] 80)
-    "*fn* F  == (%x. Abs_star(\<Union>X \<in> Rep_star(x). starrel``{%n. (F n)(X n)}))"
+    "*fn* F == Ifun (star_n F)"
 
     InternalFuns :: "('a star => 'b star) set"
     "InternalFuns == {X. \<exists>F. X = *fn* F}"
@@ -55,67 +54,41 @@
 
 subsection{*Properties of the Star-transform Applied to Sets of Reals*}
 
-lemma STAR_real_set: "*s*(UNIV::'a set) = (UNIV::'a star set)"
-by (simp add: starset_def)
-declare STAR_real_set [simp]
+lemma STAR_UNIV_set [simp]: "*s*(UNIV::'a set) = (UNIV::'a star set)"
+by (transfer UNIV_def, rule refl)
 
-lemma STAR_empty_set: "*s* {} = {}"
-by (simp add: starset_def)
-declare STAR_empty_set [simp]
+lemma STAR_empty_set [simp]: "*s* {} = {}"
+by (transfer empty_def, rule refl)
 
 lemma STAR_Un: "*s* (A Un B) = *s* A Un *s* B"
-apply (auto simp add: starset_def)
-  prefer 3 apply (blast intro: FreeUltrafilterNat_subset)
- prefer 2 apply (blast intro: FreeUltrafilterNat_subset)
-apply (drule FreeUltrafilterNat_Compl_mem)
-apply (drule bspec, assumption)
-apply (rule_tac z = x in eq_Abs_star, auto, ultra)
-done
+by (transfer Un_def, rule refl)
 
 lemma STAR_Int: "*s* (A Int B) = *s* A Int *s* B"
-apply (simp add: starset_def, auto)
-prefer 3 apply (blast intro: FreeUltrafilterNat_Int FreeUltrafilterNat_subset)
-apply (blast intro: FreeUltrafilterNat_subset)+
-done
+by (transfer Int_def, rule refl)
 
 lemma STAR_Compl: "*s* -A = -( *s* A)"
-apply (auto simp add: starset_def)
-apply (rule_tac [!] z = x in eq_Abs_star)
-apply (auto dest!: bspec, ultra)
-apply (drule FreeUltrafilterNat_Compl_mem, ultra)
-done
+by (transfer Compl_def, rule refl)
 
-lemma STAR_mem_Compl: "x \<notin> *s* F ==> x : *s* (- F)"
-by (auto simp add: STAR_Compl)
+lemma STAR_mem_Compl: "!!x. x \<notin> *s* F ==> x : *s* (- F)"
+by (transfer Compl_def, simp)
 
 lemma STAR_diff: "*s* (A - B) = *s* A - *s* B"
-by (auto simp add: Diff_eq STAR_Int STAR_Compl)
+by (transfer set_diff_def, rule refl)
 
 lemma STAR_subset: "A <= B ==> *s* A <= *s* B"
-apply (simp add: starset_def)
-apply (blast intro: FreeUltrafilterNat_subset)+
-done
+by (transfer subset_def, simp)
+
+lemma STAR_mem: "a \<in> A ==> star_of a : *s* A"
+by transfer
 
-lemma STAR_mem: "a  \<in> A ==> hypreal_of_real a : *s* A"
-apply (simp add: starset_def hypreal_of_real_def star_of_def star_n_def)
-apply (auto intro: FreeUltrafilterNat_subset)
-done
+lemma STAR_mem_iff: "(star_of x \<in> *s* A) = (x \<in> A)"
+by (transfer, rule refl)
 
-lemma STAR_hypreal_of_real_image_subset: "hypreal_of_real ` A <= *s* A"
-apply (simp add: starset_def)
-apply (auto simp add: hypreal_of_real_def star_of_def star_n_def)
-apply (blast intro: FreeUltrafilterNat_subset)
-done
+lemma STAR_star_of_image_subset: "star_of ` A <= *s* A"
+by (auto simp add: STAR_mem)
 
 lemma STAR_hypreal_of_real_Int: "*s* X Int Reals = hypreal_of_real ` X"
-apply (simp add: starset_def)
-apply (auto simp add: hypreal_of_real_def star_of_def star_n_def SReal_def)
-apply (fold star_n_def star_of_def hypreal_of_real_def)
-apply (rule imageI, rule ccontr)
-apply (drule bspec)
-apply (rule lemma_starrel_refl)
-prefer 2 apply (blast intro: FreeUltrafilterNat_subset, auto)
-done
+by (auto simp add: SReal_def STAR_mem_iff)
 
 lemma lemma_not_hyprealA: "x \<notin> hypreal_of_real ` A ==> \<forall>y \<in> A. x \<noteq> hypreal_of_real y"
 by auto
@@ -124,24 +97,19 @@
 by auto
 
 lemma STAR_real_seq_to_hypreal:
-    "\<forall>n. (X n) \<notin> M
-          ==> Abs_star(starrel``{X}) \<notin> *s* M"
-apply (simp add: starset_def)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl, auto)
+    "\<forall>n. (X n) \<notin> M ==> star_n X \<notin> *s* M"
+apply (unfold Iset_of_def star_of_def)
+apply (simp add: Iset_star_n)
 done
 
-lemma STAR_singleton: "*s* {x} = {hypreal_of_real x}"
-apply (simp add: starset_def)
-apply (auto simp add: hypreal_of_real_def star_of_def star_n_def)
-apply (rule_tac z = xa in eq_Abs_star)
-apply (auto intro: FreeUltrafilterNat_subset)
-done
-declare STAR_singleton [simp]
+lemma STAR_insert [simp]: "*s* (insert x A) = insert (star_of x) ( *s* A)"
+by (transfer insert_def Un_def, rule refl)
 
-lemma STAR_not_mem: "x \<notin> F ==> hypreal_of_real x \<notin> *s* F"
-apply (auto simp add: starset_def hypreal_of_real_def star_of_def star_n_def)
-apply (rule bexI, rule_tac [2] lemma_starrel_refl, auto)
-done
+lemma STAR_singleton: "*s* {x} = {star_of x}"
+by simp
+
+lemma STAR_not_mem: "x \<notin> F ==> star_of x \<notin> *s* F"
+by transfer
 
 lemma STAR_subset_closed: "[| x : *s* A; A <= B |] ==> x : *s* B"
 by (blast dest: STAR_subset)
@@ -150,7 +118,9 @@
    sequence) as a special case of an internal set*}
 
 lemma starset_n_starset: "\<forall>n. (As n = A) ==> *sn* As = *s* A"
-by (simp add: starset_n_def starset_def)
+apply (drule expand_fun_eq [THEN iffD2])
+apply (simp add: starset_n_def Iset_of_def star_of_def)
+done
 
 
 (*----------------------------------------------------------------*)
@@ -163,7 +133,9 @@
 (*----------------------------------------------------------------*)
 
 lemma starfun_n_starfun: "\<forall>n. (F n = f) ==> *fn* F = *f* f"
-by (simp add: starfun_n_def starfun_def)
+apply (drule expand_fun_eq [THEN iffD2])
+apply (simp add: starfun_n_def Ifun_of_def star_of_def)
+done
 
 
 (*
@@ -194,61 +166,45 @@
 
 text{*Nonstandard extension of functions*}
 
-lemma starfun_congruent: "(%X. starrel``{%n. f (X n)}) respects starrel"
-by (simp add: congruent_def, auto, ultra)
-
 lemma starfun:
-      "( *f* f) (Abs_star(starrel``{%n. X n})) =
-       Abs_star(starrel `` {%n. f (X n)})"
-apply (simp add: starfun_def)
-apply (rule_tac f = Abs_star in arg_cong)
-apply (simp add: starrel_in_hypreal [THEN Abs_star_inverse] 
-                 UN_equiv_class [OF equiv_starrel starfun_congruent])
-done
+      "( *f* f) (star_n X) = star_n (%n. f (X n))"
+by (simp add: Ifun_of_def star_of_def Ifun_star_n)
 
 lemma starfun_if_eq:
      "w \<noteq> hypreal_of_real x
-       ==> ( *f* (\<lambda>z. if z = x then a else g z)) w = ( *f* g) w" 
-apply (rule_tac z = w in eq_Abs_star)
-apply (simp add: hypreal_of_real_def star_of_def star_n_def starfun, ultra)
+       ==> ( *f* (\<lambda>z. if z = x then a else g z)) w = ( *f* g) w"
+apply (cases w)
+apply (simp add: star_of_def starfun star_n_eq_iff, ultra)
 done
 
 (*-------------------------------------------
   multiplication: ( *f) x ( *g) = *(f x g)
  ------------------------------------------*)
-lemma starfun_mult: "( *f* f) xa * ( *f* g) xa = ( *f* (%x. f x * g x)) xa"
-by (rule_tac z = xa in eq_Abs_star, simp add: starfun hypreal_mult)
-
+lemma starfun_mult: "!!x. ( *f* f) x * ( *f* g) x = ( *f* (%x. f x * g x)) x"
+by (transfer, rule refl)
 declare starfun_mult [symmetric, simp]
 
 (*---------------------------------------
   addition: ( *f) + ( *g) = *(f + g)
  ---------------------------------------*)
-lemma starfun_add: "( *f* f) xa + ( *f* g) xa = ( *f* (%x. f x + g x)) xa"
-by (rule_tac z = xa in eq_Abs_star, simp add: starfun hypreal_add)
+lemma starfun_add: "!!x. ( *f* f) x + ( *f* g) x = ( *f* (%x. f x + g x)) x"
+by (transfer, rule refl)
 declare starfun_add [symmetric, simp]
 
 (*--------------------------------------------
   subtraction: ( *f) + -( *g) = *(f + -g)
  -------------------------------------------*)
-
-lemma starfun_minus: "- ( *f* f) x = ( *f* (%x. - f x)) x"
-apply (rule_tac z = x in eq_Abs_star)
-apply (auto simp add: starfun hypreal_minus)
-done
+lemma starfun_minus: "!!x. - ( *f* f) x = ( *f* (%x. - f x)) x"
+by (transfer, rule refl)
 declare starfun_minus [symmetric, simp]
 
 (*FIXME: delete*)
-lemma starfun_add_minus: "( *f* f) xa + -( *f* g) xa = ( *f* (%x. f x + -g x)) xa"
-apply (simp (no_asm))
-done
+lemma starfun_add_minus: "!!x. ( *f* f) x + -( *f* g) x = ( *f* (%x. f x + -g x)) x"
+by (transfer, rule refl)
 declare starfun_add_minus [symmetric, simp]
 
-lemma starfun_diff:
-  "( *f* f) xa  - ( *f* g) xa = ( *f* (%x. f x - g x)) xa"
-apply (rule_tac z = xa in eq_Abs_star)
-apply (simp add: starfun hypreal_diff)
-done
+lemma starfun_diff: "!!x. ( *f* f) x  - ( *f* g) x = ( *f* (%x. f x - g x)) x"
+by (transfer, rule refl)
 declare starfun_diff [symmetric, simp]
 
 (*--------------------------------------
@@ -256,44 +212,33 @@
  ---------------------------------------*)
 
 lemma starfun_o2: "(%x. ( *f* f) (( *f* g) x)) = *f* (%x. f (g x))"
-apply (rule ext)
-apply (rule_tac z = x in eq_Abs_star)
-apply (auto simp add: starfun)
-done
+by (transfer, rule refl)
 
 lemma starfun_o: "( *f* f) o ( *f* g) = ( *f* (f o g))"
-apply (simp add: o_def)
-apply (simp (no_asm) add: starfun_o2)
-done
+by (transfer o_def, rule refl)
 
 text{*NS extension of constant function*}
-lemma starfun_const_fun: "( *f* (%x. k)) xa = hypreal_of_real  k"
-apply (rule_tac z = xa in eq_Abs_star)
-apply (auto simp add: starfun hypreal_of_real_def star_of_def star_n_def)
-done
-
-declare starfun_const_fun [simp]
+lemma starfun_const_fun [simp]: "!!x. ( *f* (%x. k)) x = star_of k"
+by (transfer, rule refl)
 
 text{*the NS extension of the identity function*}
 
-lemma starfun_Idfun_approx: "x @= hypreal_of_real a ==> ( *f* (%x. x)) x @= hypreal_of_real  a"
-apply (rule_tac z = x in eq_Abs_star)
-apply (auto simp add: starfun)
-done
+lemma starfun_Id [simp]: "!!x. ( *f* (%x. x)) x = x"
+by (transfer, rule refl)
 
-lemma starfun_Id: "( *f* (%x. x)) x = x"
-apply (rule_tac z = x in eq_Abs_star)
-apply (auto simp add: starfun)
-done
-declare starfun_Id [simp]
+(* this is trivial, given starfun_Id *)
+lemma starfun_Idfun_approx:
+  "x @= hypreal_of_real a ==> ( *f* (%x. x)) x @= hypreal_of_real a"
+by (simp only: starfun_Id)
 
 text{*The Star-function is a (nonstandard) extension of the function*}
 
 lemma is_starext_starfun: "is_starext ( *f* f) f"
 apply (simp add: is_starext_def, auto)
-apply (rule_tac z = x in eq_Abs_star)
-apply (rule_tac z = y in eq_Abs_star)
-apply (auto intro!: bexI simp add: starfun)
+apply (rule_tac x = x in star_cases)
+apply (rule_tac x = y in star_cases)
+apply (auto intro!: bexI [OF _ Rep_star_star_n]
+            simp add: starfun star_n_eq_iff)
 done
 
 text{*Any nonstandard extension is in fact the Star-function*}
@@ -301,7 +246,7 @@
 lemma is_starfun_starext: "is_starext F f ==> F = *f* f"
 apply (simp add: is_starext_def)
 apply (rule ext)
-apply (rule_tac z = x in eq_Abs_star)
+apply (rule_tac x = x in star_cases)
 apply (drule_tac x = x in spec)
 apply (drule_tac x = "( *f* f) x" in spec)
 apply (auto dest!: FreeUltrafilterNat_Compl_mem simp add: starfun, ultra)
@@ -313,36 +258,31 @@
 text{*extented function has same solution as its standard
    version for real arguments. i.e they are the same
    for all real arguments*}
-lemma starfun_eq: "( *f* f) (hypreal_of_real a) = hypreal_of_real (f a)"
-by (auto simp add: starfun hypreal_of_real_def star_of_def star_n_def)
+lemma starfun_eq [simp]: "( *f* f) (star_of a) = star_of (f a)"
+by (transfer, rule refl)
 
-declare starfun_eq [simp]
-
-lemma starfun_approx: "( *f* f) (hypreal_of_real a) @= hypreal_of_real (f a)"
-by auto
+lemma starfun_approx: "( *f* f) (star_of a) @= hypreal_of_real (f a)"
+by simp
 
 (* useful for NS definition of derivatives *)
-lemma starfun_lambda_cancel: "( *f* (%h. f (x + h))) xa  = ( *f* f) (hypreal_of_real  x + xa)"
-apply (rule_tac z = xa in eq_Abs_star)
-apply (auto simp add: starfun hypreal_of_real_def star_of_def star_n_def hypreal_add)
-done
+lemma starfun_lambda_cancel:
+  "!!x'. ( *f* (%h. f (x + h))) x'  = ( *f* f) (star_of x + x')"
+by (transfer, rule refl)
 
-lemma starfun_lambda_cancel2: "( *f* (%h. f(g(x + h)))) xa = ( *f* (f o g)) (hypreal_of_real x + xa)"
-apply (rule_tac z = xa in eq_Abs_star)
-apply (auto simp add: starfun hypreal_of_real_def star_of_def star_n_def hypreal_add)
-done
+lemma starfun_lambda_cancel2:
+  "( *f* (%h. f(g(x + h)))) x' = ( *f* (f o g)) (star_of x + x')"
+by (unfold o_def, rule starfun_lambda_cancel)
 
-lemma starfun_mult_HFinite_approx: "[| ( *f* f) xa @= l; ( *f* g) xa @= m;
+lemma starfun_mult_HFinite_approx: "[| ( *f* f) x @= l; ( *f* g) x @= m;
                   l: HFinite; m: HFinite
-               |] ==>  ( *f* (%x. f x * g x)) xa @= l * m"
-apply (drule approx_mult_HFinite, assumption+)
+               |] ==>  ( *f* (%x. f x * g x)) x @= l * m"
+apply (drule (3) approx_mult_HFinite)
 apply (auto intro: approx_HFinite [OF _ approx_sym])
 done
 
-lemma starfun_add_approx: "[| ( *f* f) xa @= l; ( *f* g) xa @= m
-               |] ==>  ( *f* (%x. f x + g x)) xa @= l + m"
-apply (auto intro: approx_add)
-done
+lemma starfun_add_approx: "[| ( *f* f) x @= l; ( *f* g) x @= m
+               |] ==>  ( *f* (%x. f x + g x)) x @= l + m"
+by (auto intro: approx_add)
 
 text{*Examples: hrabs is nonstandard extension of rabs
               inverse is nonstandard extension of inverse*}
@@ -354,53 +294,32 @@
 lemma starfun_rabs_hrabs: "*f* abs = abs"
 by (rule hrabs_is_starext_rabs [THEN is_starext_starfun_iff [THEN iffD1], symmetric])
 
-lemma starfun_inverse_inverse: "( *f* inverse) x = inverse(x)"
-apply (rule_tac z = x in eq_Abs_star)
-apply (simp add: starfun)
-apply (fold star_n_def)
-apply (unfold star_inverse_def Ifun_of_def star_of_def)
-apply (simp add: Ifun_star_n)
-done
-declare starfun_inverse_inverse [simp]
+lemma starfun_inverse_inverse [simp]: "( *f* inverse) x = inverse(x)"
+by (unfold star_inverse_def, rule refl)
 
-lemma starfun_inverse: "inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x"
-apply (rule_tac z = x in eq_Abs_star)
-apply (auto simp add: starfun hypreal_inverse2)
-done
+lemma starfun_inverse: "!!x. inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x"
+by (transfer, rule refl)
 declare starfun_inverse [symmetric, simp]
 
-lemma starfun_divide: "( *f* f) xa  / ( *f* g) xa = ( *f* (%x. f x / g x)) xa"
-apply (rule_tac z = xa in eq_Abs_star)
-apply (simp add: starfun)
-apply (fold star_n_def)
-apply (unfold star_divide_def Ifun2_of_def star_of_def)
-apply (simp add: Ifun_star_n)
-done
+lemma starfun_divide: "!!x. ( *f* f) x / ( *f* g) x = ( *f* (%x. f x / g x)) x"
+by (transfer, rule refl)
 declare starfun_divide [symmetric, simp]
 
-lemma starfun_inverse2: "inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x"
-apply (rule_tac z = x in eq_Abs_star)
-apply (auto intro: FreeUltrafilterNat_subset dest!: FreeUltrafilterNat_Compl_mem simp add: starfun hypreal_inverse hypreal_zero_def)
-done
+lemma starfun_inverse2: "!!x. inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x"
+by (transfer, rule refl)
 
 text{*General lemma/theorem needed for proofs in elementary
     topology of the reals*}
 lemma starfun_mem_starset:
-      "( *f* f) x : *s* A ==> x : *s* {x. f x  \<in> A}"
-apply (simp add: starset_def)
-apply (rule_tac z = x in eq_Abs_star)
-apply (auto simp add: starfun)
-apply (rename_tac "X")
-apply (drule_tac x = "%n. f (X n) " in bspec)
-apply (auto, ultra)
-done
+      "!!x. ( *f* f) x : *s* A ==> x : *s* {x. f x  \<in> A}"
+by (transfer, simp)
 
 text{*Alternative definition for hrabs with rabs function
    applied entrywise to equivalence class representative.
    This is easily proved using starfun and ns extension thm*}
 lemma hypreal_hrabs:
-     "abs (Abs_star (starrel``{X})) = Abs_star(starrel `` {%n. abs (X n)})"
-by (rule hypreal_hrabs)
+     "abs (star_n X) = star_n (%n. abs (X n))"
+by (simp only: starfun_rabs_hrabs [symmetric] starfun)
 
 text{*nonstandard extension of set through nonstandard extension
    of rabs function i.e hrabs. A more general result should be
@@ -409,26 +328,12 @@
 lemma STAR_rabs_add_minus:
    "*s* {x. abs (x + - y) < r} =
      {x. abs(x + -hypreal_of_real y) < hypreal_of_real r}"
-apply (simp add: starset_def, safe)
-apply (rule_tac [!] z = x in eq_Abs_star)
-apply (auto intro!: exI dest!: bspec
-             simp add: hypreal_minus hypreal_of_real_def hypreal_add
-                       star_of_def star_n_def 
-                       hypreal_hrabs hypreal_less)
-apply ultra
-done
+by (transfer, rule refl)
 
 lemma STAR_starfun_rabs_add_minus:
   "*s* {x. abs (f x + - y) < r} =
        {x. abs(( *f* f) x + -hypreal_of_real y) < hypreal_of_real r}"
-apply (simp add: starset_def, safe)
-apply (rule_tac [!] z = x in eq_Abs_star)
-apply (auto intro!: exI dest!: bspec
-             simp add: hypreal_minus hypreal_of_real_def hypreal_add
-                       star_of_def star_n_def
-                       hypreal_hrabs hypreal_less starfun)
-apply ultra
-done
+by (transfer, rule refl)
 
 text{*Another characterization of Infinitesimal and one of @= relation.
    In this theory since @{text hypreal_hrabs} proved here. Maybe
@@ -438,42 +343,39 @@
       (\<exists>X \<in> Rep_star(x).
         \<forall>m. {n. abs(X n) < inverse(real(Suc m))}
                 \<in>  FreeUltrafilterNat)"
-apply (rule_tac z = x in eq_Abs_star)
+apply (cases x)
 apply (auto intro!: bexI lemma_starrel_refl 
-            simp add: Infinitesimal_hypreal_of_nat_iff hypreal_of_real_def
-     star_of_def star_n_def
-     hypreal_inverse hypreal_hrabs hypreal_less hypreal_of_nat_eq)
+            simp add: Infinitesimal_hypreal_of_nat_iff star_of_def
+     star_n_inverse star_n_abs star_n_less hypreal_of_nat_eq)
 apply (drule_tac x = n in spec, ultra)
 done
 
-lemma approx_FreeUltrafilterNat_iff: "(Abs_star(starrel``{X}) @= Abs_star(starrel``{Y})) =
+lemma approx_FreeUltrafilterNat_iff: "star_n X @= star_n Y =
       (\<forall>m. {n. abs (X n + - Y n) <
                   inverse(real(Suc m))} : FreeUltrafilterNat)"
 apply (subst approx_minus_iff)
 apply (rule mem_infmal_iff [THEN subst])
-apply (auto simp add: hypreal_minus hypreal_add Infinitesimal_FreeUltrafilterNat_iff2)
+apply (auto simp add: star_n_minus star_n_add Infinitesimal_FreeUltrafilterNat_iff2)
 apply (drule_tac x = m in spec, ultra)
 done
 
 lemma inj_starfun: "inj starfun"
 apply (rule inj_onI)
 apply (rule ext, rule ccontr)
-apply (drule_tac x = "Abs_star (starrel ``{%n. xa}) " in fun_cong)
-apply (auto simp add: starfun)
+apply (drule_tac x = "star_n (%n. xa)" in fun_cong)
+apply (auto simp add: starfun star_n_eq_iff)
 done
 
 ML
 {*
-val starset_def = thm"starset_def";
 val starset_n_def = thm"starset_n_def";
 val InternalSets_def = thm"InternalSets_def";
 val is_starext_def = thm"is_starext_def";
-val starfun_def = thm"starfun_def";
 val starfun_n_def = thm"starfun_n_def";
 val InternalFuns_def = thm"InternalFuns_def";
 
 val no_choice = thm "no_choice";
-val STAR_real_set = thm "STAR_real_set";
+val STAR_UNIV_set = thm "STAR_UNIV_set";
 val STAR_empty_set = thm "STAR_empty_set";
 val STAR_Un = thm "STAR_Un";
 val STAR_Int = thm "STAR_Int";
@@ -482,7 +384,7 @@
 val STAR_diff = thm "STAR_diff";
 val STAR_subset = thm "STAR_subset";
 val STAR_mem = thm "STAR_mem";
-val STAR_hypreal_of_real_image_subset = thm "STAR_hypreal_of_real_image_subset";
+val STAR_star_of_image_subset = thm "STAR_star_of_image_subset";
 val STAR_hypreal_of_real_Int = thm "STAR_hypreal_of_real_Int";
 val STAR_real_seq_to_hypreal = thm "STAR_real_seq_to_hypreal";
 val STAR_singleton = thm "STAR_singleton";
@@ -492,7 +394,6 @@
 val starfun_n_starfun = thm "starfun_n_starfun";
 val hrabs_is_starext_rabs = thm "hrabs_is_starext_rabs";
 val Rep_star_FreeUltrafilterNat = thm "Rep_star_FreeUltrafilterNat";
-val starfun_congruent = thm "starfun_congruent";
 val starfun = thm "starfun";
 val starfun_mult = thm "starfun_mult";
 val starfun_add = thm "starfun_add";
--- a/src/HOL/Hyperreal/StarType.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/StarType.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -89,7 +89,7 @@
   star_of :: "'a \<Rightarrow> 'a star"
   "star_of x \<equiv> star_n (\<lambda>n. x)"
 
-theorem star_cases:
+theorem star_cases [case_names star_n, cases type: star]:
   "(\<And>X. x = star_n X \<Longrightarrow> P) \<Longrightarrow> P"
 by (unfold star_n_def, rule eq_Abs_star[of x], blast)
 
@@ -312,6 +312,22 @@
 lemmas star_of_le_1   = star_of_le   [of _ 1, simplified star_of_one]
 lemmas star_of_eq_1   = star_of_eq   [of _ 1, simplified star_of_one]
 
+text{*As above, for numerals*}
+
+lemmas star_of_number_less =
+  star_of_less [of "number_of b", simplified star_of_number_of, standard]
+lemmas star_of_number_le   =
+  star_of_le   [of "number_of b", simplified star_of_number_of, standard]
+lemmas star_of_number_eq   =
+  star_of_eq   [of "number_of b", simplified star_of_number_of, standard]
+
+lemmas star_of_less_number =
+  star_of_less [of _ "number_of b", simplified star_of_number_of, standard]
+lemmas star_of_le_number   =
+  star_of_le   [of _ "number_of b", simplified star_of_number_of, standard]
+lemmas star_of_eq_number   =
+  star_of_eq   [of _ "number_of b", simplified star_of_number_of, standard]
+
 lemmas star_of_simps =
   star_of_add     star_of_diff    star_of_minus
   star_of_mult    star_of_divide  star_of_inverse
@@ -323,6 +339,8 @@
   star_of_less_0  star_of_le_0    star_of_eq_0
   star_of_1_less  star_of_1_le    star_of_1_eq
   star_of_less_1  star_of_le_1    star_of_eq_1
+  star_of_number_less star_of_number_le star_of_number_eq
+  star_of_less_number star_of_le_number star_of_eq_number
 
 declare star_of_simps [simp]
 
--- a/src/HOL/Hyperreal/Transcendental.thy	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/Transcendental.thy	Fri Sep 09 19:34:22 2005 +0200
@@ -2473,8 +2473,8 @@
 by (erule DERIV_fun_exp)
 
 lemma STAR_exp_ln: "0 < z ==> ( *f* (%x. exp (ln x))) z = z"
-apply (rule_tac z = z in eq_Abs_star)
-apply (auto simp add: starfun hypreal_zero_def hypreal_less)
+apply (cases z)
+apply (auto simp add: starfun star_n_zero_num star_n_less star_n_eq_iff)
 done
 
 lemma hypreal_add_Infinitesimal_gt_zero:
--- a/src/HOL/Hyperreal/hypreal_arith.ML	Fri Sep 09 17:47:37 2005 +0200
+++ b/src/HOL/Hyperreal/hypreal_arith.ML	Fri Sep 09 19:34:22 2005 +0200
@@ -13,9 +13,9 @@
 
 local
 
-val real_inj_thms = [hypreal_of_real_le_iff RS iffD2, 
-                     hypreal_of_real_less_iff RS iffD2,
-                     hypreal_of_real_eq_iff RS iffD2];
+val real_inj_thms = [thm "star_of_le" RS iffD2, 
+                     thm "star_of_less" RS iffD2,
+                     thm "star_of_eq" RS iffD2];
 
 in