first stage in tidying up Real and Hyperreal.
authorpaulson
Tue, 12 Dec 2000 12:01:19 +0100
changeset 10648 a8c647cfa31f
parent 10647 a4529a251b6f
child 10649 fb27b5547765
first stage in tidying up Real and Hyperreal. Factor cancellation simprocs inverse #0 = #0 simprules for division corrected ambigous syntax definitions in Hyperreal
src/HOL/Real/Hyperreal/HyperDef.ML
src/HOL/Real/Hyperreal/Lim.ML
src/HOL/Real/Hyperreal/Lim.thy
src/HOL/Real/Hyperreal/NSA.ML
src/HOL/Real/Hyperreal/SEQ.ML
src/HOL/Real/Hyperreal/SEQ.thy
src/HOL/Real/RealAbs.ML
src/HOL/Real/RealArith.ML
src/HOL/Real/RealBin.ML
src/HOL/Real/RealDef.ML
src/HOL/Real/RealDef.thy
src/HOL/Real/RealOrd.ML
src/HOL/Real/RealPow.ML
--- a/src/HOL/Real/Hyperreal/HyperDef.ML	Tue Dec 12 11:59:25 2000 +0100
+++ b/src/HOL/Real/Hyperreal/HyperDef.ML	Tue Dec 12 12:01:19 2000 +0100
@@ -338,16 +338,15 @@
 by (rotate_tac 1 1);
 by (asm_full_simp_tac (simpset() addsimps 
     [hypreal_inverse,hypreal_zero_def] addsplits [split_if]) 1);
-by (ultra_tac (claset() addDs (map rename_numerals [real_inverse_not_zero,real_inverse_inverse]),
-	       simpset()) 1);
+by (ultra_tac (claset() addDs (map rename_numerals [real_inverse_not_zero]),
+	       simpset() addsimps [real_inverse_inverse]) 1);
 qed "hypreal_inverse_inverse";
 
 Addsimps [hypreal_inverse_inverse];
 
 Goalw [hypreal_one_def] "inverse(1hr) = 1hr";
 by (full_simp_tac (simpset() addsimps [hypreal_inverse,
-       real_zero_not_eq_one RS not_sym] 
-                   addsplits [split_if]) 1);
+                                       real_zero_not_eq_one RS not_sym]) 1);
 qed "hypreal_inverse_1";
 Addsimps [hypreal_inverse_1];
 
@@ -403,13 +402,11 @@
 
 Goalw [hypreal_zero_def] "z + -z = (0::hypreal)";
 by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_minus,
-        hypreal_add]) 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_minus, hypreal_add]) 1);
 qed "hypreal_add_minus";
 
 Goal "-z + z = (0::hypreal)";
-by (simp_tac (simpset() addsimps 
-    [hypreal_add_commute,hypreal_add_minus]) 1);
+by (simp_tac (simpset() addsimps [hypreal_add_commute, hypreal_add_minus]) 1);
 qed "hypreal_add_minus_left";
 
 Addsimps [hypreal_add_minus,hypreal_add_minus_left,
@@ -448,8 +445,9 @@
 Goal "-(x + (y::hypreal)) = -x + -y";
 by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
 by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_minus,
-    hypreal_add,real_minus_add_distrib]));
+by (auto_tac (claset(),
+              simpset() addsimps [hypreal_minus, hypreal_add,
+                                  real_minus_add_distrib]));
 qed "hypreal_minus_add_distrib";
 Addsimps [hypreal_minus_add_distrib];
 
@@ -460,7 +458,7 @@
 Goal "(x + - (y::hypreal)) + (y + - z) = x + -z";
 by (res_inst_tac [("w1","y")] (hypreal_add_commute RS subst) 1);
 by (simp_tac (simpset() addsimps [hypreal_add_left_commute,
-    hypreal_add_assoc]) 1);
+                                  hypreal_add_assoc]) 1);
 by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
 qed "hypreal_add_minus_cancel1";
 
@@ -488,15 +486,16 @@
 Addsimps [hypreal_add_minus_cancelc];
 
 Goal "(z + -x) + (y + -z) = (y + -(x::hypreal))";
-by (full_simp_tac (simpset() addsimps [hypreal_minus_add_distrib
-    RS sym, hypreal_add_left_cancel] @ hypreal_add_ac 
-    delsimps [hypreal_minus_add_distrib]) 1); 
+by (full_simp_tac
+    (simpset() addsimps [hypreal_minus_add_distrib RS sym, 
+                         hypreal_add_left_cancel] @ hypreal_add_ac 
+               delsimps [hypreal_minus_add_distrib]) 1); 
 qed "hypreal_add_minus_cancel3";
 Addsimps [hypreal_add_minus_cancel3];
 
 Goal "(y + (x::hypreal)= z + x) = (y = z)";
 by (simp_tac (simpset() addsimps [hypreal_add_commute,
-    hypreal_add_left_cancel]) 1);
+                                  hypreal_add_left_cancel]) 1);
 qed "hypreal_add_right_cancel";
 
 Goal "z + (y + -z) = (y::hypreal)";
@@ -549,7 +548,8 @@
 
 qed_goal "hypreal_mult_left_commute" (the_context ())
     "(z1::hypreal) * (z2 * z3) = z2 * (z1 * z3)"
- (fn _ => [rtac (hypreal_mult_commute RS trans) 1, rtac (hypreal_mult_assoc RS trans) 1,
+ (fn _ => [rtac (hypreal_mult_commute RS trans) 1, 
+           rtac (hypreal_mult_assoc RS trans) 1,
            rtac (hypreal_mult_commute RS arg_cong) 1]);
 
 (* hypreal multiplication is an AC operator *)
--- a/src/HOL/Real/Hyperreal/Lim.ML	Tue Dec 12 11:59:25 2000 +0100
+++ b/src/HOL/Real/Hyperreal/Lim.ML	Tue Dec 12 12:01:19 2000 +0100
@@ -9,25 +9,30 @@
 fun ARITH_PROVE str = prove_goal thy str 
                       (fn prems => [cut_facts_tac prems 1,arith_tac 1]);
 
+Goal "(x + - a = (#0::real)) = (x=a)";
+by (arith_tac 1);
+qed "real_add_minus_iff";
+Addsimps [real_add_minus_iff];
+
 (*---------------------------------------------------------------
    Theory of limits, continuity and differentiation of 
    real=>real functions 
  ----------------------------------------------------------------*)
 Goalw [LIM_def] "(f -- a --> L) = (ALL r. #0 < r --> \
-\    (EX s. #0 < s & (ALL x. (#0 < abs(x + -a) & (abs(x + -a) < s) \
+\    (EX s. #0 < s & (ALL x. (x ~= a & (abs(x + -a) < s) \
 \          --> abs(f x + -L) < r))))";
-by (Blast_tac 1);
+by Auto_tac;
 qed "LIM_iff";
 
 Goalw [LIM_def] 
       "!!a. [| f -- a --> L; #0 < r |] \
-\           ==> (EX s. #0 < s & (ALL x. (#0 < abs(x + -a) \
+\           ==> (EX s. #0 < s & (ALL x. (x ~= a \
 \           & (abs(x + -a) < s) --> abs(f x + -L) < r)))";
-by (Blast_tac 1);
+by Auto_tac;
 qed "LIMD";
 
 Goalw [LIM_def] "(%x. k) -- x --> k";
-by (Auto_tac);
+by Auto_tac;
 qed "LIM_const";
 Addsimps [LIM_const];
 
@@ -44,7 +49,7 @@
 by (Step_tac 1);
 by (REPEAT(dres_inst_tac [("x","r*inverse(#1 + #1)")] spec 1));
 by (dtac (rename_numerals (real_zero_less_two RS real_inverse_gt_zero 
-    RSN (2,real_mult_less_mono2))) 1);
+          RSN (2,real_mult_less_mono2))) 1);
 by (Asm_full_simp_tac 1);
 by (Clarify_tac 1);
 by (res_inst_tac [("R1.0","s"),("R2.0","sa")] 
@@ -66,7 +71,8 @@
 
 Goalw [LIM_def] "f -- a --> L ==> (%x. -f(x)) -- a --> -L";
 by (full_simp_tac (simpset() addsimps [real_minus_add_distrib RS sym,
-    abs_minus_cancel] delsimps [real_minus_add_distrib,real_minus_minus]) 1);
+                                       abs_minus_cancel] 
+                    delsimps [real_minus_add_distrib,real_minus_minus]) 1);
 qed "LIM_minus";
 
 (*----------------------------------------------
@@ -81,8 +87,7 @@
      LIM_zero
  ----------------------------------------------*)
 Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> #0";
-by (res_inst_tac [("z1","l")] (rename_numerals 
-    (real_add_minus RS subst)) 1);
+by (res_inst_tac [("z1","l")] (rename_numerals (real_add_minus RS subst)) 1);
 by (rtac LIM_add_minus 1 THEN Auto_tac);
 qed "LIM_zero";
 
@@ -155,25 +160,20 @@
 by (Auto_tac);
 qed "LIM_self";
 
-Goal "(#0::real) < abs (x + -a) ==> x ~= a";
-by (arith_tac 1);
-qed "lemma_rabs_not_eq";
-
 (*--------------------------------------------------------------
    Limits are equal for functions equal except at limit point
  --------------------------------------------------------------*)
 Goalw [LIM_def] 
       "[| ALL x. x ~= a --> (f x = g x) |] \
 \           ==> (f -- a --> l) = (g -- a --> l)";
-by (auto_tac (claset(),simpset() addsimps [lemma_rabs_not_eq]));
+by (auto_tac (claset(), simpset() addsimps [real_add_minus_iff]));
 qed "LIM_equal";
 
 Goal "[| (%x. f(x) + -g(x)) -- a --> #0; \
 \        g -- a --> l |] \
 \      ==> f -- a --> l";
 by (dtac LIM_add 1 THEN assume_tac 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [real_add_assoc]));
+by (auto_tac (claset(), simpset() addsimps [real_add_assoc]));
 qed "LIM_trans";
 
 (***-------------------------------------------------------------***)
@@ -184,54 +184,47 @@
  --------------------------------------------------------------*)
 Goalw [LIM_def,NSLIM_def,inf_close_def] 
       "f -- x --> L ==> f -- x --NS> L";
-by (asm_full_simp_tac (simpset() addsimps 
-    [Infinitesimal_FreeUltrafilterNat_iff]) 1);
+by (asm_full_simp_tac
+    (simpset() addsimps [Infinitesimal_FreeUltrafilterNat_iff]) 1);
 by (Step_tac 1);
 by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [starfun,
-    hypreal_minus,hypreal_of_real_minus RS sym,
-    hypreal_of_real_def,hypreal_add]));
+by (auto_tac (claset(),
+      simpset() addsimps [real_add_minus_iff,
+                          starfun, hypreal_minus,hypreal_of_real_minus RS sym,
+                          hypreal_of_real_def, hypreal_add]));
 by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2 THEN Step_tac 1);
 by (dres_inst_tac [("x","u")] spec 1 THEN Clarify_tac 1);
 by (dres_inst_tac [("x","s")] spec 1 THEN Clarify_tac 1);
-by (subgoal_tac "ALL n::nat. (#0::real) < abs ((xa n) + - x) & \
-\  abs ((xa n) + - x) < s --> abs (f (xa n) + - L) < u" 1);
+by (subgoal_tac "ALL n::nat. (xa n) ~= x & \
+\                    abs ((xa n) + - x) < s --> abs (f (xa n) + - L) < u" 1);
 by (Blast_tac 2);
 by (dtac FreeUltrafilterNat_all 1);
-by (Ultra_tac 1 THEN arith_tac 1);
+by (Ultra_tac 1);
 qed "LIM_NSLIM";
  
 (*---------------------------------------------------------------------
     Limit: NS definition ==> standard definition
  ---------------------------------------------------------------------*)
 
-Goal "ALL s. #0 < s --> (EX xa.  #0 < abs (xa + - x) & \
+Goal "ALL s. #0 < s --> (EX xa.  xa ~= x & \
 \        abs (xa + - x) < s  & r <= abs (f xa + -L)) \
-\     ==> ALL n::nat. EX xa.  #0 < abs (xa + - x) & \
+\     ==> ALL n::nat. EX xa.  xa ~= x & \
 \             abs(xa + -x) < inverse(real_of_posnat n) & r <= abs(f xa + -L)";
 by (Step_tac 1);
 by (cut_inst_tac [("n1","n")] (real_of_posnat_gt_zero RS real_inverse_gt_zero) 1);
 by Auto_tac;
 val lemma_LIM = result();
 
-Goal "ALL s. #0 < s --> (EX xa.  #0 < abs (xa + - x) & \
+Goal "ALL s. #0 < s --> (EX xa.  xa ~= x & \
 \        abs (xa + - x) < s  & r <= abs (f xa + -L)) \
-\     ==> EX X. ALL n::nat. #0 < abs (X n + - x) & \
+\     ==> EX X. ALL n::nat. X n ~= x & \
 \               abs(X n + -x) < inverse(real_of_posnat n) & r <= abs(f (X n) + -L)";
 by (dtac lemma_LIM 1);
 by (dtac choice 1);
 by (Blast_tac 1);
 val lemma_skolemize_LIM2 = result();
 
-Goal "{n. f (X n) + - L = Y n} Int \
-\         {n. #0 < abs (X n + - x) & \
-\             abs (X n + - x) < inverse (real_of_posnat  n) & \
-\             r <= abs (f (X n) + - L)} <= \
-\         {n. r <= abs (Y n)}";
-by (Auto_tac);
-val lemma_Int = result ();
-
-Goal "!!X. ALL n. #0 < abs (X n + - x) & \
+Goal "!!X. ALL n. X n ~= x & \
 \         abs (X n + - x) < inverse (real_of_posnat  n) & \
 \         r <= abs (f (X n) + - L) ==> \
 \         ALL n. abs (X n + - x) < inverse (real_of_posnat  n)";
@@ -244,25 +237,24 @@
 
 Goalw [LIM_def,NSLIM_def,inf_close_def] 
       "!!f. f -- x --NS> L ==> f -- x --> L";
-by (asm_full_simp_tac (simpset() addsimps 
-    [Infinitesimal_FreeUltrafilterNat_iff]) 1);
+by (asm_full_simp_tac
+    (simpset() addsimps [Infinitesimal_FreeUltrafilterNat_iff]) 1);
 by (EVERY1[Step_tac, rtac ccontr, Asm_full_simp_tac]);
 by (fold_tac [real_le_def]);
 by (dtac lemma_skolemize_LIM2 1);
 by (Step_tac 1);
 by (dres_inst_tac [("x","Abs_hypreal(hyprel^^{X})")] spec 1);
-by (asm_full_simp_tac (simpset() addsimps [starfun,
-    hypreal_minus,hypreal_of_real_minus RS sym,
-    hypreal_of_real_def,hypreal_add]) 1);
+by (asm_full_simp_tac
+    (simpset() addsimps [starfun,
+                         hypreal_minus,hypreal_of_real_minus RS sym,
+                         hypreal_of_real_def,hypreal_add]) 1);
 by (Step_tac 1);
-by (dtac FreeUltrafilterNat_Ex_P 1 THEN etac exE 1);
-by (dres_inst_tac [("x","n")] spec 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [real_add_minus,real_less_not_refl]) 1);
 by (dtac (lemma_simp RS real_seq_to_hypreal_Infinitesimal) 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [Infinitesimal_FreeUltrafilterNat_iff,hypreal_of_real_def,
-     hypreal_minus,hypreal_of_real_minus,hypreal_add]) 1);
+by (asm_full_simp_tac
+    (simpset() addsimps 
+       [Infinitesimal_FreeUltrafilterNat_iff,hypreal_of_real_def,
+        hypreal_minus,hypreal_of_real_minus,hypreal_add]) 1);
+by (Blast_tac 1); 
 by (rotate_tac 2 1);
 by (dres_inst_tac [("x","r")] spec 1);
 by (Clarify_tac 1);
@@ -270,6 +262,7 @@
 by (Ultra_tac 1);
 qed "NSLIM_LIM";
 
+
 (**** Key result ****)
 Goal "(f -- x --> L) = (f -- x --NS> L)";
 by (blast_tac (claset() addIs [LIM_NSLIM,NSLIM_LIM]) 1);
@@ -726,7 +719,7 @@
 
 Goalw [isUCont_def,isCont_def,LIM_def]
      "isUCont f ==> EX x. isCont f x";
-by (Fast_tac 1);
+by (Force_tac 1);
 qed "isUCont_isCont";
 
 Goalw [isNSUCont_def,isUCont_def,inf_close_def] 
@@ -1330,9 +1323,9 @@
     NSDERIV_chain]) 1);
 qed "DERIV_chain";
 
-Goal "[| DERIV f g x :> Da; DERIV g x :> Db |] \
+Goal "[| DERIV f (g x) :> Da; DERIV g x :> Db |] \
 \     ==> DERIV (%x. f (g x)) x :> Da * Db";
-by (auto_tac (claset() addDs [DERIV_chain],simpset() addsimps [o_def]));
+by (auto_tac (claset() addDs [DERIV_chain], simpset() addsimps [o_def]));
 qed "DERIV_chain2";
 
 Goal "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E";
@@ -1355,13 +1348,13 @@
 qed "DERIV_Id";
 Addsimps [DERIV_Id];
 
-Goal "DERIV op * c x :> c";
+Goal "DERIV (op * c) x :> c";
 by (cut_inst_tac [("c","c"),("x","x")] (DERIV_Id RS DERIV_cmult) 1);
 by (Asm_full_simp_tac 1);
 qed "DERIV_cmult_Id";
 Addsimps [DERIV_cmult_Id];
 
-Goal "NSDERIV op * c x :> c";
+Goal "NSDERIV (op * c) x :> c";
 by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff]) 1);
 qed "NSDERIV_cmult_Id";
 Addsimps [NSDERIV_cmult_Id];
@@ -1496,13 +1489,426 @@
 by (asm_full_simp_tac (simpset() addsimps [isNSCont_def]) 1);
 qed "CARAT_DERIVD";
  
-(*--------------------------------------------------------------------*)
-(* In this theory, still have to include Bisection theorem, IVT, MVT, *)
-(* Rolle etc.                                                         *)
-(*--------------------------------------------------------------------*)
+
+
+(*----------------------------------------------------------------------------*)
+(* Useful lemmas about nested intervals and proof by bisection (cf.Harrison)  *)
+(*----------------------------------------------------------------------------*)
+
+Goal "(ALL n. (f::nat=>real) n <= f (Suc n)) --> f m <= f(m + no)";
+by (induct_tac "no" 1);
+by (auto_tac (claset() addIs [real_le_trans],simpset()));
+qed_spec_mp "lemma_f_mono_add";
+
+
+(* IMPROVE?: long proof! *)
+Goal "[| ALL n. f(n) <= f(Suc n); \
+\        ALL n. g(Suc n) <= g(n); \
+\        ALL n. f(n) <= g(n) |] \
+\     ==> EX l m. l <= m &  ((ALL n. f(n) <= l) & f ----> l) & \
+\                           ((ALL n. m <= g(n)) & g ----> m)";
+by (subgoal_tac "Bseq f" 1);
+by (res_inst_tac [("k","f 0"),("K","g 0")] BseqI2 2 THEN rtac allI 2);
+by (induct_tac "n" 2);
+by (auto_tac (claset() addIs [real_le_trans],simpset()));
+by (res_inst_tac [("j","g(Suc na)")] real_le_trans 2);
+by (induct_tac "na" 3);
+by (auto_tac (claset() addIs [real_le_trans],simpset()));
+by (subgoal_tac "monoseq f" 1);
+by (subgoal_tac "monoseq g" 1);
+by (subgoal_tac "Bseq g" 1);
+by (res_inst_tac [("k","f 0"),("K","g 0")] BseqI2 2 THEN rtac allI 2);
+by (induct_tac "n" 2);
+by (auto_tac (claset() addIs [real_le_trans],simpset()));
+by (res_inst_tac [("j","f(Suc na)")] real_le_trans 2);
+by (induct_tac "na" 2);
+by (auto_tac (claset() addIs [real_le_trans],simpset()));
+by (auto_tac (claset() addSDs [Bseq_monoseq_convergent],
+    simpset() addsimps [convergent_LIMSEQ_iff]));
+by (res_inst_tac [("x","lim f")] exI 1);
+by (res_inst_tac [("x","lim g")] exI 1);
+by (auto_tac (claset() addIs [LIMSEQ_le],simpset()));
+by (rtac real_leI 1 THEN rtac real_leI 2);
+by (auto_tac (claset(),simpset() addsimps [LIMSEQ_iff,monoseq_Suc]));
+by (ALLGOALS (dtac real_less_sum_gt_zero));
+by (dres_inst_tac [("x","f n + - lim f")] spec 1);
+by (rotate_tac 4 2);
+by (dres_inst_tac [("x","lim g + - g n")] spec 2);
+by (Step_tac 1);
+by (ALLGOALS(rotate_tac 5));
+by (ALLGOALS(dres_inst_tac [("x","no + n")] spec));
+by (Auto_tac);
+by (subgoal_tac "0 <= f(no + n) - lim f" 1);
+by (induct_tac "no" 2);
+by (auto_tac (claset() addIs [real_le_trans],
+    simpset() addsimps [real_diff_def]));
+by (dtac abs_eqI1 1 THEN Asm_full_simp_tac 1);
+by (dres_inst_tac [("i"," f (no + n)"),("no1","no")] (lemma_f_mono_add 
+    RSN (2,real_less_le_trans)) 1);
+by (subgoal_tac "0 <= lim g + - g(no + n)" 3);
+by (induct_tac "no" 4);
+by (auto_tac (claset() addIs [real_le_trans],
+    simpset() addsimps [real_diff_def,abs_minus_add_cancel]));
+by (dtac abs_eqI1 2 THEN Asm_full_simp_tac 2);
+by (dres_inst_tac [("i","- g (no + n)"),("no1","no")] (lemma_f_mono_add 
+    RSN (2,real_less_le_trans)) 2);
+by (auto_tac (claset(),simpset() addsimps [add_commute]));
+qed "lemma_nest";
+
+Goal "[| ALL n. f(n) <= f(Suc n); \
+\        ALL n. g(Suc n) <= g(n); \
+\        ALL n. f(n) <= g(n); \
+\        (%n. f(n) - g(n)) ----> 0 |] \
+\     ==> EX l. ((ALL n. f(n) <= l) & f ----> l) & \
+\               ((ALL n. l <= g(n)) & g ----> l)";
+by (dtac lemma_nest 1 THEN Auto_tac);
+by (subgoal_tac "l = m" 1);
+by (dres_inst_tac [("X","f")] LIMSEQ_diff 2);
+by (auto_tac (claset() addIs [LIMSEQ_unique],simpset()));
+qed "lemma_nest_unique";
+
+Goal "EX! fn. (fn 0 = e) & (ALL n. fn (Suc n) = f n (fn n))";
+by (rtac ex1I 1);
+by (rtac conjI 1 THEN rtac allI 2);
+by (rtac def_nat_rec_0 1 THEN rtac def_nat_rec_Suc 2);
+by (Auto_tac);
+by (rtac ext 1 THEN induct_tac "n" 1);
+by (Auto_tac);
+qed "nat_Axiom";
+
+
+(****************************************************************
+REPLACING THE VERSION IN REALORD.ML
+****************************************************************)
+
+(*NEW VERSION with #2*)
+Goal "x < y ==> x < (x + y) / (#2::real)";
+by Auto_tac;
+(*proof was
+by (dres_inst_tac [("C","x")] real_add_less_mono2 1);
+by (dtac (real_add_self RS subst) 1);
+by (dtac (real_zero_less_two RS real_inverse_gt_zero RS 
+          real_mult_less_mono1) 1);
+by (asm_full_simp_tac
+    (simpset() addsimps [real_mult_assoc, real_inverse_eq_divide]) 1);
+*)
+qed "real_less_half_sum";
+
+
+(*Replaces "inverse #nn" by #1/#nn *)
+Addsimps [inst "x" "number_of ?w" real_inverse_eq_divide];
+
+Goal "((x::real) = y / (#2 * z)) = (#2 * x = y/z)";
+by Auto_tac;  
+by (dres_inst_tac [("f","%u. (#1/#2)*u")] arg_cong 1); 
+by Auto_tac;  
+qed "eq_divide_2_times_iff";
 
 
- 
+val [prem1,prem2] =
+Goal "[| !!a b c. [| a <= b; b <= c; P(a,b); P(b,c)|] ==> P(a,c); \
+\        ALL x. EX d::real. 0 < d & \
+\               (ALL a b. a <= x & x <= b & (b - a) < d --> P(a,b)) \
+\     |] ==> ALL a b. a <= b --> P(a,b)";
+by (Step_tac 1);
+by (cut_inst_tac [("e","(a,b)"),
+    ("f","%n fn. if P(fst fn, (fst fn + snd fn)/#2) \
+\              then ((fst fn + snd fn) /#2,snd fn) \
+\              else (fst fn,(fst fn + snd fn)/#2)")] 
+    nat_Axiom 1);
+by (etac ex1E 1 THEN etac conjE 1 THEN rtac ccontr 1);
+(* set up 1st premise of lemma_nest_unique *)
+by (subgoal_tac "ALL n. fst((fn::nat =>(real*real)) n) <= snd (fn n)" 1);
+by (rtac allI 2);
+by (induct_tac "n" 2);
+by (Asm_simp_tac 2);
+by (dres_inst_tac [("x","na")] spec 2);
+by (case_tac "P (fst (fn na), (fst (fn na) + snd (fn na)) / #2)" 2);
+by (Asm_simp_tac 3);
+by (Asm_simp_tac 2);
+(* 2nd premise *) 
+by (subgoal_tac "ALL n. fst(fn n) <= fst (fn (Suc n))" 1);
+by (rtac allI 2);
+by (induct_tac "n" 2);
+by (dres_inst_tac [("x","0")] spec 2);
+by (Asm_simp_tac 3);  (*super slow, but proved!*)
+by (Asm_simp_tac 2);
+(* 3rd premise? [lcp] *)
+by (subgoal_tac "ALL n. ~P(fst(fn n),snd(fn n))" 1);
+by (rtac allI 2);
+by (induct_tac "n" 2);
+by (Asm_simp_tac 2);
+by (Asm_simp_tac 2 THEN rtac impI 2);
+by (rtac ccontr 2);
+by (res_inst_tac [("Pa","P (fst (fn na), snd (fn na))")] swap 2);
+by (assume_tac 2);
+by (res_inst_tac [("b","(fst(fn na) + snd(fn na))/#2")] prem1 2);
+by (Asm_full_simp_tac 4); 
+by (Asm_full_simp_tac 4); 
+by (Asm_full_simp_tac 2); 
+by (Asm_simp_tac 2); 
+(* 3rd premise [looks like the 4th to lcp!] *)
+by (subgoal_tac "ALL n. snd(fn (Suc n)) <= snd (fn n)" 1);
+by (rtac allI 2);
+by (induct_tac "n" 2);
+by (Asm_simp_tac 2);
+by (Asm_simp_tac 2);
+(* FIXME: simplification takes a very long time! *)
+by (ALLGOALS(thin_tac "ALL y. \
+\            y 0 = (a, b) & \
+\            (ALL n. \
+\                y (Suc n) = \
+\                (if P (fst (y n), (fst (y n) + snd (y n)) /#2) \
+\                 then ((fst (y n) + snd (y n)) /#2, snd (y n)) \
+\                 else (fst (y n),\
+\                       (fst (y n) + snd (y n)) /#2))) --> \
+\            y = fn"));
+(*another premise? the 5th? lcp*)
+by (subgoal_tac "ALL n. snd(fn n) - fst(fn n) = \
+\                       (b - a) * inverse(#2 ^ n)" 1);
+by (rtac allI 2);
+by (induct_tac "n" 2);
+by (Asm_simp_tac 2);
+by (asm_full_simp_tac
+    (simpset() addsimps [eq_divide_2_times_iff, real_inverse_eq_divide,
+                         real_diff_mult_distrib2]) 2);
+(*end of the premises [lcp]*)
+by (dtac lemma_nest_unique 1);
+by (REPEAT(assume_tac 1));
+by (Step_tac 2);
+by (rtac LIMSEQ_minus_cancel 1);
+by (asm_simp_tac (simpset() addsimps [CLAIM "-((a::real) - b) = (b - a)",
+    realpow_inverse]) 1);
+by (subgoal_tac "#0 = (b-a) * #0" 1);
+by (eres_inst_tac [("t","#0")] ssubst 1); 
+by (rtac (LIMSEQ_const RS LIMSEQ_mult) 1);
+by (rtac LIMSEQ_realpow_zero 1);
+by (Asm_full_simp_tac 3); 
+by (EVERY1[Simp_tac, Simp_tac]);
+by (cut_facts_tac [prem2] 1);
+by (dres_inst_tac [("x","l")] spec 1 THEN Step_tac 1);
+by (rewtac LIMSEQ_def);
+by (dres_inst_tac [("x","d/#2")] spec 1);
+by (dres_inst_tac [("x","d/#2")] spec 1);
+by (dtac real_less_half_sum 1);
+by (thin_tac "ALL n. \
+\             fn (Suc n) = \
+\             (if P (fst (fn n), (fst (fn n) + snd (fn n))/#2) \
+\              then ((fst (fn n) + snd (fn n))/#2, snd (fn n)) \
+\              else (fst (fn n), \
+\                    (fst (fn n) + snd (fn n))/#2))" 1);
+by (Step_tac 1);
+by (ALLGOALS(Asm_full_simp_tac));
+by (dres_inst_tac [("x","fst(fn (no + noa))")] spec 1);
+by (dres_inst_tac [("x","snd(fn (no + noa))")] spec 1);
+by (Step_tac 1);
+by (ALLGOALS(Asm_full_simp_tac));
+by (res_inst_tac [("j","abs(fst(fn(no + noa)) - l) + \
+\                       abs(snd(fn(no + noa)) - l)")] real_le_less_trans 1);
+by (rtac (real_sum_of_halves RS subst) 2);
+by (rewtac real_diff_def);
+by (rtac real_add_less_mono 2);
+
+by (Asm_full_simp_tac 2); 
+by (Asm_full_simp_tac 2); 
+by (res_inst_tac [("y1","fst(fn (no + noa)) ")] 
+    (abs_minus_add_cancel RS subst) 1);
+by (subgoal_tac "#0 <= (l + - fst (fn (no + noa)))" 1);
+by (subgoal_tac "#0 <= (snd (fn (no + noa)) + - l)" 1);
+by (asm_simp_tac (simpset() addsimps [abs_eqI1] @ real_add_ac) 1);
+by (asm_simp_tac (simpset() addsimps real_add_ac) 1);
+by (res_inst_tac [("C","l")] real_le_add_right_cancel 1);
+by (res_inst_tac [("C"," fst (fn (no + noa))")] real_le_add_right_cancel 2);
+by (ALLGOALS(asm_simp_tac (simpset() addsimps real_add_ac)));
+qed "lemma_BOLZANO";
 
 
- 
\ No newline at end of file
+Goal "((ALL a b c. (a <= b & b <= c & P(a,b) & P(b,c)) --> P(a,c)) & \
+\         (ALL x. EX d::real. 0 < d & \
+\               (ALL a b. a <= x & x <= b & (b - a) < d --> P(a,b)))) \
+\         --> (ALL a b. a <= b --> P(a,b))";
+by (Step_tac 1);
+by (rtac (lemma_BOLZANO RS allE) 1);
+by (assume_tac 2);
+by (ALLGOALS(Blast_tac));
+qed "lemma_BOLZANO2";
+
+(*----------------------------------------------------------------------------*)
+(* Intermediate Value Theorem (prove contrapositive by bisection)             *)
+(*----------------------------------------------------------------------------*)
+(* proved elsewhere surely? *)
+Goal "!!P. (~P ==> ~Q) ==> Q ==> P";
+by (auto_tac (claset() addIs [ccontr],simpset()));
+qed "contrapos2";
+
+(*
+see junk.ML
+Goal "[| f(a) <= y & y <= f(b); \
+\        a <= b; \
+\        (ALL x. a <= x & x <= b --> isCont f x) \
+\     |] ==> EX x. a <= x & x <= b & f(x) = y";
+by (rtac contrapos2 1);
+by (assume_tac 2);
+by (cut_inst_tac [("P","%(u,v). a <= u & u <= v & v <= b --> \
+\   ~(f(u) <= y & y <= f(v))")] lemma_BOLZANO2 1);
+by (Step_tac 1);
+by (ALLGOALS(Asm_full_simp_tac));
+by (Blast_tac 2);
+by (asm_full_simp_tac (simpset() addsimps [isCont_iff,LIM_def]) 1);
+by (rtac ccontr 1);
+by (subgoal_tac "a <= x & x <= b" 1);
+by (Asm_full_simp_tac 2);
+by (rotate_tac 3 2);
+by (dres_inst_tac [("x","1r")] spec 2);
+by (Step_tac 2);
+by (asm_full_simp_tac (simpset() addsimps []) 2);
+by (asm_full_simp_tac (simpset() addsimps []) 2);
+by (REPEAT(blast_tac (claset() addIs [real_le_trans]) 2));
+by (REPEAT(dres_inst_tac [("x","x")] spec 1));
+by (Asm_full_simp_tac 1);
+(**)
+by (rotate_tac 4 1);
+by (dres_inst_tac [("x","abs(y - f x)")] spec 1);
+by (Step_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [real_less_le,
+                                           abs_ge_zero,real_diff_def]) 1);
+by (dtac (sym RS (abs_zero_iff RS iffD2)) 1);
+by (arith_tac 1);
+by (dres_inst_tac [("x","s")] spec 1);
+by (Step_tac 1);
+by (cut_inst_tac [("R1.0","f x"),("R2.0","y")] real_linear 1);
+by (Auto_tac);
+by (dres_inst_tac [("x","ba - x")] spec 1);
+by (auto_tac (claset(),simpset() addsimps [abs_iff,real_diff_le_eq,
+       (real_diff_def RS meta_eq_to_obj_eq) RS sym,real_less_le_iff,
+       real_le_diff_eq,CLAIM "(~(x::real) <= y) = (y < x)",
+       CLAIM "(-x < -y) = ((y::real) < x)",
+       CLAIM "(-(y - x)) = (x - (y::real))"]));
+by (dres_inst_tac [("z","x"),("w","ba")] real_le_anti_sym 1);
+by (assume_tac 1 THEN Asm_full_simp_tac 1);
+by (dres_inst_tac [("x","aa - x")] spec 1);
+by (case_tac "x <= aa" 1);
+by (auto_tac (claset(),simpset() addsimps [real_minus_zero_le_iff2,
+    real_le_diff_eq,real_diff_le_eq,(real_le_def RS meta_eq_to_obj_eq) 
+    RS sym]));
+by (dres_inst_tac [("z","x"),("w","aa")] real_le_anti_sym 1);
+by (assume_tac 1 THEN Asm_full_simp_tac 1);
+qed "IVT";
+
+Goal "[| f(b) <= y & y <= f(a); \
+\        a <= b; \
+\        (ALL x. a <= x & x <= b --> isCont f x) \
+\     |] ==> EX x. a <= x & x <= b & f(x) = y";
+by (subgoal_tac "- f a <= -y & -y <= - f b" 1);
+by (thin_tac "f b <= y & y <= f a" 1);
+by (dres_inst_tac [("f","%x. - f x")] IVT 1);
+by (auto_tac (claset() addIs [isCont_minus],simpset()));
+qed "IVT2";
+
+Goal "(f(a) <= y & y <= f(b) & a <= b & \
+\     (ALL x. a <= x & x <= b --> isCont f x)) \
+\     --> (EX x. a <= x & x <= b & f(x) = y)";
+by (blast_tac (claset() addIs [IVT]) 1);
+qed "IVT_objl";
+
+Goal "(f(b) <= y & y <= f(a) & a <= b & \
+\     (ALL x. a <= x & x <= b --> isCont f x)) \
+\     --> (EX x. a <= x & x <= b & f(x) = y)";
+by (blast_tac (claset() addIs [IVT2]) 1);
+qed "IVT2_objl";
+
+(*----------------------------------------------------------------------------*)
+(* By bisection, function continuous on closed interval is bounded above      *)
+(*----------------------------------------------------------------------------*)
+
+Goal "abs (real_of_nat x) = real_of_nat x";
+by (auto_tac (claset() addIs [abs_eqI1],simpset()));
+qed "abs_real_of_nat_cancel";
+Addsimps [abs_real_of_nat_cancel];
+
+Goal "~ abs(x) + 1r < x";
+by (rtac real_leD 1);
+by (auto_tac (claset() addIs [abs_ge_self RS real_le_trans],simpset()));
+qed "abs_add_one_not_less_self";
+Addsimps [abs_add_one_not_less_self];
+
+
+Goal "[| a <= b; ALL x. a <= x & x <= b --> isCont f x |]\
+\     ==> EX M. ALL x. a <= x & x <= b --> f(x) <= M";
+by (cut_inst_tac [("P","%(u,v). a <= u & u <= v & v <= b --> \
+\                         (EX M. ALL x. u <= x & x <= v --> f x <= M)")] 
+    lemma_BOLZANO2 1);
+by (Step_tac 1);
+by (ALLGOALS(Asm_full_simp_tac));
+by (cut_inst_tac [("x","M"),("y","Ma")] 
+    (CLAIM "x <= y | y <= (x::real)") 1);
+by (Step_tac 1);
+by (res_inst_tac [("x","Ma")] exI 1);
+by (Step_tac 1);
+by (cut_inst_tac [("x","xb"),("y","xa")] 
+    (CLAIM "x <= y | y <= (x::real)") 1);
+by (Step_tac 1);
+by (rtac real_le_trans 1 THEN assume_tac 2);
+by (Asm_full_simp_tac 1);
+by (Asm_full_simp_tac 1);
+by (res_inst_tac [("x","M")] exI 1);
+by (Step_tac 1);
+by (cut_inst_tac [("x","xb"),("y","xa")] 
+    (CLAIM "x <= y | y <= (x::real)") 1);
+by (Step_tac 1);
+by (Asm_full_simp_tac 1);
+by (rtac real_le_trans 1 THEN assume_tac 2);
+by (Asm_full_simp_tac 1);
+by (case_tac "a <= x & x <= b" 1);
+by (res_inst_tac [("x","#1")] exI 2);
+by (auto_tac (claset(),
+              simpset() addsimps [LIM_def,isCont_iff]));
+by (dres_inst_tac [("x","x")] spec 1 THEN Auto_tac);
+by (thin_tac "ALL M. EX x. a <= x & x <= b & ~ f x <= M" 1);
+by (dres_inst_tac [("x","#1")] spec 1);
+by Auto_tac;  
+by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
+by (res_inst_tac [("x","abs(f x) + #1")] exI 1 THEN Step_tac 1);
+by (dres_inst_tac [("x","xa - x")] spec 1 THEN Auto_tac);
+by (res_inst_tac [("j","abs(f xa)")] real_le_trans 3);
+by (res_inst_tac [("j","abs(f x) + abs(f(xa) - f(x))")] real_le_trans 4);
+by (auto_tac (claset() addIs [abs_triangle_ineq RSN (2, real_le_trans)],
+              simpset() addsimps [real_diff_def,abs_ge_self]));
+(*arith_tac problem: this step should not be needed*)
+by (asm_full_simp_tac (simpset() addsimps [rename_numerals (real_eq_minus_iff2 RS sym)]) 1);
+by (auto_tac (claset(),
+              simpset() addsimps [abs_real_def] addsplits [split_if_asm]));
+qed "isCont_bounded";
+
+(*----------------------------------------------------------------------------*)
+(* Refine the above to existence of least upper bound                         *)
+(*----------------------------------------------------------------------------*)
+
+Goal "((EX x. x : S) & (EX y. isUb UNIV S (y::real))) --> \
+\     (EX t. isLub UNIV S t)";
+by (blast_tac (claset() addIs [reals_complete]) 1);
+qed "lemma_reals_complete";
+
+Goal "[| a <= b; ALL x. a <= x & x <= b --> isCont f x |] \
+\        ==> EX M. (ALL x. a <= x & x <= b --> f(x) <= M) & \
+\                  (ALL N. N < M --> (EX x. a <= x & x <= b & N < f(x)))";
+by (cut_inst_tac [("S","Collect (%y. EX x. a <= x & x <= b & y = f x)")]
+    lemma_reals_complete 1);
+by (Auto_tac);
+by (dtac isCont_bounded 1 THEN assume_tac 1);
+by (auto_tac (claset(),simpset() addsimps [isUb_def,leastP_def,
+    isLub_def,setge_def,setle_def]));
+by (rtac exI 1 THEN Auto_tac);
+by (REPEAT(dtac spec 1) THEN Auto_tac);
+by (dres_inst_tac [("x","x")] spec 1);
+by (auto_tac (claset() addSIs [real_leI],simpset()));
+qed "isCont_has_Ub";
+
+(*----------------------------------------------------------------------------*)
+(* Now show that it attains its upper bound                                   *)
+(*----------------------------------------------------------------------------*)
+
+
+
+
+ *)
\ No newline at end of file
--- a/src/HOL/Real/Hyperreal/Lim.thy	Tue Dec 12 11:59:25 2000 +0100
+++ b/src/HOL/Real/Hyperreal/Lim.thy	Tue Dec 12 12:01:19 2000 +0100
@@ -11,12 +11,15 @@
          Limits, continuity and differentiation: standard and NS definitions
       -----------------------------------------------------------------------*)
 constdefs
-      LIM :: [real=>real,real,real] => bool    ("((_)/ -- (_)/ --> (_))" 60)
-      "f -- a --> L == (ALL r. #0 < r --> 
-                          (EX s. #0 < s & (ALL x. (#0 < abs(x + -a) & (abs(x + -a) < s)
-                                --> abs(f x + -L) < r))))"
+      LIM :: [real=>real,real,real] => bool
+                                    ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60)
+      "f -- a --> L ==
+         ALL r. #0 < r --> 
+                 (EX s. #0 < s & (ALL x. (x ~= a & (abs(x + -a) < s)
+                              --> abs(f x + -L) < r)))"
 
-      NSLIM :: [real=>real,real,real] => bool  ("((_)/ -- (_)/ --NS> (_))" 60)
+      NSLIM :: [real=>real,real,real] => bool
+                                  ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60)
       "f -- a --NS> L == (ALL x. (x ~= hypreal_of_real a & 
                           x @= hypreal_of_real a --> (*f* f) x @= hypreal_of_real L))"   
 
@@ -29,10 +32,12 @@
                                (*f* f) y @= hypreal_of_real (f a))"
       
       (* differentiation: D is derivative of function f at x *)
-      deriv:: [real=>real,real,real] => bool   ("(DERIV (_)/ (_)/ :> (_))" 60)
+      deriv:: [real=>real,real,real] => bool
+                                ("(DERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60)
       "DERIV f x :> D == ((%h. (f(x + h) + -f(x))*inverse(h)) -- #0 --> D)"
 
-      nsderiv :: [real=>real,real,real] => bool   ("(NSDERIV (_)/ (_)/ :> (_))" 60)
+      nsderiv :: [real=>real,real,real] => bool
+                                ("(NSDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60)
       "NSDERIV f x :> D == (ALL h: Infinitesimal - {0}. 
                             ((*f* f)(hypreal_of_real x + h) + 
                              -hypreal_of_real (f x))*inverse(h) @= hypreal_of_real D)"
--- a/src/HOL/Real/Hyperreal/NSA.ML	Tue Dec 12 11:59:25 2000 +0100
+++ b/src/HOL/Real/Hyperreal/NSA.ML	Tue Dec 12 12:01:19 2000 +0100
@@ -2060,10 +2060,10 @@
          Infinitesimals as smaller than 1/n for all n::nat (> 0)
  ------------------------------------------------------------------------*)
 Goal "(ALL r. 0 < r --> x < r) = (ALL n. x < inverse(real_of_posnat n))";
-by (auto_tac (claset(),simpset() addsimps 
-    [rename_numerals (real_of_posnat_gt_zero RS real_inverse_gt_zero)]));
+by (auto_tac (claset(),
+        simpset() addsimps [rename_numerals real_of_posnat_gt_zero]));
 by (blast_tac (claset() addSDs [rename_numerals reals_Archimedean] 
-    addIs [real_less_trans]) 1);
+                        addIs [order_less_trans]) 1);
 qed "lemma_Infinitesimal";
 
 Goal "(ALL r: SReal. 0 < r --> x < r) = \
@@ -2075,15 +2075,15 @@
           (hypreal_of_real_less_iff RS iffD1) RSN(2,impE)) 1);
 by (assume_tac 2);
 by (asm_full_simp_tac (simpset() addsimps 
-   [real_not_refl2 RS not_sym RS hypreal_of_real_inverse RS sym,
-   rename_numerals real_of_posnat_gt_zero,hypreal_of_real_zero,
-    hypreal_of_real_of_posnat]) 1);
+       [real_not_refl2 RS not_sym RS hypreal_of_real_inverse RS sym,
+        rename_numerals real_of_posnat_gt_zero,hypreal_of_real_zero,
+        hypreal_of_real_of_posnat]) 1);
 by (auto_tac (claset() addSDs [reals_Archimedean],
-    simpset() addsimps [SReal_iff,hypreal_of_real_zero RS sym]));
+              simpset() addsimps [SReal_iff,hypreal_of_real_zero RS sym]));
 by (dtac (hypreal_of_real_less_iff RS iffD1) 1);
 by (asm_full_simp_tac (simpset() addsimps 
-   [real_not_refl2 RS not_sym RS hypreal_of_real_inverse RS sym,
-   rename_numerals real_of_posnat_gt_zero,hypreal_of_real_of_posnat])1);
+       [real_not_refl2 RS not_sym RS hypreal_of_real_inverse RS sym,
+        rename_numerals real_of_posnat_gt_zero,hypreal_of_real_of_posnat])1);
 by (blast_tac (claset() addIs [hypreal_less_trans]) 1);
 qed "lemma_Infinitesimal2";
 
@@ -2117,7 +2117,7 @@
 by (cut_inst_tac [("x","u")] reals_Archimedean2 1);
 by (Step_tac 1);
 by (rtac (finite_real_of_posnat_segment RSN (2,finite_subset)) 1);
-by (auto_tac (claset() addDs [real_less_trans],simpset()));
+by (auto_tac (claset() addDs [order_less_trans],simpset()));
 qed "finite_real_of_posnat_less_real";
 
 Goal "{n. real_of_posnat n <= u} = \
@@ -2265,7 +2265,7 @@
 \    ==> Abs_hypreal(hyprel^^{X}) + -hypreal_of_real x : Infinitesimal";
 by (auto_tac (claset() addSIs [bexI] addDs [rename_numerals 
     FreeUltrafilterNat_inverse_real_of_posnat,
-    FreeUltrafilterNat_all,FreeUltrafilterNat_Int] addIs [real_less_trans,
+    FreeUltrafilterNat_all,FreeUltrafilterNat_Int] addIs [order_less_trans,
     FreeUltrafilterNat_subset],simpset() addsimps [hypreal_minus,
     hypreal_of_real_minus RS sym,hypreal_of_real_def,hypreal_add,
     Infinitesimal_FreeUltrafilterNat_iff,hypreal_inverse,hypreal_of_real_of_posnat]));
@@ -2287,11 +2287,13 @@
 Goal "ALL n. abs(X n + -Y n) < inverse(real_of_posnat n) \ 
 \     ==> Abs_hypreal(hyprel^^{X}) + \
 \         -Abs_hypreal(hyprel^^{Y}) : Infinitesimal";
-by (auto_tac (claset() addSIs [bexI] addDs [rename_numerals 
-    FreeUltrafilterNat_inverse_real_of_posnat,
-    FreeUltrafilterNat_all,FreeUltrafilterNat_Int] addIs [real_less_trans,
-    FreeUltrafilterNat_subset],simpset() addsimps 
-    [Infinitesimal_FreeUltrafilterNat_iff,hypreal_minus,hypreal_add, 
-    hypreal_inverse,hypreal_of_real_of_posnat]));
+by (auto_tac (claset() addSIs [bexI] 
+                  addDs [rename_numerals 
+                         FreeUltrafilterNat_inverse_real_of_posnat,
+                         FreeUltrafilterNat_all,FreeUltrafilterNat_Int] 
+        addIs [order_less_trans, FreeUltrafilterNat_subset],
+     simpset() addsimps 
+            [Infinitesimal_FreeUltrafilterNat_iff,hypreal_minus,hypreal_add, 
+             hypreal_inverse,hypreal_of_real_of_posnat]));
 qed "real_seq_to_hypreal_Infinitesimal2";
  
\ No newline at end of file
--- a/src/HOL/Real/Hyperreal/SEQ.ML	Tue Dec 12 11:59:25 2000 +0100
+++ b/src/HOL/Real/Hyperreal/SEQ.ML	Tue Dec 12 12:01:19 2000 +0100
@@ -1155,11 +1155,11 @@
 by (etac NSLIMSEQ_imp_Suc 1);
 qed "LIMSEQ_imp_Suc";
 
-Goal "(%n. f(Suc n) ----> l) = (f ----> l)";
+Goal "((%n. f(Suc n)) ----> l) = (f ----> l)";
 by (blast_tac (claset() addIs [LIMSEQ_imp_Suc,LIMSEQ_Suc]) 1);
 qed "LIMSEQ_Suc_iff";
 
-Goal "(%n. f(Suc n) ----NS> l) = (f ----NS> l)";
+Goal "((%n. f(Suc n)) ----NS> l) = (f ----NS> l)";
 by (blast_tac (claset() addIs [NSLIMSEQ_imp_Suc,NSLIMSEQ_Suc]) 1);
 qed "NSLIMSEQ_Suc_iff";
 
@@ -1215,7 +1215,8 @@
 by (forw_inst_tac [("x","f n")] (rename_numerals real_inverse_gt_zero) 1);
 by (asm_simp_tac (simpset() addsimps [abs_eqI2]) 1);
 by (res_inst_tac [("t","r")] (real_inverse_inverse RS subst) 1);
-by (auto_tac (claset() addIs [real_inverse_less_iff RS iffD1], simpset()));
+by (auto_tac (claset() addIs [real_inverse_less_iff RS iffD2], 
+              simpset() delsimps [real_inverse_inverse]));
 qed "LIMSEQ_inverse_zero";
 
 Goal "ALL y. EX N. ALL n. N <= n --> y < f(n) \
--- a/src/HOL/Real/Hyperreal/SEQ.thy	Tue Dec 12 11:59:25 2000 +0100
+++ b/src/HOL/Real/Hyperreal/SEQ.thy	Tue Dec 12 12:01:19 2000 +0100
@@ -9,11 +9,11 @@
 constdefs
 
   (* Standard definition of convergence of sequence *)           
-  LIMSEQ :: [nat=>real,real] => bool    ("((_)/ ----> (_))" 60)
+  LIMSEQ :: [nat=>real,real] => bool    ("((_)/ ----> (_))" [60, 60] 60)
   "X ----> L == (ALL r. #0 < r --> (EX no. ALL n. no <= n --> abs (X n + -L) < r))"
   
   (* Nonstandard definition of convergence of sequence *)
-  NSLIMSEQ :: [nat=>real,real] => bool    ("((_)/ ----NS> (_))" 60)
+  NSLIMSEQ :: [nat=>real,real] => bool    ("((_)/ ----NS> (_))" [60, 60] 60)
   "X ----NS> L == (ALL N: HNatInfinite. (*fNat* X) N @= hypreal_of_real L)"
 
   (* define value of limit using choice operator*)
--- a/src/HOL/Real/RealAbs.ML	Tue Dec 12 11:59:25 2000 +0100
+++ b/src/HOL/Real/RealAbs.ML	Tue Dec 12 12:01:19 2000 +0100
@@ -90,15 +90,14 @@
 	      simpset() addsimps [real_0_le_mult_iff]));
 qed "abs_mult";
 
-Goalw [abs_real_def] "x~= (#0::real) ==> abs(inverse(x)) = inverse(abs(x))";
+Goalw [abs_real_def] "abs(inverse(x::real)) = inverse(abs(x))";
+by (real_div_undefined_case_tac "x=0" 1);
 by (auto_tac (claset(), 
-              simpset() addsimps [real_le_less] @ 
-	    (map rename_numerals [real_minus_inverse, real_inverse_gt_zero])));
-by (etac (rename_numerals (real_inverse_less_zero RSN (2,real_less_asym))) 1);
-by (arith_tac 1);
+              simpset() addsimps [real_minus_inverse, real_le_less] @ 
+	          (map rename_numerals [INVERSE_ZERO, real_inverse_gt_zero])));
 qed "abs_inverse";
 
-Goal "y ~= #0 ==> abs (x * inverse y) = (abs x) * inverse (abs (y::real))";
+Goal "abs (x * inverse y) = (abs x) * inverse (abs (y::real))";
 by (asm_simp_tac (simpset() addsimps [abs_mult, abs_inverse]) 1);
 qed "abs_mult_inverse";
 
@@ -234,6 +233,12 @@
 qed "abs_le_zero_iff";
 Addsimps [abs_le_zero_iff];
 
+Goal "((#0::real) < abs x) = (x ~= 0)";
+by (simp_tac (simpset() addsimps [abs_real_def]) 1);
+by (arith_tac 1);
+qed "real_0_less_abs_iff";
+Addsimps [real_0_less_abs_iff];
+
 Goal "abs (real_of_nat x) = real_of_nat x";
 by (auto_tac (claset() addIs [abs_eqI1],simpset()
     addsimps [rename_numerals real_of_nat_ge_zero]));
--- a/src/HOL/Real/RealArith.ML	Tue Dec 12 11:59:25 2000 +0100
+++ b/src/HOL/Real/RealArith.ML	Tue Dec 12 12:01:19 2000 +0100
@@ -1,5 +1,372 @@
-(*useful??*)
-Goal "(z = z + w) = (w = (#0::real))";
-by Auto_tac;
-qed "real_add_left_cancel0";
-Addsimps [real_add_left_cancel0];
+(*  Title:      HOL/Real/RealArith.ML
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1999  University of Cambridge
+
+Assorted facts that need binary literals and the arithmetic decision procedure
+
+Also, common factor cancellation
+*)
+
+(** Division and inverse **)
+
+Goal "((#0::real) < inverse x) = (#0 < x)";
+by (case_tac "x=#0" 1);
+by (asm_simp_tac (HOL_ss addsimps [rename_numerals INVERSE_ZERO]) 1); 
+by (auto_tac (claset() addDs [rename_numerals real_inverse_less_zero], 
+              simpset() addsimps [linorder_neq_iff, 
+                                  rename_numerals real_inverse_gt_zero]));  
+qed "real_0_less_inverse_iff";
+AddIffs [real_0_less_inverse_iff];
+
+Goal "(inverse x < (#0::real)) = (x < #0)";
+by (case_tac "x=#0" 1);
+by (asm_simp_tac (HOL_ss addsimps [rename_numerals INVERSE_ZERO]) 1); 
+by (auto_tac (claset() addDs [rename_numerals real_inverse_less_zero], 
+              simpset() addsimps [linorder_neq_iff, 
+                                  rename_numerals real_inverse_gt_zero]));  
+qed "real_inverse_less_0_iff";
+AddIffs [real_inverse_less_0_iff];
+
+Goalw [real_divide_def] "x/(#0::real) = #0";
+by (stac (rename_numerals INVERSE_ZERO) 1); 
+by (Simp_tac 1); 
+qed "REAL_DIVIDE_ZERO";
+
+(*generalize?*)
+Goal "((#0::real) < #1/x) = (#0 < x)";
+by (simp_tac (simpset() addsimps [real_divide_def]) 1);
+qed "real_0_less_recip_iff";
+AddIffs [real_0_less_recip_iff];
+
+Goal "(#1/x < (#0::real)) = (x < #0)";
+by (simp_tac (simpset() addsimps [real_divide_def]) 1);
+qed "real_recip_less_0_iff";
+AddIffs [real_recip_less_0_iff];
+
+Goal "inverse (x::real) = #1/x";
+by (simp_tac (simpset() addsimps [real_divide_def]) 1); 
+qed "real_inverse_eq_divide";
+
+Goal "(x::real)/#1 = x";
+by (simp_tac (simpset() addsimps [real_divide_def]) 1); 
+qed "real_divide_1";
+Addsimps [real_divide_1];
+
+
+(**** Factor cancellation theorems for "real" ****)
+
+(** Cancellation laws for k*m < k*n and m*k < n*k, also for <= and =,
+    but not (yet?) for k*m < n*k. **)
+
+bind_thm ("real_minus_less_minus", real_less_swap_iff RS sym);
+bind_thm ("real_mult_minus_right", real_minus_mult_eq2 RS sym);
+
+Goal "[| i<j;  k < (0::real) |] ==> j*k < i*k";
+by (rtac (real_minus_less_minus RS iffD1) 1);
+by (auto_tac (claset(), 
+              simpset() delsimps [real_minus_mult_eq2 RS sym]
+                        addsimps [real_minus_mult_eq2])); 
+qed "real_mult_less_mono1_neg";
+
+Goal "[| i<j;  k < (0::real) |] ==> k*j < k*i";
+by (rtac (real_minus_less_minus RS iffD1) 1);
+by (auto_tac (claset(), 
+              simpset() delsimps [real_minus_mult_eq1 RS sym]
+                            addsimps [real_minus_mult_eq1]));;
+qed "real_mult_less_mono2_neg";
+
+Goal "[| i <= j;  (0::real) <= k |] ==> i*k <= j*k";
+by (auto_tac (claset(), 
+              simpset() addsimps [order_le_less, real_mult_less_mono1]));  
+qed "real_mult_le_mono1";
+
+Goal "[| i <= j;  k <= (0::real) |] ==> j*k <= i*k";
+by (auto_tac (claset(), 
+              simpset() addsimps [order_le_less, real_mult_less_mono1_neg]));  
+qed "real_mult_le_mono1_neg";
+
+Goal "[| i <= j;  (0::real) <= k |] ==> k*i <= k*j";
+by (dtac real_mult_le_mono1 1);
+by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [real_mult_commute])));
+qed "real_mult_le_mono2";
+
+Goal "[| i <= j;  k <= (0::real) |] ==> k*j <= k*i";
+by (dtac real_mult_le_mono1_neg 1);
+by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [real_mult_commute])));
+qed "real_mult_le_mono2_neg";
+
+Goal "(m*k < n*k) = (((#0::real) < k & m<n) | (k < #0 & n<m))";
+by (case_tac "k = (0::real)" 1);
+by (auto_tac (claset(), 
+              simpset() addsimps [linorder_neq_iff, 
+                          real_mult_less_mono1, real_mult_less_mono1_neg]));  
+by (auto_tac (claset(), 
+              simpset() addsimps [linorder_not_less,
+				  inst "y1" "m*k" (linorder_not_le RS sym),
+                                  inst "y1" "m" (linorder_not_le RS sym)]));
+by (TRYALL (etac notE));
+by (auto_tac (claset(), 
+              simpset() addsimps [order_less_imp_le, real_mult_le_mono1,
+                                            real_mult_le_mono1_neg]));  
+qed "real_mult_less_cancel2";
+
+Goal "(m*k <= n*k) = (((#0::real) < k --> m<=n) & (k < #0 --> n<=m))";
+by (simp_tac (simpset() addsimps [linorder_not_less RS sym, 
+                                  real_mult_less_cancel2]) 1);
+qed "real_mult_le_cancel2";
+
+Goal "(k*m < k*n) = (((#0::real) < k & m<n) | (k < #0 & n<m))";
+by (simp_tac (simpset() addsimps [inst "z" "k" real_mult_commute, 
+                                  real_mult_less_cancel2]) 1);
+qed "real_mult_less_cancel1";
+
+Goal "!!k::real. (k*m <= k*n) = ((#0 < k --> m<=n) & (k < #0 --> n<=m))";
+by (simp_tac (simpset() addsimps [linorder_not_less RS sym, 
+                                  real_mult_less_cancel1]) 1);
+qed "real_mult_le_cancel1";
+
+Goal "!!k::real. (k*m = k*n) = (k = #0 | m=n)";
+by (case_tac "k=0" 1);
+by (auto_tac (claset(), simpset() addsimps [real_mult_left_cancel]));  
+qed "real_mult_eq_cancel1";
+
+Goal "!!k::real. (m*k = n*k) = (k = #0 | m=n)";
+by (case_tac "k=0" 1);
+by (auto_tac (claset(), simpset() addsimps [real_mult_right_cancel]));  
+qed "real_mult_eq_cancel2";
+
+Goal "!!k::real. k~=#0 ==> (k*m) / (k*n) = (m/n)";
+by (asm_simp_tac
+    (simpset() addsimps [real_divide_def, real_inverse_distrib]) 1); 
+by (subgoal_tac "k * m * (inverse k * inverse n) = \
+\                (k * inverse k) * (m * inverse n)" 1);
+by (asm_full_simp_tac (simpset() addsimps []) 1); 
+by (asm_full_simp_tac (HOL_ss addsimps real_mult_ac) 1); 
+qed "real_mult_div_cancel1";
+
+
+local
+  open Real_Numeral_Simprocs
+in
+
+val rel_real_number_of = [eq_real_number_of, less_real_number_of, 
+                          le_real_number_of_eq_not_less];
+
+structure CancelNumeralFactorCommon =
+  struct
+  val mk_coeff		= mk_coeff
+  val dest_coeff	= dest_coeff 1
+  val trans_tac         = trans_tac
+  val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_plus_1s))
+            THEN ALLGOALS
+                 (simp_tac 
+                  (HOL_ss addsimps [eq_real_number_of, 
+                                    real_mult_minus_right RS sym]@
+    [mult_real_number_of, real_mult_number_of_left]@bin_simps@real_mult_ac))
+  val numeral_simp_tac	= 
+         ALLGOALS (simp_tac (HOL_ss addsimps rel_real_number_of@bin_simps))
+  val simplify_meta_eq  = simplify_meta_eq
+  end
+
+structure DivCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = prove_conv "realdiv_cancel_numeral_factor"
+  val mk_bal   = HOLogic.mk_binop "HOL.divide"
+  val dest_bal = HOLogic.dest_bin "HOL.divide" HOLogic.realT
+  val cancel = real_mult_div_cancel1 RS trans
+  val neg_exchanges = false
+)
+
+structure EqCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = prove_conv "realeq_cancel_numeral_factor"
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" HOLogic.realT
+  val cancel = real_mult_eq_cancel1 RS trans
+  val neg_exchanges = false
+)
+
+structure LessCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = prove_conv "realless_cancel_numeral_factor"
+  val mk_bal   = HOLogic.mk_binrel "op <"
+  val dest_bal = HOLogic.dest_bin "op <" HOLogic.realT
+  val cancel = real_mult_less_cancel1 RS trans
+  val neg_exchanges = true
+)
+
+structure LeCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = prove_conv "realle_cancel_numeral_factor"
+  val mk_bal   = HOLogic.mk_binrel "op <="
+  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.realT
+  val cancel = real_mult_le_cancel1 RS trans
+  val neg_exchanges = true
+)
+
+val real_cancel_numeral_factors = 
+  map prep_simproc
+   [("realeq_cancel_numeral_factors",
+     prep_pats ["(l::real) * m = n", "(l::real) = m * n"], 
+     EqCancelNumeralFactor.proc),
+    ("realless_cancel_numeral_factors", 
+     prep_pats ["(l::real) * m < n", "(l::real) < m * n"], 
+     LessCancelNumeralFactor.proc),
+    ("realle_cancel_numeral_factors", 
+     prep_pats ["(l::real) * m <= n", "(l::real) <= m * n"], 
+     LeCancelNumeralFactor.proc),
+    ("realdiv_cancel_numeral_factors", 
+     prep_pats ["((l::real) * m) / n", "(l::real) / (m * n)"], 
+     DivCancelNumeralFactor.proc)];
+
+end;
+
+Addsimprocs real_cancel_numeral_factors;
+
+
+(*examples:
+print_depth 22;
+set timing;
+set trace_simp;
+fun test s = (Goal s; by (Simp_tac 1)); 
+
+test "#0 <= (y::real) * #-2";
+test "#9*x = #12 * (y::real)";
+test "(#9*x) / (#12 * (y::real)) = z";
+test "#9*x < #12 * (y::real)";
+test "#9*x <= #12 * (y::real)";
+
+test "#-99*x = #132 * (y::real)";
+test "(#-99*x) / (#132 * (y::real)) = z";
+test "#-99*x < #132 * (y::real)";
+test "#-99*x <= #132 * (y::real)";
+
+test "#999*x = #-396 * (y::real)";
+test "(#999*x) / (#-396 * (y::real)) = z";
+test "#999*x < #-396 * (y::real)";
+test "#999*x <= #-396 * (y::real)";
+
+test "#-99*x = #-81 * (y::real)";
+test "(#-99*x) / (#-81 * (y::real)) = z";
+test "#-99*x <= #-81 * (y::real)";
+test "#-99*x < #-81 * (y::real)";
+
+test "#-2 * x = #-1 * (y::real)";
+test "#-2 * x = -(y::real)";
+test "(#-2 * x) / (#-1 * (y::real)) = z";
+test "#-2 * x < -(y::real)";
+test "#-2 * x <= #-1 * (y::real)";
+test "-x < #-23 * (y::real)";
+test "-x <= #-23 * (y::real)";
+*)
+
+(*** Simplification of inequalities involving literal divisors ***)
+
+Goal "#0<z ==> ((x::real) <= y/z) = (x*z <= y)";
+by (subgoal_tac "(x*z <= y) = (x*z <= (y/z)*z)" 1);
+by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac real_mult_le_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "pos_real_le_divide_eq";
+Addsimps [inst "z" "number_of ?w" pos_real_le_divide_eq];
+
+Goal "z<#0 ==> ((x::real) <= y/z) = (y <= x*z)";
+by (subgoal_tac "(y <= x*z) = ((y/z)*z <= x*z)" 1);
+by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac real_mult_le_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "neg_real_le_divide_eq";
+Addsimps [inst "z" "number_of ?w" neg_real_le_divide_eq];
+
+Goal "#0<z ==> (y/z <= (x::real)) = (y <= x*z)";
+by (subgoal_tac "(y <= x*z) = ((y/z)*z <= x*z)" 1);
+by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac real_mult_le_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "pos_real_divide_le_eq";
+Addsimps [inst "z" "number_of ?w" pos_real_divide_le_eq];
+
+Goal "z<#0 ==> (y/z <= (x::real)) = (x*z <= y)";
+by (subgoal_tac "(x*z <= y) = (x*z <= (y/z)*z)" 1);
+by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac real_mult_le_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "neg_real_divide_le_eq";
+Addsimps [inst "z" "number_of ?w" neg_real_divide_le_eq];
+
+Goal "#0<z ==> ((x::real) < y/z) = (x*z < y)";
+by (subgoal_tac "(x*z < y) = (x*z < (y/z)*z)" 1);
+by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac real_mult_less_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "pos_real_less_divide_eq";
+Addsimps [inst "z" "number_of ?w" pos_real_less_divide_eq];
+
+Goal "z<#0 ==> ((x::real) < y/z) = (y < x*z)";
+by (subgoal_tac "(y < x*z) = ((y/z)*z < x*z)" 1);
+by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac real_mult_less_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "neg_real_less_divide_eq";
+Addsimps [inst "z" "number_of ?w" neg_real_less_divide_eq];
+
+Goal "#0<z ==> (y/z < (x::real)) = (y < x*z)";
+by (subgoal_tac "(y < x*z) = ((y/z)*z < x*z)" 1);
+by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac real_mult_less_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "pos_real_divide_less_eq";
+Addsimps [inst "z" "number_of ?w" pos_real_divide_less_eq];
+
+Goal "z<#0 ==> (y/z < (x::real)) = (x*z < y)";
+by (subgoal_tac "(x*z < y) = (x*z < (y/z)*z)" 1);
+by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac real_mult_less_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "neg_real_divide_less_eq";
+Addsimps [inst "z" "number_of ?w" neg_real_divide_less_eq];
+
+Goal "z~=#0 ==> ((x::real) = y/z) = (x*z = y)";
+by (subgoal_tac "(x*z = y) = (x*z = (y/z)*z)" 1);
+by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac real_mult_eq_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "real_eq_divide_eq";
+Addsimps [inst "z" "number_of ?w" real_eq_divide_eq];
+
+Goal "z~=#0 ==> (y/z = (x::real)) = (y = x*z)";
+by (subgoal_tac "(y = x*z) = ((y/z)*z = x*z)" 1);
+by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2); 
+by (etac ssubst 1);
+by (stac real_mult_eq_cancel2 1); 
+by (Asm_simp_tac 1); 
+qed "real_divide_eq_eq";
+Addsimps [inst "z" "number_of ?w" real_divide_eq_eq];
+
+
+(** Moved from RealOrd.ML to use #0 **)
+
+Goal "[| #0 < r; #0 < x|] ==> (inverse x < inverse (r::real)) = (r < x)";
+by (auto_tac (claset() addIs [real_inverse_less_swap], simpset()));
+by (res_inst_tac [("t","r")] (real_inverse_inverse RS subst) 1);
+by (res_inst_tac [("t","x")] (real_inverse_inverse RS subst) 1);
+by (auto_tac (claset() addIs [real_inverse_less_swap],
+	      simpset() delsimps [real_inverse_inverse]
+			addsimps [real_inverse_gt_zero]));
+qed "real_inverse_less_iff";
+
+Goal "[| #0 < r; #0 < x|] ==> (inverse x <= inverse r) = (r <= (x::real))";
+by (asm_simp_tac (simpset() addsimps [linorder_not_less RS sym, 
+                                      real_inverse_less_iff]) 1); 
+qed "real_inverse_le_iff";
+
--- a/src/HOL/Real/RealBin.ML	Tue Dec 12 11:59:25 2000 +0100
+++ b/src/HOL/Real/RealBin.ML	Tue Dec 12 12:01:19 2000 +0100
@@ -344,14 +344,16 @@
 (*Simplify #1*n and n*#1 to n*)
 val add_0s = map rename_numerals
                  [real_add_zero_left, real_add_zero_right];
-val mult_1s = map rename_numerals
-                  [real_mult_1, real_mult_1_right, 
-		   real_mult_minus_1, real_mult_minus_1_right];
+val mult_plus_1s = map rename_numerals
+                      [real_mult_1, real_mult_1_right];
+val mult_minus_1s = map rename_numerals
+                      [real_mult_minus_1, real_mult_minus_1_right];
+val mult_1s = mult_plus_1s @ mult_minus_1s;
 
 (*To perform binary arithmetic*)
 val bin_simps =
     [add_real_number_of, real_add_number_of_left, minus_real_number_of, 
-     diff_real_number_of] @ 
+     diff_real_number_of, mult_real_number_of, real_mult_number_of_left] @ 
     bin_arith_simps @ bin_rel_simps;
 
 (*To evaluate binary negations of coefficients*)
--- a/src/HOL/Real/RealDef.ML	Tue Dec 12 11:59:25 2000 +0100
+++ b/src/HOL/Real/RealDef.ML	Tue Dec 12 12:01:19 2000 +0100
@@ -488,9 +488,9 @@
     @ preal_add_ac @ preal_mult_ac));
 qed "real_mult_inv_right_ex";
 
-Goal "!!(x::real). x ~= 0 ==> EX y. y*x = 1r";
-by (asm_simp_tac (simpset() addsimps [real_mult_commute,
-				      real_mult_inv_right_ex]) 1);
+Goal "x ~= 0 ==> EX y. y*x = 1r";
+by (dtac real_mult_inv_right_ex 1); 
+by (auto_tac (claset(), simpset() addsimps [real_mult_commute]));  
 qed "real_mult_inv_left_ex";
 
 Goalw [real_inverse_def] "x ~= 0 ==> inverse(x)*x = 1r";
@@ -499,23 +499,40 @@
 by (rtac someI2 1);
 by Auto_tac;
 qed "real_mult_inv_left";
+Addsimps [real_mult_inv_left];
 
 Goal "x ~= 0 ==> x*inverse(x) = 1r";
-by (auto_tac (claset() addIs [real_mult_commute RS subst],
-              simpset() addsimps [real_mult_inv_left]));
+by (stac real_mult_commute 1);
+by (auto_tac (claset(), simpset() addsimps [real_mult_inv_left]));
 qed "real_mult_inv_right";
+Addsimps [real_mult_inv_right];
+
+(** Inverse of zero!  Useful to simplify certain equations **)
+
+Goalw [real_inverse_def] "inverse 0 = (0::real)";
+by (rtac someI2 1);
+by (auto_tac (claset(), simpset() addsimps [real_zero_not_eq_one]));  
+qed "INVERSE_ZERO";
+
+Goal "a / (0::real) = 0";
+by (simp_tac (simpset() addsimps [real_divide_def, INVERSE_ZERO]) 1);
+qed "DIVISION_BY_ZERO";  (*NOT for adding to default simpset*)
+
+fun real_div_undefined_case_tac s i =
+  case_tac s i THEN 
+  asm_simp_tac (simpset() addsimps [DIVISION_BY_ZERO, INVERSE_ZERO]) i;
+
 
 Goal "(c::real) ~= 0 ==> (c*a=c*b) = (a=b)";
 by Auto_tac;
 by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1);
-by (asm_full_simp_tac (simpset() addsimps [real_mult_inv_right] @ real_mult_ac)  1);
+by (asm_full_simp_tac (simpset() addsimps real_mult_ac)  1);
 qed "real_mult_left_cancel";
     
 Goal "(c::real) ~= 0 ==> (a*c=b*c) = (a=b)";
 by (Step_tac 1);
 by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1);
-by (asm_full_simp_tac
-    (simpset() addsimps [real_mult_inv_right] @ real_mult_ac)  1);
+by (asm_full_simp_tac (simpset() addsimps real_mult_ac)  1);
 qed "real_mult_right_cancel";
 
 Goal "c*a ~= c*b ==> a ~= b";
@@ -531,12 +548,9 @@
 by (etac exE 1);
 by (rtac someI2 1);
 by (auto_tac (claset(),
-	      simpset() addsimps [real_mult_0,
-    real_zero_not_eq_one]));
+	      simpset() addsimps [real_mult_0, real_zero_not_eq_one]));
 qed "real_inverse_not_zero";
 
-Addsimps [real_mult_inv_left,real_mult_inv_right];
-
 Goal "[| x ~= 0; y ~= 0 |] ==> x * y ~= (0::real)";
 by (Step_tac 1);
 by (dres_inst_tac [("f","%z. inverse x*z")] arg_cong 1);
@@ -545,11 +559,13 @@
 
 bind_thm ("real_mult_not_zeroE",real_mult_not_zero RS notE);
 
-Goal "x ~= 0 ==> inverse(inverse (x::real)) = x";
+Goal "inverse(inverse (x::real)) = x";
+by (real_div_undefined_case_tac "x=0" 1);
 by (res_inst_tac [("c1","inverse x")] (real_mult_right_cancel RS iffD1) 1);
 by (etac real_inverse_not_zero 1);
 by (auto_tac (claset() addDs [real_inverse_not_zero],simpset()));
 qed "real_inverse_inverse";
+Addsimps [real_inverse_inverse];
 
 Goalw [real_inverse_def] "inverse(1r) = 1r";
 by (cut_facts_tac [real_zero_not_eq_one RS 
@@ -559,13 +575,16 @@
 qed "real_inverse_1";
 Addsimps [real_inverse_1];
 
-Goal "x ~= 0 ==> inverse(-x) = -inverse(x::real)";
+Goal "inverse(-x) = -inverse(x::real)";
+by (real_div_undefined_case_tac "x=0" 1);
 by (res_inst_tac [("c1","-x")] (real_mult_right_cancel RS iffD1) 1);
 by (stac real_mult_inv_left 2);
 by Auto_tac;
 qed "real_minus_inverse";
 
-Goal "[| x ~= 0; y ~= 0 |] ==> inverse(x*y) = inverse(x)*inverse(y::real)";
+Goal "inverse(x*y) = inverse(x)*inverse(y::real)";
+by (real_div_undefined_case_tac "x=0" 1);
+by (real_div_undefined_case_tac "y=0" 1);
 by (forw_inst_tac [("y","y")] real_mult_not_zero 1 THEN assume_tac 1);
 by (res_inst_tac [("c1","x")] (real_mult_left_cancel RS iffD1) 1);
 by (auto_tac (claset(),simpset() addsimps [real_mult_assoc RS sym]));
@@ -574,6 +593,52 @@
 by (asm_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 1);
 qed "real_inverse_distrib";
 
+Goal "(x::real) * (y/z) = (x*y)/z";
+by (simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 1); 
+qed "real_times_divide1_eq";
+
+Goal "(y/z) * (x::real) = (y*x)/z";
+by (simp_tac (simpset() addsimps [real_divide_def]@real_mult_ac) 1); 
+qed "real_times_divide2_eq";
+
+Addsimps [real_times_divide1_eq, real_times_divide2_eq];
+
+Goal "(x::real) / (y/z) = (x*z)/y";
+by (simp_tac (simpset() addsimps [real_divide_def, real_inverse_distrib]@
+                                 real_mult_ac) 1); 
+qed "real_divide_divide1_eq";
+
+Goal "((x::real) / y) / z = x/(y*z)";
+by (simp_tac (simpset() addsimps [real_divide_def, real_inverse_distrib, 
+                                  real_mult_assoc]) 1); 
+qed "real_divide_divide2_eq";
+
+Addsimps [real_divide_divide1_eq, real_divide_divide2_eq];
+
+(** As with multiplication, pull minus signs OUT of the / operator **)
+
+Goal "(-x) / (y::real) = - (x/y)";
+by (simp_tac (simpset() addsimps [real_divide_def]) 1); 
+qed "real_minus_divide_eq";
+Addsimps [real_minus_divide_eq];
+
+Goal "(x / -(y::real)) = - (x/y)";
+by (simp_tac (simpset() addsimps [real_divide_def, real_minus_inverse]) 1); 
+qed "real_divide_minus_eq";
+Addsimps [real_divide_minus_eq];
+
+Goal "(x+y)/(z::real) = x/z + y/z";
+by (simp_tac (simpset() addsimps [real_divide_def, real_add_mult_distrib]) 1); 
+qed "real_add_divide_distrib";
+
+(*The following would e.g. convert (x+y)/2 to x/2 + y/2.  Sometimes that
+  leads to cancellations in x or y, but is also prevents "multiplying out"
+  the 2 in e.g. (x+y)/2 = 5.
+  
+Addsimps [inst "z" "number_of ?w" real_add_divide_distrib];
+**)
+
+
 (*---------------------------------------------------------
      Theorems for ordering
  --------------------------------------------------------*)
--- a/src/HOL/Real/RealDef.thy	Tue Dec 12 11:59:25 2000 +0100
+++ b/src/HOL/Real/RealDef.thy	Tue Dec 12 12:01:19 2000 +0100
@@ -37,7 +37,7 @@
   "R - (S::real) == R + - S"
 
   real_inverse_def
-  "inverse (R::real) == (@S. R ~= 0 & S*R = 1r)"
+  "inverse (R::real) == (SOME S. (R = 0 & S = 0) | S * R = 1r)"
 
   real_divide_def
   "R / (S::real) == R * inverse S"
--- a/src/HOL/Real/RealOrd.ML	Tue Dec 12 11:59:25 2000 +0100
+++ b/src/HOL/Real/RealOrd.ML	Tue Dec 12 12:01:19 2000 +0100
@@ -419,7 +419,7 @@
 by (ftac real_not_refl2 1);
 by (dtac (real_minus_zero_less_iff RS iffD2) 1);
 by (rtac (real_minus_zero_less_iff RS iffD1) 1);
-by (dtac (real_minus_inverse RS sym) 1);
+by (stac (real_minus_inverse RS sym) 1);
 by (auto_tac (claset() addIs [real_inverse_gt_zero], simpset()));
 qed "real_inverse_less_zero";
 
@@ -503,14 +503,6 @@
 Addsimps [real_mult_less_iff1,real_mult_less_iff2];
 
 (* 05/00 *)
-Goalw [real_le_def] "[| (0::real) < z; x*z<=y*z |] ==> x<=y";
-by (Auto_tac);
-qed "real_mult_le_cancel1";
-
-Goalw [real_le_def] "[| (0::real) < z; z*x<=z*y |] ==> x<=y";
-by (Auto_tac);
-qed "real_mult_le_cancel2";
-
 Goalw [real_le_def] "(0::real) < z ==> (x*z <= y*z) = (x <= y)";
 by (Auto_tac);
 qed "real_mult_le_cancel_iff1";
@@ -709,16 +701,6 @@
          not_sym RS real_mult_inv_left,real_mult_assoc RS sym]) 1);
 qed "real_inverse_less_swap";
 
-Goal "[| 0 < r; 0 < x|] ==> (r < x) = (inverse x < inverse (r::real))";
-by (auto_tac (claset() addIs [real_inverse_less_swap],simpset()));
-by (res_inst_tac [("t","r")] (real_inverse_inverse RS subst) 1);
-by (etac (real_not_refl2 RS not_sym) 1);
-by (res_inst_tac [("t","x")] (real_inverse_inverse RS subst) 1);
-by (etac (real_not_refl2 RS not_sym) 1);
-by (auto_tac (claset() addIs [real_inverse_less_swap],
-	      simpset() addsimps [real_inverse_gt_zero]));
-qed "real_inverse_less_iff";
-
 Goal "r < r + inverse (real_of_posnat n)";
 by (res_inst_tac [("C","-r")] real_less_add_left_cancel 1);
 by (full_simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1);
--- a/src/HOL/Real/RealPow.ML	Tue Dec 12 11:59:25 2000 +0100
+++ b/src/HOL/Real/RealPow.ML	Tue Dec 12 12:01:19 2000 +0100
@@ -20,13 +20,10 @@
 by (auto_tac (claset() addDs [realpow_not_zero], simpset()));
 qed "realpow_zero_zero";
 
-Goal "(r::real) ~= #0 --> inverse (r ^ n) = (inverse r) ^ n";
+Goal "inverse ((r::real) ^ n) = (inverse r) ^ n";
 by (induct_tac "n" 1);
-by (Auto_tac);
-by (forw_inst_tac [("n","n")] realpow_not_zero 1);
-by (auto_tac (claset() addDs [rename_numerals real_inverse_distrib],
-	      simpset()));
-qed_spec_mp "realpow_inverse";
+by (auto_tac (claset(), simpset() addsimps [real_inverse_distrib]));
+qed "realpow_inverse";
 
 Goal "abs (r::real) ^ n = abs (r ^ n)";
 by (induct_tac "n" 1);
@@ -414,12 +411,6 @@
 Addsimps [realpow_real_of_nat_two_pos];
 
 
-(***	REALORD.ML, AFTER real_inverse_less_iff ***)
-Goal "[| 0 < r; 0 < x|] ==> (inverse x <= inverse r) = (r <= (x::real))";
-by (asm_simp_tac (simpset() addsimps [linorder_not_less RS sym, 
-                                      real_inverse_less_iff]) 1); 
-qed "real_inverse_le_iff";
-
 Goal "(#0::real) <= x --> #0 <= y --> x ^ Suc n <= y ^ Suc n --> x <= y";
 by (induct_tac "n" 1);
 by Auto_tac;