replacing HOL/Real/PRat, PNat by the rational number development
authorpaulson
Tue, 27 Jan 2004 15:39:51 +0100
changeset 14365 3d4df8c166ae
parent 14364 fc62df0bf353
child 14366 dd4e0f2c071a
replacing HOL/Real/PRat, PNat by the rational number development of Markus Wenzel
src/HOL/Complex/CLim.ML
src/HOL/Complex/NSCA.ML
src/HOL/HOL.thy
src/HOL/Hyperreal/HRealAbs.ML
src/HOL/Hyperreal/HRealAbs.thy
src/HOL/Hyperreal/HyperDef.thy
src/HOL/Hyperreal/IntFloor.ML
src/HOL/Hyperreal/Integration.ML
src/HOL/Hyperreal/Lim.ML
src/HOL/Hyperreal/Log.ML
src/HOL/Hyperreal/MacLaurin.ML
src/HOL/Hyperreal/NSA.ML
src/HOL/Hyperreal/NthRoot.thy
src/HOL/Hyperreal/Poly.ML
src/HOL/Hyperreal/SEQ.ML
src/HOL/Hyperreal/Series.ML
src/HOL/Hyperreal/Transcendental.ML
src/HOL/Integ/Bin.thy
src/HOL/Integ/Equiv.thy
src/HOL/Integ/Int.thy
src/HOL/Integ/IntArith.thy
src/HOL/Integ/NatBin.thy
src/HOL/Integ/nat_simprocs.ML
src/HOL/IsaMakefile
src/HOL/Library/Library.thy
src/HOL/Library/Rational_Numbers.thy
src/HOL/MicroJava/BV/Kildall.thy
src/HOL/Real/PNat.ML
src/HOL/Real/PNat.thy
src/HOL/Real/PRat.ML
src/HOL/Real/PRat.thy
src/HOL/Real/PReal.thy
src/HOL/Real/RComplete.ML
src/HOL/Real/RComplete.thy
src/HOL/Real/RatArith.thy
src/HOL/Real/Rational.thy
src/HOL/Real/RealArith.thy
src/HOL/Real/RealBin.ML
src/HOL/Real/RealBin.thy
src/HOL/Real/RealDef.thy
src/HOL/Real/RealInt.thy
src/HOL/Real/rat_arith.ML
src/HOL/Real/real_arith.ML
src/HOL/Ring_and_Field.thy
--- a/src/HOL/Complex/CLim.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Complex/CLim.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -83,7 +83,7 @@
     CInfinitesimal_hcmod_iff,hcomplex_of_complex_def,
     Infinitesimal_FreeUltrafilterNat_iff,hcmod]));
 by (EVERY1[rtac ccontr, Asm_full_simp_tac]);
-by (fold_tac [real_le_def]);
+by (asm_full_simp_tac (simpset() addsimps [linorder_not_less]) 1);
 by (dtac lemma_skolemize_CLIM2 1);
 by (Step_tac 1);
 by (dres_inst_tac [("x","X")] spec 1);
@@ -159,7 +159,7 @@
     CInfinitesimal_hcmod_iff,hcmod,Infinitesimal_approx_minus RS sym,
     Infinitesimal_FreeUltrafilterNat_iff]));
 by (EVERY1[rtac ccontr, Asm_full_simp_tac]);
-by (fold_tac [real_le_def]);
+by (asm_full_simp_tac (simpset() addsimps [linorder_not_less]) 1);
 by (dtac lemma_skolemize_CRLIM2 1);
 by (Step_tac 1);
 by (dres_inst_tac [("x","X")] spec 1);
--- a/src/HOL/Complex/NSCA.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Complex/NSCA.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -863,17 +863,18 @@
 
 Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) : CFinite \
 \     ==> Abs_hypreal(hyprel `` {%n. Re(X n)}) : HFinite";
-by (auto_tac (claset(),simpset() addsimps [CFinite_hcmod_iff,
-    hcmod,HFinite_FreeUltrafilterNat_iff]));
+by (auto_tac (claset(),
+ simpset() addsimps [CFinite_hcmod_iff,hcmod,HFinite_FreeUltrafilterNat_iff]));
 by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
 by (res_inst_tac [("x","u")] exI 1 THEN Auto_tac);
 by (Ultra_tac 1);
 by (dtac sym 1 THEN res_inst_tac [("z","X x")] eq_Abs_complex 1);
-by (auto_tac (claset(),simpset() addsimps [complex_mod,numeral_2_eq_2] delsimps [realpow_Suc]));
+by (auto_tac (claset(),
+    simpset() addsimps [complex_mod,numeral_2_eq_2] delsimps [realpow_Suc]));
 by (rtac ccontr 1 THEN dtac (linorder_not_less RS iffD1) 1);
 by (dtac order_less_le_trans 1 THEN assume_tac 1);
 by (dtac (real_sqrt_ge_abs1 RSN (2,order_less_le_trans)) 1);
-by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2 RS sym]));
+by (auto_tac ((claset(),simpset() addsimps [numeral_2_eq_2 RS sym]) addIffs [order_less_irrefl]));
 qed "CFinite_HFinite_Re";
 
 Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) : CFinite \
@@ -888,7 +889,7 @@
 by (rtac ccontr 1 THEN dtac (linorder_not_less RS iffD1) 1);
 by (dtac order_less_le_trans 1 THEN assume_tac 1);
 by (dtac (real_sqrt_ge_abs2 RSN (2,order_less_le_trans)) 1);
-by Auto_tac;
+by (auto_tac (clasimpset() addIffs [order_less_irrefl]));
 qed "CFinite_HFinite_Im";
 
 Goal "[| Abs_hypreal(hyprel `` {%n. Re(X n)}) : HFinite; \
--- a/src/HOL/HOL.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/HOL.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -804,10 +804,13 @@
   apply (insert linorder_linear, blast)
   done
 
+lemma linorder_le_cases [case_names le ge]:
+    "((x::'a::linorder) \<le> y ==> P) ==> (y \<le> x ==> P) ==> P"
+  by (insert linorder_linear, blast)
+
 lemma linorder_cases [case_names less equal greater]:
     "((x::'a::linorder) < y ==> P) ==> (x = y ==> P) ==> (y < x ==> P) ==> P"
-  apply (insert linorder_less_linear, blast)
-  done
+  by (insert linorder_less_linear, blast)
 
 lemma linorder_not_less: "!!x::'a::linorder. (~ x < y) = (y <= x)"
   apply (simp add: order_less_le)
--- a/src/HOL/Hyperreal/HRealAbs.ML	Tue Jan 27 09:44:14 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,138 +0,0 @@
-(*  Title       : HRealAbs.ML
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Absolute value function for the hyperreals
-                  Similar to RealAbs.thy
-*) 
-
-(*------------------------------------------------------------
-  absolute value on hyperreals as pointwise operation on 
-  equivalence class representative
- ------------------------------------------------------------*)
-
-Goalw [hrabs_def]
-     "abs (number_of v :: hypreal) = \
-\       (if neg (number_of v) then number_of (bin_minus v) \
-\        else number_of v)";
-by (Simp_tac 1); 
-qed "hrabs_number_of";
-Addsimps [hrabs_number_of];
-
-(*------------------------------------------------------------
-   Properties of the absolute value function over the reals
-   (adapted version of previously proved theorems about abs)
- ------------------------------------------------------------*)
-
-Goal "(0::hypreal)<=x ==> abs x = x";
-by (asm_simp_tac (simpset() addsimps [hrabs_def]) 1); 
-qed "hrabs_eqI1";
-
-Goal "(0::hypreal)<x ==> abs x = x";
-by (asm_simp_tac (simpset() addsimps [order_less_imp_le, hrabs_eqI1]) 1);
-qed "hrabs_eqI2";
-
-Goal "x<(0::hypreal) ==> abs x = -x";
-by (asm_simp_tac (simpset() addsimps [hypreal_le_def, hrabs_def]) 1); 
-qed "hrabs_minus_eqI2";
-
-Goal "x<=(0::hypreal) ==> abs x = -x";
-by (auto_tac (claset() addDs [order_antisym], simpset() addsimps [hrabs_def]));qed "hrabs_minus_eqI1";
-
-Addsimps [abs_mult];
-
-Goalw [hrabs_def] "[| abs x < r; abs y < s |] ==> abs(x+y) < r + (s::hypreal)";
-by (asm_full_simp_tac (simpset() addsplits [split_if_asm]) 1); 
-qed "hrabs_add_less";
-
-Goal "abs x < r ==> (0::hypreal) < r";
-by (blast_tac (claset() addSIs [order_le_less_trans, abs_ge_zero]) 1);
-qed "hrabs_less_gt_zero";
-
-Goal "abs x = (x::hypreal) | abs x = -x";
-by (simp_tac (simpset() addsimps [hrabs_def]) 1); 
-qed "hrabs_disj";
-
-Goal "abs x = (y::hypreal) ==> x = y | -x = y";
-by (asm_full_simp_tac (simpset() addsimps [hrabs_def] 
-                                 addsplits [split_if_asm]) 1); 
-qed "hrabs_eq_disj";
-
-(* Needed in Geom.ML *)
-Goal "(y::hypreal) + - x + (y + - z) = abs (x + - z) ==> y = z | x = y";
-by (asm_full_simp_tac (simpset() addsimps [hrabs_def] 
-                                 addsplits [split_if_asm]) 1); 
-qed "hrabs_add_lemma_disj";
-
-(* Needed in Geom.ML?? *)
-Goal "(x::hypreal) + - y + (z + - y) = abs (x + - z) ==> y = z | x = y";
-by (asm_full_simp_tac (simpset() addsimps [hrabs_def] 
-                                 addsplits [split_if_asm]) 1); 
-qed "hrabs_add_lemma_disj2";
-
- 
-(*----------------------------------------------------------
-    Relating hrabs to abs through embedding of IR into IR*
- ----------------------------------------------------------*)
-Goalw [hypreal_of_real_def] 
-    "abs (hypreal_of_real r) = hypreal_of_real (abs r)";
-by (auto_tac (claset(), simpset() addsimps [hypreal_hrabs]));
-qed "hypreal_of_real_hrabs";
-
-
-(*----------------------------------------------------------------------------
-             Embedding of the naturals in the hyperreals
- ----------------------------------------------------------------------------*)
-
-Goal "hypreal_of_nat (m + n) = hypreal_of_nat m + hypreal_of_nat n";
-by (simp_tac (simpset() addsimps [hypreal_of_nat_def]) 1);
-qed "hypreal_of_nat_add";
-Addsimps [hypreal_of_nat_add];
-
-Goal "hypreal_of_nat (m * n) = hypreal_of_nat m * hypreal_of_nat n";
-by (simp_tac (simpset() addsimps [hypreal_of_nat_def]) 1);
-qed "hypreal_of_nat_mult";
-Addsimps [hypreal_of_nat_mult];
-
-Goalw [hypreal_of_nat_def] 
-      "(n < m) = (hypreal_of_nat n < hypreal_of_nat m)";
-by (auto_tac (claset() addIs [hypreal_add_less_mono1], simpset()));
-qed "hypreal_of_nat_less_iff";
-Addsimps [hypreal_of_nat_less_iff RS sym];
-
-(*------------------------------------------------------------*)
-(* naturals embedded in hyperreals                            *)
-(* is a hyperreal c.f. NS extension                           *)
-(*------------------------------------------------------------*)
-
-Goalw [hypreal_of_nat_def, hypreal_of_real_def, real_of_nat_def] 
-     "hypreal_of_nat  m = Abs_hypreal(hyprel``{%n. real m})";
-by Auto_tac;
-qed "hypreal_of_nat_iff";
-
-Goal "inj hypreal_of_nat";
-by (simp_tac (simpset() addsimps [inj_on_def, hypreal_of_nat_def]) 1);
-qed "inj_hypreal_of_nat";
-
-Goalw [hypreal_of_nat_def] 
-     "hypreal_of_nat (Suc n) = hypreal_of_nat n + (1::hypreal)";
-by (simp_tac (simpset() addsimps [real_of_nat_Suc]) 1);
-qed "hypreal_of_nat_Suc";
-
-(*"neg" is used in rewrite rules for binary comparisons*)
-Goal "hypreal_of_nat (number_of v :: nat) = \
-\        (if neg (number_of v) then 0 \
-\         else (number_of v :: hypreal))";
-by (simp_tac (simpset() addsimps [hypreal_of_nat_def]) 1);
-qed "hypreal_of_nat_number_of";
-Addsimps [hypreal_of_nat_number_of];
-
-Goal "hypreal_of_nat 0 = 0";
-by (simp_tac (simpset() delsimps [numeral_0_eq_0]
-                        addsimps [numeral_0_eq_0 RS sym]) 1);
-qed "hypreal_of_nat_zero";
-Addsimps [hypreal_of_nat_zero];
-
-Goal "hypreal_of_nat 1 = 1";
-by (simp_tac (simpset() addsimps [hypreal_of_nat_Suc]) 1); 
-qed "hypreal_of_nat_one";
-Addsimps [hypreal_of_nat_one];
--- a/src/HOL/Hyperreal/HRealAbs.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Hyperreal/HRealAbs.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -4,11 +4,146 @@
     Description : Absolute value function for the hyperreals
 *) 
 
-HRealAbs = HyperArith + RealArith + 
+theory HRealAbs = HyperArith:
 
 constdefs
-  
-  hypreal_of_nat :: nat => hypreal                   
+
+  hypreal_of_nat :: "nat => hypreal"
   "hypreal_of_nat (n::nat) == hypreal_of_real (real n)"
 
-end
\ No newline at end of file
+
+lemma hrabs_number_of [simp]:
+     "abs (number_of v :: hypreal) =
+        (if neg (number_of v) then number_of (bin_minus v)
+         else number_of v)"
+by (simp add: hrabs_def)
+
+
+(*------------------------------------------------------------
+   Properties of the absolute value function over the reals
+   (adapted version of previously proved theorems about abs)
+ ------------------------------------------------------------*)
+
+lemma hrabs_eqI1: "(0::hypreal)<=x ==> abs x = x"
+by (simp add: hrabs_def)
+
+lemma hrabs_eqI2: "(0::hypreal)<x ==> abs x = x"
+by (simp add: order_less_imp_le hrabs_eqI1)
+
+lemma hrabs_minus_eqI2: "x<(0::hypreal) ==> abs x = -x"
+by (simp add: hypreal_le_def hrabs_def)
+
+lemma hrabs_minus_eqI1: "x<=(0::hypreal) ==> abs x = -x"
+by (auto dest: order_antisym simp add: hrabs_def)
+
+declare abs_mult [simp]
+
+lemma hrabs_add_less: "[| abs x < r; abs y < s |] ==> abs(x+y) < r + (s::hypreal)"
+apply (unfold hrabs_def)
+apply (simp split add: split_if_asm)
+done
+
+lemma hrabs_less_gt_zero: "abs x < r ==> (0::hypreal) < r"
+by (blast intro!: order_le_less_trans abs_ge_zero)
+
+lemma hrabs_disj: "abs x = (x::hypreal) | abs x = -x"
+by (simp add: hrabs_def)
+
+lemma hrabs_eq_disj: "abs x = (y::hypreal) ==> x = y | -x = y"
+by (simp add: hrabs_def split add: split_if_asm)
+
+(* Needed in Geom.ML *)
+lemma hrabs_add_lemma_disj: "(y::hypreal) + - x + (y + - z) = abs (x + - z) ==> y = z | x = y"
+by (simp add: hrabs_def split add: split_if_asm)
+
+(* Needed in Geom.ML?? *)
+lemma hrabs_add_lemma_disj2: "(x::hypreal) + - y + (z + - y) = abs (x + - z) ==> y = z | x = y"
+by (simp add: hrabs_def split add: split_if_asm)
+
+
+(*----------------------------------------------------------
+    Relating hrabs to abs through embedding of IR into IR*
+ ----------------------------------------------------------*)
+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
+
+
+(*----------------------------------------------------------------------------
+             Embedding of the naturals in the hyperreals
+ ----------------------------------------------------------------------------*)
+
+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)
+
+lemma hypreal_of_nat_mult: "hypreal_of_nat (m * n) = hypreal_of_nat m * hypreal_of_nat n"
+by (simp add: hypreal_of_nat_def)
+declare hypreal_of_nat_mult [simp]
+
+lemma hypreal_of_nat_less_iff:
+      "(n < m) = (hypreal_of_nat n < hypreal_of_nat m)"
+apply (simp add: hypreal_of_nat_def)
+done
+declare hypreal_of_nat_less_iff [symmetric, simp]
+
+(*------------------------------------------------------------*)
+(* naturals embedded in hyperreals                            *)
+(* is a hyperreal c.f. NS extension                           *)
+(*------------------------------------------------------------*)
+
+lemma hypreal_of_nat_iff:
+     "hypreal_of_nat  m = Abs_hypreal(hyprel``{%n. real m})"
+by (simp add: hypreal_of_nat_def hypreal_of_real_def real_of_nat_def)
+
+lemma inj_hypreal_of_nat: "inj hypreal_of_nat"
+by (simp add: inj_on_def hypreal_of_nat_def)
+
+lemma hypreal_of_nat_Suc:
+     "hypreal_of_nat (Suc n) = hypreal_of_nat n + (1::hypreal)"
+by (simp add: hypreal_of_nat_def real_of_nat_Suc)
+
+(*"neg" is used in rewrite rules for binary comparisons*)
+lemma hypreal_of_nat_number_of [simp]:
+     "hypreal_of_nat (number_of v :: nat) =
+         (if neg (number_of v) then 0
+          else (number_of v :: hypreal))"
+by (simp add: hypreal_of_nat_def)
+
+lemma hypreal_of_nat_zero [simp]: "hypreal_of_nat 0 = 0"
+by (simp del: numeral_0_eq_0 add: numeral_0_eq_0 [symmetric])
+
+lemma hypreal_of_nat_one [simp]: "hypreal_of_nat 1 = 1"
+by (simp add: hypreal_of_nat_Suc)
+
+
+ML
+{*
+val hypreal_of_nat_def = thm"hypreal_of_nat_def";
+
+val hrabs_number_of = thm "hrabs_number_of";
+val hrabs_eqI1 = thm "hrabs_eqI1";
+val hrabs_eqI2 = thm "hrabs_eqI2";
+val hrabs_minus_eqI2 = thm "hrabs_minus_eqI2";
+val hrabs_minus_eqI1 = thm "hrabs_minus_eqI1";
+val hrabs_add_less = thm "hrabs_add_less";
+val hrabs_less_gt_zero = thm "hrabs_less_gt_zero";
+val hrabs_disj = thm "hrabs_disj";
+val hrabs_eq_disj = thm "hrabs_eq_disj";
+val hrabs_add_lemma_disj = thm "hrabs_add_lemma_disj";
+val hrabs_add_lemma_disj2 = thm "hrabs_add_lemma_disj2";
+val hypreal_of_real_hrabs = thm "hypreal_of_real_hrabs";
+val hypreal_of_nat_add = thm "hypreal_of_nat_add";
+val hypreal_of_nat_mult = thm "hypreal_of_nat_mult";
+val hypreal_of_nat_less_iff = thm "hypreal_of_nat_less_iff";
+val hypreal_of_nat_iff = thm "hypreal_of_nat_iff";
+val inj_hypreal_of_nat = thm "inj_hypreal_of_nat";
+val hypreal_of_nat_Suc = thm "hypreal_of_nat_Suc";
+val hypreal_of_nat_number_of = thm "hypreal_of_nat_number_of";
+val hypreal_of_nat_zero = thm "hypreal_of_nat_zero";
+val hypreal_of_nat_one = thm "hypreal_of_nat_one";
+*}
+
+end
--- a/src/HOL/Hyperreal/HyperDef.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Hyperreal/HyperDef.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -82,7 +82,7 @@
                                Y \<in> Rep_hypreal(Q) &
                                {n::nat. X n < Y n} \<in> FreeUltrafilterNat"
   hypreal_le_def:
-  "P <= (Q::hypreal) == ~(Q < P)"
+  "P \<le> (Q::hypreal) == ~(Q < P)"
 
   hrabs_def:  "abs (r::hypreal) == (if 0 \<le> r then r else -r)"
 
@@ -149,7 +149,7 @@
 done
 
 lemma FreeUltrafilterNat_subset:
-     "[| X: FreeUltrafilterNat;  X <= Y |]  
+     "[| X: FreeUltrafilterNat;  X \<subseteq> Y |]  
       ==> Y \<in> FreeUltrafilterNat"
 apply (cut_tac FreeUltrafilterNat_mem)
 apply (blast dest: FreeUltrafilter_Ultrafilter Ultrafilter_Filter mem_FiltersetD2)
@@ -227,21 +227,20 @@
 
 text{*Proving that @{term hyprel} is an equivalence relation*}
 
-lemma hyprel_iff: "((X,Y): hyprel) = ({n. X n = Y n}: FreeUltrafilterNat)"
+lemma hyprel_iff: "((X,Y) \<in> hyprel) = ({n. X n = Y n}: FreeUltrafilterNat)"
 by (unfold hyprel_def, fast)
 
-lemma hyprel_refl: "(x,x): hyprel"
+lemma hyprel_refl: "(x,x) \<in> hyprel"
 apply (unfold hyprel_def)
 apply (auto simp add: FreeUltrafilterNat_Nat_set)
 done
 
-lemma hyprel_sym [rule_format (no_asm)]: "(x,y): hyprel --> (y,x):hyprel"
+lemma hyprel_sym [rule_format (no_asm)]: "(x,y) \<in> hyprel --> (y,x) \<in> hyprel"
 by (simp add: hyprel_def eq_commute)
 
 lemma hyprel_trans: 
-      "[|(x,y): hyprel; (y,z):hyprel|] ==> (x,z):hyprel"
-apply (unfold hyprel_def, auto, ultra)
-done
+      "[|(x,y) \<in> hyprel; (y,z) \<in> hyprel|] ==> (x,z) \<in> hyprel"
+by (unfold hyprel_def, auto, ultra)
 
 lemma equiv_hyprel: "equiv UNIV hyprel"
 apply (simp add: equiv_def refl_def sym_def trans_def hyprel_refl)
@@ -536,7 +535,7 @@
 
 subsection{*Trichotomy: the hyperreals are Linearly Ordered*}
 
-lemma lemma_hyprel_0_mem: "\<exists>x. x: hyprel `` {%n. 0}"
+lemma lemma_hyprel_0_mem: "\<exists>x. x \<in> hyprel `` {%n. 0}"
 apply (unfold hyprel_def)
 apply (rule_tac x = "%n. 0" in exI, safe)
 apply (auto intro!: FreeUltrafilterNat_Nat_set)
@@ -588,64 +587,57 @@
 lemma hypreal_neq_iff: "((w::hypreal) \<noteq> z) = (w<z | z<w)"
 by (cut_tac hypreal_linear, blast)
 
-lemma hypreal_linear_less2: "!!(x::hypreal). [| x < y ==> P;  x = y ==> P;  
-           y < x ==> P |] ==> P"
-apply (cut_tac x = x and y = y in hypreal_linear, auto)
-done
-
 
 subsection{*Properties of The @{text "\<le>"} Relation*}
 
 lemma hypreal_le: 
-      "(Abs_hypreal(hyprel``{%n. X n}) <=  
-            Abs_hypreal(hyprel``{%n. Y n})) =  
-       ({n. X n <= Y n} \<in> FreeUltrafilterNat)"
-apply (unfold hypreal_le_def real_le_def)
-apply (auto simp add: hypreal_less)
+      "(Abs_hypreal(hyprel``{%n. X n}) \<le> Abs_hypreal(hyprel``{%n. Y n})) =  
+       ({n. X n \<le> Y n} \<in> FreeUltrafilterNat)"
+apply (auto simp add: hypreal_less hypreal_le_def linorder_not_less[symmetric])
 apply (ultra+)
 done
 
-lemma hypreal_le_imp_less_or_eq: "!!(x::hypreal). x <= y ==> x < y | x = y"
+lemma hypreal_le_imp_less_or_eq: "!!(x::hypreal). x \<le> y ==> x < y | x = y"
 apply (unfold hypreal_le_def)
 apply (cut_tac hypreal_linear)
 apply (fast elim: hypreal_less_irrefl hypreal_less_asym)
 done
 
-lemma hypreal_less_or_eq_imp_le: "z<w | z=w ==> z <=(w::hypreal)"
+lemma hypreal_less_or_eq_imp_le: "z<w | z=w ==> z \<le>(w::hypreal)"
 apply (unfold hypreal_le_def)
 apply (cut_tac hypreal_linear)
 apply (fast elim: hypreal_less_irrefl hypreal_less_asym)
 done
 
-lemma hypreal_le_eq_less_or_eq: "(x <= (y::hypreal)) = (x < y | x=y)"
+lemma hypreal_le_eq_less_or_eq: "(x \<le> (y::hypreal)) = (x < y | x=y)"
 by (blast intro!: hypreal_less_or_eq_imp_le dest: hypreal_le_imp_less_or_eq) 
 
 lemmas hypreal_le_less = hypreal_le_eq_less_or_eq
 
-lemma hypreal_le_refl: "w <= (w::hypreal)"
+lemma hypreal_le_refl: "w \<le> (w::hypreal)"
 by (simp add: hypreal_le_eq_less_or_eq)
 
 (* Axiom 'linorder_linear' of class 'linorder': *)
-lemma hypreal_le_linear: "(z::hypreal) <= w | w <= z"
+lemma hypreal_le_linear: "(z::hypreal) \<le> w | w \<le> z"
 apply (simp add: hypreal_le_less)
 apply (cut_tac hypreal_linear, blast)
 done
 
-lemma hypreal_le_trans: "[| i <= j; j <= k |] ==> i <= (k::hypreal)"
+lemma hypreal_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::hypreal)"
 apply (drule hypreal_le_imp_less_or_eq) 
 apply (drule hypreal_le_imp_less_or_eq) 
 apply (rule hypreal_less_or_eq_imp_le) 
 apply (blast intro: hypreal_less_trans) 
 done
 
-lemma hypreal_le_anti_sym: "[| z <= w; w <= z |] ==> z = (w::hypreal)"
+lemma hypreal_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::hypreal)"
 apply (drule hypreal_le_imp_less_or_eq) 
 apply (drule hypreal_le_imp_less_or_eq) 
 apply (fast elim: hypreal_less_irrefl hypreal_less_asym)
 done
 
 (* Axiom 'order_less_le' of class 'order': *)
-lemma hypreal_less_le: "((w::hypreal) < z) = (w <= z & w \<noteq> z)"
+lemma hypreal_less_le: "((w::hypreal) < z) = (w \<le> z & w \<noteq> z)"
 apply (simp add: hypreal_le_def hypreal_neq_iff)
 apply (blast intro: hypreal_less_asym)
 done
@@ -794,9 +786,8 @@
 done
 
 lemma hypreal_of_real_le_iff [simp]: 
-     "(hypreal_of_real z1 <= hypreal_of_real z2) = (z1 <= z2)"
-apply (unfold hypreal_le_def real_le_def, auto)
-done
+     "(hypreal_of_real z1 \<le> hypreal_of_real z2) = (z1 \<le> z2)"
+by (force simp add: hypreal_less hypreal_le_def linorder_not_less[symmetric])
 
 lemma hypreal_of_real_eq_iff [simp]:
      "(hypreal_of_real z1 = hypreal_of_real z2) = (z1 = z2)"
@@ -952,8 +943,6 @@
 val hypreal_eq_minus_iff3 = thm "hypreal_eq_minus_iff3";
 val hypreal_not_eq_minus_iff = thm "hypreal_not_eq_minus_iff";
 val hypreal_linear = thm "hypreal_linear";
-val hypreal_neq_iff = thm "hypreal_neq_iff";
-val hypreal_linear_less2 = thm "hypreal_linear_less2";
 val hypreal_le = thm "hypreal_le";
 val hypreal_le_imp_less_or_eq = thm "hypreal_le_imp_less_or_eq";
 val hypreal_le_eq_less_or_eq = thm "hypreal_le_eq_less_or_eq";
--- a/src/HOL/Hyperreal/IntFloor.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Hyperreal/IntFloor.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -20,15 +20,13 @@
 qed "number_of_less_real_of_int_iff2";
 Addsimps [number_of_less_real_of_int_iff2];
 
-Goalw [real_le_def,zle_def] 
-   "((number_of n) <= real (m::int)) = (number_of n <= m)";
-by Auto_tac;
+Goal "((number_of n) <= real (m::int)) = (number_of n <= m)";
+by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym]));
 qed "number_of_le_real_of_int_iff";
 Addsimps [number_of_le_real_of_int_iff];
 
-Goalw [real_le_def,zle_def] 
-   "(real (m::int) <= (number_of n)) = (m <= number_of n)";
-by Auto_tac;
+Goal "(real (m::int) <= (number_of n)) = (m <= number_of n)";
+by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym]));
 qed "number_of_le_real_of_int_iff2";
 Addsimps [number_of_le_real_of_int_iff2];
 
--- a/src/HOL/Hyperreal/Integration.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Hyperreal/Integration.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -20,7 +20,7 @@
 
 Goalw [partition_def] 
    "a <= b ==> partition(a,b)(%n. if n = 0 then a else b)";
-by (auto_tac (claset(),simpset() addsimps [real_le_less]));
+by (auto_tac (claset(),simpset() addsimps [order_le_less]));
 qed "partition_single";
 Addsimps [partition_single];
 
@@ -79,7 +79,7 @@
 by (ALLGOALS(dtac (ARITH_PROVE "Suc m <= n ==> m < n")));
 by (ALLGOALS(dtac less_le_trans));
 by (assume_tac 1 THEN assume_tac 2);
-by (ALLGOALS(blast_tac (claset() addIs [real_less_trans])));
+by (ALLGOALS(blast_tac (claset() addIs [order_less_trans])));
 qed_spec_mp "lemma_partition_lt_gen";
 
 Goal "m < n ==> EX d. n = m + Suc d";
@@ -97,7 +97,7 @@
 by (Blast_tac 1);
 by (blast_tac (claset() addSDs [leI] addDs 
     [(ARITH_PROVE "m <= n ==> m <= Suc n"),
-    le_less_trans,real_less_trans]) 1);
+    le_less_trans,order_less_trans]) 1);
 qed_spec_mp "partition_lt";
 
 Goal "partition(a,b) D ==> a <= b";
@@ -587,7 +587,7 @@
 by (auto_tac (claset() addSIs [Least_equality RS sym,partition_rhs],simpset()));
 by (rtac ccontr 1);
 by (dtac partition_ub_lt 1);
-by Auto_tac;
+by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym]));
 qed "partition_psize_Least";
 
 Goal "partition (a, c) D ==> ~ (EX n. c < D(n))";
--- a/src/HOL/Hyperreal/Lim.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Hyperreal/Lim.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -32,8 +32,7 @@
 by (REPEAT(dres_inst_tac [("x","r/2")] spec 1));
 by (Asm_full_simp_tac 1);
 by (Clarify_tac 1);
-by (res_inst_tac [("R1.0","s"),("R2.0","sa")] 
-    real_linear_less2 1);
+by (res_inst_tac [("x","s"),("y","sa")] linorder_cases 1);
 by (res_inst_tac [("x","s")] exI 1);
 by (res_inst_tac [("x","sa")] exI 2);
 by (res_inst_tac [("x","sa")] exI 3);
@@ -75,7 +74,7 @@
    Limit not zero
  --------------------------*)
 Goalw [LIM_def] "k \\<noteq> 0 ==> ~ ((%x. k) -- x --> 0)";
-by (res_inst_tac [("R1.0","k"),("R2.0","0")] real_linear_less2 1);
+by (res_inst_tac [("x","k"),("y","0")] linorder_cases 1);
 by (auto_tac (claset(), simpset() addsimps [real_abs_def]));
 by (res_inst_tac [("x","-k")] exI 1);
 by (res_inst_tac [("x","k")] exI 2);
@@ -116,8 +115,8 @@
 by (cut_facts_tac [real_zero_less_one] 1);
 by (asm_full_simp_tac (simpset() addsimps [abs_mult]) 1);
 by (Clarify_tac 1);
-by (res_inst_tac [("R1.0","s"),("R2.0","sa")] 
-    real_linear_less2 1);
+by (res_inst_tac [("x","s"),("y","sa")] 
+    linorder_cases 1);
 by (res_inst_tac [("x","s")] exI 1);
 by (res_inst_tac [("x","sa")] exI 2);
 by (res_inst_tac [("x","sa")] exI 3);
@@ -216,7 +215,7 @@
 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 (asm_full_simp_tac (simpset() addsimps [linorder_not_less]) 1);
 by (dtac lemma_skolemize_LIM2 1);
 by Safe_tac;
 by (dres_inst_tac [("x","Abs_hypreal(hyprel``{X})")] spec 1);
@@ -722,7 +721,7 @@
 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 (asm_full_simp_tac (simpset() addsimps [linorder_not_less]) 1);
 by (dtac lemma_skolemize_LIM2u 1);
 by Safe_tac;
 by (dres_inst_tac [("x","Abs_hypreal(hyprel``{X})")] spec 1);
@@ -1904,7 +1903,7 @@
 Goal "[| DERIV f x :> l; \
 \        \\<exists>d. 0 < d & (\\<forall>y. abs(x - y) < d --> f(y) \\<le> f(x)) |] \
 \     ==> l = 0";
-by (res_inst_tac [("R1.0","l"),("R2.0","0")] real_linear_less2 1);
+by (res_inst_tac [("x","l"),("y","0")] linorder_cases 1);
 by Safe_tac;
 by (dtac DERIV_left_dec 1);
 by (dtac DERIV_left_inc 3);
@@ -2113,13 +2112,13 @@
 qed "DERIV_isconst2";
 
 Goal "\\<forall>x. DERIV f x :> 0 ==> f(x) = f(y)";
-by (res_inst_tac [("R1.0","x"),("R2.0","y")] real_linear_less2 1);
+by (res_inst_tac [("x","x"),("y","y")] linorder_cases 1);
 by (rtac sym 1);
 by (auto_tac (claset() addIs [DERIV_isCont,DERIV_isconst_end],simpset()));
 qed "DERIV_isconst_all";
 
 Goal "[|a \\<noteq> b; \\<forall>x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b - a) * k";
-by (res_inst_tac [("R1.0","a"),("R2.0","b")] real_linear_less2 1);
+by (res_inst_tac [("x","a"),("y","b")] linorder_cases 1);
 by Auto_tac;
 by (ALLGOALS(dres_inst_tac [("f","f")] MVT));
 by (auto_tac (claset() addDs [DERIV_isCont,DERIV_unique],simpset() addsimps 
@@ -2148,7 +2147,7 @@
 (* Gallileo's "trick": average velocity = av. of end velocities *)
 Goal "[|a \\<noteq> (b::real); \\<forall>x. DERIV v x :> k|] \
 \     ==> v((a + b)/2) = (v a + v b)/2";
-by (res_inst_tac [("R1.0","a"),("R2.0","b")] real_linear_less2 1);
+by (res_inst_tac [("x","a"),("y","b")] linorder_cases 1);
 by Safe_tac;
 by (ftac DERIV_const_ratio_const2 1 THEN assume_tac 1);
 by (ftac DERIV_const_ratio_const2 2 THEN assume_tac 2);
--- a/src/HOL/Hyperreal/Log.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Hyperreal/Log.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -81,8 +81,8 @@
 qed "powr_less_cancel_iff";
 Addsimps [powr_less_cancel_iff];
 
-Goalw [real_le_def] "1 < x ==> (x powr a <= x powr b) = (a <= b)";
-by (Auto_tac);
+Goal "1 < x ==> (x powr a <= x powr b) = (a <= b)";
+by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym]));
 qed "powr_le_cancel_iff";
 Addsimps [powr_le_cancel_iff];
 
@@ -156,7 +156,7 @@
 Addsimps [log_less_cancel_iff];
 
 Goal "[| 1 < a; 0 < x; 0 < y |] ==> (log a x <= log a y) = (x <= y)";
-by (auto_tac (claset(),simpset() addsimps [real_le_def]));
+by (auto_tac (claset(),simpset() addsimps [linorder_not_less RS sym]));
 qed "log_le_cancel_iff";
 Addsimps [log_le_cancel_iff];
 
--- a/src/HOL/Hyperreal/MacLaurin.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Hyperreal/MacLaurin.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -63,8 +63,8 @@
 by (asm_simp_tac (HOL_ss addsimps 
     [CLAIM "(a::real) * (b * (c * d)) = (d * a) * (b * c)"]
      delsimps [realpow_Suc]) 2);
-by (stac real_mult_inv_left 2);
-by (stac real_mult_inv_left 3);
+by (stac left_inverse 2);
+by (stac left_inverse 3);
 by (rtac (real_not_refl2 RS not_sym) 2);
 by (etac zero_less_power 2);
 by (rtac real_of_nat_fact_not_zero 2);
@@ -345,7 +345,7 @@
 \     |] ==> EX t. 0 < abs t & abs t < abs x & \
 \              f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + \
 \                    (diff n t / real (fact n)) * x ^ n";
-by (res_inst_tac [("R1.0","x"),("R2.0","0")] real_linear_less2 1);
+by (res_inst_tac [("x","x"),("y","0")] linorder_cases 1);
 by (Blast_tac 2);
 by (dtac Maclaurin_minus 1);
 by (dtac Maclaurin 5);
--- a/src/HOL/Hyperreal/NSA.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Hyperreal/NSA.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -1088,7 +1088,7 @@
 Goal "(x::hypreal): Reals ==> isLub Reals {s. s: Reals & s < x} x";
 by (auto_tac (claset() addSIs [isLubI2,lemma_SReal_ub,setgeI], simpset()));
 by (ftac isUbD2a 1);
-by (res_inst_tac [("x","x"),("y","y")] hypreal_linear_less2 1);
+by (res_inst_tac [("x","x"),("y","y")] linorder_cases 1);
 by (auto_tac (claset() addSIs [order_less_imp_le], simpset()));
 by (EVERY1[dtac SReal_dense, assume_tac, assume_tac, Step_tac]);
 by (dres_inst_tac [("y","r")] isUbD 1);
@@ -1980,7 +1980,7 @@
 qed "lemma_Int_HI";
 
 Goal "{n. u < abs (X n)} Int {n. abs (X n) < (u::real)} = {}";
-by (auto_tac (claset() addIs [real_less_asym], simpset()));
+by (auto_tac (claset() addIs [order_less_asym], simpset()));
 qed "lemma_Int_HIa";
 
 Goal "EX X: Rep_hypreal x. ALL u. \
--- a/src/HOL/Hyperreal/NthRoot.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Hyperreal/NthRoot.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -45,7 +45,8 @@
 apply (assumption+)
 apply (drule real_le_trans , assumption)
 apply (drule_tac y = "y ^ n" in order_less_le_trans)
-apply (assumption , erule real_less_irrefl)
+apply (assumption)
+apply (simp); 
 apply (drule_tac n = "n" in zero_less_one [THEN realpow_less])
 apply auto
 done
@@ -53,8 +54,8 @@
 lemma nth_realpow_isLub_ex:
      "[| (0::real) < a; 0 < n |]  
       ==> \<exists>u. isLub (UNIV::real set) {x. x ^ n <= a & 0 < x} u"
-apply (blast intro: lemma_nth_realpow_isUb_ex lemma_nth_realpow_non_empty reals_complete)
-done
+by (blast intro: lemma_nth_realpow_isUb_ex lemma_nth_realpow_non_empty reals_complete)
+
  
 subsection{*First Half -- Lemmas First*}
 
@@ -62,7 +63,7 @@
      "isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u  
            ==> u + inverse(real (Suc k)) ~: {x. x ^ n <= a & 0 < x}"
 apply (safe , drule isLubD2 , blast)
-apply (simp add: real_le_def)
+apply (simp add: linorder_not_less [symmetric])
 done
 
 lemma lemma_nth_realpow_isLub_gt_zero:
@@ -78,8 +79,8 @@
          0 < a; 0 < n |] ==> ALL k. a <= (u + inverse(real (Suc k))) ^ n"
 apply (safe)
 apply (frule lemma_nth_realpow_seq , safe)
-apply (auto elim: real_less_asym simp add: real_le_def)
-apply (simp add: real_le_def [symmetric])
+apply (auto elim: order_less_asym simp add: linorder_not_less [symmetric])
+apply (simp add: linorder_not_less)
 apply (rule order_less_trans [of _ 0])
 apply (auto intro: lemma_nth_realpow_isLub_gt_zero)
 done
@@ -103,14 +104,14 @@
 apply (drule isLub_le_isUb)
 apply assumption
 apply (drule order_less_le_trans)
-apply (auto simp add: real_less_not_refl)
+apply (auto)
 done
 
 lemma not_isUb_less_ex:
      "~ isUb (UNIV::real set) S u ==> \<exists>x \<in> S. u < x"
 apply (rule ccontr , erule swap)
 apply (rule setleI [THEN isUbI])
-apply (auto simp add: real_le_def)
+apply (auto simp add: linorder_not_less [symmetric])
 done
 
 lemma real_mult_less_self: "0 < r ==> r * (1 + -inverse(real (Suc n))) < r"
@@ -174,7 +175,7 @@
 apply auto
 apply (drule_tac x = "r" in realpow_less)
 apply (drule_tac [4] x = "y" in realpow_less)
-apply (auto simp add: real_less_not_refl)
+apply (auto)
 done
 
 ML
--- a/src/HOL/Hyperreal/Poly.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Hyperreal/Poly.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -226,7 +226,7 @@
 \     ==> EX x. a < x & x < b & (poly p x = 0)";
 by (cut_inst_tac [("f","%x. poly p x"),("a","a"),("b","b"),("y","0")] 
     IVT_objl 1);
-by (auto_tac (claset(),simpset() addsimps [real_le_less]));
+by (auto_tac (claset(),simpset() addsimps [order_le_less]));
 qed "poly_IVT_pos";
 
 Goal "[| a < b; 0 < poly p a; poly p b < 0 |] \
--- a/src/HOL/Hyperreal/SEQ.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Hyperreal/SEQ.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -146,7 +146,7 @@
 by (dtac choice 1 THEN Step_tac 1);
 by (dres_inst_tac [("x","Abs_hypnat(hypnatrel``{f})")] bspec 1);
 by (dtac (approx_minus_iff RS iffD1) 2);
-by (fold_tac [real_le_def]);
+by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [linorder_not_less])));
 by (blast_tac (claset() addIs [HNatInfinite_NSLIMSEQ]) 1);
 by (blast_tac (claset() addIs [lemmaLIM3]) 1);
 qed "NSLIMSEQ_LIMSEQ";
@@ -528,7 +528,7 @@
 Goalw [Bseq_def,NSBseq_def] 
       "NSBseq X ==> Bseq X";
 by (rtac ccontr 1);
-by (auto_tac (claset(), simpset() addsimps [real_le_def]));
+by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym]));
 by (dtac lemmaNSBseq2 1 THEN Step_tac 1);
 by (forw_inst_tac [("X","X"),("f","f")] real_seq_to_hypreal_HInfinite 1);
 by (dtac (HNatInfinite_skolem_f RSN (2,bspec)) 1 THEN assume_tac 1);
@@ -658,7 +658,7 @@
 \              |] ==> EX m. U + -T < X m & X m < U";
 by (dtac lemma_converg2 1 THEN assume_tac 1);
 by (rtac ccontr 1 THEN Asm_full_simp_tac 1);
-by (fold_tac [real_le_def]);
+by (asm_full_simp_tac (simpset() addsimps [linorder_not_less]) 1);
 by (dtac lemma_converg3 1);
 by (dtac isLub_le_isUb 1 THEN assume_tac 1);
 by (auto_tac (claset() addDs [order_less_le_trans],
--- a/src/HOL/Hyperreal/Series.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Hyperreal/Series.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -382,7 +382,7 @@
 by (subgoal_tac "0 <= sumr 0 (Suc (Suc 0) * Suc no + n) f + - suminf f" 1);
 by (dtac (abs_eqI1) 1 );
 by (Asm_full_simp_tac 1);
-by (auto_tac (claset(),simpset() addsimps [real_le_def]));
+by (auto_tac (claset(),simpset() addsimps [linorder_not_less RS sym]));
 qed "sumr_pos_lt_pair";
 
 (*-----------------------------------------------------------------
--- a/src/HOL/Hyperreal/Transcendental.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Hyperreal/Transcendental.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -34,10 +34,10 @@
 by (rtac some_equality 1);
 by (forw_inst_tac [("n","n")] zero_less_power 2);
 by (auto_tac (claset(),simpset() addsimps [zero_less_mult_iff]));
-by (res_inst_tac [("R1.0","u"),("R2.0","x")] real_linear_less2 1);
+by (res_inst_tac [("x","u"),("y","x")] linorder_cases 1);
 by (dres_inst_tac [("n1","n"),("x","u")] (zero_less_Suc RSN  (3, realpow_less)) 1);
 by (dres_inst_tac [("n1","n"),("x","x")] (zero_less_Suc RSN (3, realpow_less)) 4);
-by (auto_tac (claset(),simpset() addsimps [real_less_not_refl])); 
+by (auto_tac (claset(),simpset() addsimps [order_less_irrefl])); 
 qed "real_root_pos";
 
 Goal "0 <= x ==> root(Suc n) (x ^ (Suc n)) = x";
@@ -62,10 +62,10 @@
 by (rtac some_equality 1);
 by Auto_tac;
 by (rtac ccontr 1);
-by (res_inst_tac [("R1.0","u"),("R2.0","1")] real_linear_less2 1);
+by (res_inst_tac [("x","u"),("y","1")] linorder_cases 1);
 by (dres_inst_tac [("n","n")] realpow_Suc_less_one 1);
 by (dres_inst_tac [("n","n")] power_gt1_lemma 4);
-by (auto_tac (claset(),simpset() addsimps [real_less_not_refl]));
+by (auto_tac (claset(),simpset() addsimps [order_less_irrefl]));
 qed "real_root_one";
 Addsimps [real_root_one];
 
@@ -160,14 +160,14 @@
 by (rtac someI2 1 THEN Step_tac 1 THEN Blast_tac 2);
 by (Asm_full_simp_tac 1 THEN Asm_full_simp_tac 1);
 by (res_inst_tac [("a","xa * x")] someI2 1);
-by (auto_tac (claset() addEs [real_less_asym],
+by (auto_tac (claset() addEs [order_less_asym],
     simpset() addsimps mult_ac@[power_mult_distrib RS sym,realpow_two_disj,
     zero_less_power, real_mult_order] delsimps [realpow_Suc]));
 qed "real_sqrt_mult_distrib";
 
 Goal "[|0<=x; 0<=y |] ==> sqrt(x*y) =  sqrt(x) * sqrt(y)";
 by (auto_tac (claset() addIs [ real_sqrt_mult_distrib],
-    simpset() addsimps [real_le_less]));
+    simpset() addsimps [order_le_less]));
 qed "real_sqrt_mult_distrib2";
 
 Goal "(r * r = 0) = (r = (0::real))";
@@ -184,7 +184,7 @@
 
 Goal "0 <= x ==> 0 <= sqrt(x)";
 by (auto_tac (claset() addIs [real_sqrt_gt_zero],
-    simpset() addsimps [real_le_less]));
+    simpset() addsimps [order_le_less]));
 qed "real_sqrt_ge_zero";
 
 Goal "0 <= sqrt (x ^ 2 + y ^ 2)";
@@ -226,7 +226,7 @@
 
 Goal "0 < x ==> sqrt x ~= 0";
 by (ftac real_sqrt_pow2_gt_zero 1);
-by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2, real_less_not_refl]));
+by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2, order_less_irrefl]));
 qed "real_sqrt_not_eq_zero";
 
 Goal "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x";
@@ -621,7 +621,7 @@
 by (rtac real_le_trans 2 THEN assume_tac 3 THEN Auto_tac);
 by (res_inst_tac [("x","e")] exI 1 THEN Auto_tac);
 by (res_inst_tac [("y","K * abs x")] order_le_less_trans 1);
-by (res_inst_tac [("R2.0","K * e")] real_less_trans 2);
+by (res_inst_tac [("y","K * e")] order_less_trans 2);
 by (res_inst_tac [("z","inverse K")] (CLAIM_SIMP 
     "[|(0::real) <z; z*x<z*y |] ==> x<y" [mult_less_cancel_left]) 3);
 by (asm_full_simp_tac (simpset() addsimps [mult_assoc RS sym]) 4);
@@ -1032,7 +1032,7 @@
 
 Goal "0 < exp x";
 by (simp_tac (simpset() addsimps 
-    [CLAIM_SIMP "(x < y) = (x <= y & y ~= (x::real))" [real_le_less]]) 1);
+    [CLAIM_SIMP "(x < y) = (x <= y & y ~= (x::real))" [order_le_less]]) 1);
 qed "exp_gt_zero";
 Addsimps [exp_gt_zero];
 
@@ -1073,8 +1073,8 @@
 qed "exp_less_cancel_iff";
 AddIffs [exp_less_cancel_iff];
 
-Goalw [real_le_def] "(exp(x) <= exp(y)) = (x <= y)";
-by (Auto_tac);
+Goal "(exp(x) <= exp(y)) = (x <= y)";
+by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym]));
 qed "exp_le_cancel_iff";
 AddIffs [exp_le_cancel_iff];
 
@@ -1094,7 +1094,7 @@
 qed "lemma_exp_total";
 
 Goal "0 < y ==> EX x. exp x = y";
-by (res_inst_tac [("R1.0","1"),("R2.0","y")] real_linear_less2 1);
+by (res_inst_tac [("x","1"),("y","y")] linorder_cases 1);
 by (dtac (order_less_imp_le RS lemma_exp_total) 1);
 by (res_inst_tac [("x","0")] exI 2);
 by (ftac real_inverse_gt_one 3);
@@ -1156,7 +1156,8 @@
 qed "ln_less_cancel_iff";
 Addsimps [ln_less_cancel_iff];
 
-Goalw [real_le_def] "[| 0 < x; 0 < y|] ==> (ln x <= ln y) = (x <= y)";
+Goal "[| 0 < x; 0 < y|] ==> (ln x <= ln y) = (x <= y)";
+by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym]));
 by (Auto_tac);
 qed "ln_le_cancel_iff";
 Addsimps [ln_le_cancel_iff];
@@ -1348,8 +1349,8 @@
 by (arith_tac 1);
 qed "real_gt_one_ge_zero_add_less";
 
-Goalw [real_le_def] "abs(sin x) <= 1";
-by (rtac notI 1);
+Goal "abs(sin x) <= 1";
+by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym]));
 by (dres_inst_tac [("n","Suc 0")] power_gt1 1); 
 by (auto_tac (claset(),simpset() delsimps [realpow_Suc]));
 by (dres_inst_tac [("r1","cos x")] (realpow_two_le RSN 
@@ -1376,8 +1377,8 @@
 qed "sin_le_one";
 Addsimps [sin_le_one];
 
-Goalw [real_le_def] "abs(cos x) <= 1";
-by (rtac notI 1);
+Goal "abs(cos x) <= 1";
+by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym]));
 by (dres_inst_tac [("n","Suc 0")] power_gt1 1); 
 by (auto_tac (claset(),simpset() delsimps [realpow_Suc]));
 by (dres_inst_tac [("r1","sin x")] (realpow_two_le RSN 
@@ -1599,7 +1600,7 @@
 by (asm_simp_tac (simpset() addsimps [mult_less_cancel_left]) 1); 
 by (rtac real_mult_less_mono 1); (*mult_strict_mono would be stronger*)
 by (ALLGOALS(Asm_simp_tac));
-by (TRYALL(rtac real_less_trans));
+by (TRYALL(rtac order_less_trans));
 by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc] delsimps [fact_Suc]));
 qed "sin_gt_zero";
 
@@ -1626,9 +1627,9 @@
 by (dtac sums_minus 1);
 by (rtac (CLAIM "- x < -y ==> (y::real) < x") 1);
 by (ftac sums_unique 1 THEN Auto_tac);
-by (res_inst_tac [("R2.0",
+by (res_inst_tac [("y",
     "sumr 0 (Suc (Suc (Suc 0))) (%n. -((- 1) ^ n /(real (fact(2 * n))) \
-\               * 2 ^ (2 * n)))")] real_less_trans 1);
+\               * 2 ^ (2 * n)))")] order_less_trans 1);
 by (simp_tac (simpset() addsimps [fact_num_eq_if,realpow_num_eq_if] 
     delsimps [fact_Suc,realpow_Suc]) 1);
 by (simp_tac (simpset() addsimps [real_mult_assoc] 
@@ -1832,7 +1833,7 @@
 Goal "[| 0 < x; x < pi/2 |] ==> 0 < sin x";
 by (rtac sin_gt_zero 1);
 by (assume_tac 1); 
-by (rtac real_less_trans 1 THEN assume_tac 1);
+by (rtac order_less_trans 1 THEN assume_tac 1);
 by (rtac pi_half_less_two 1); 
 qed "sin_gt_zero2";
 
@@ -1866,7 +1867,7 @@
 qed "cos_gt_zero";
 
 Goal "[| -(pi/2) < x; x < pi/2 |] ==> 0 < cos x";
-by (res_inst_tac [("R1.0","x"),("R2.0","0")] real_linear_less2 1);
+by (res_inst_tac [("x","x"),("y","0")] linorder_cases 1);
 by (rtac (cos_minus RS subst) 1);
 by (rtac cos_gt_zero 1);
 by (rtac (CLAIM "-y < x ==> -x < (y::real)") 2);
@@ -1874,7 +1875,7 @@
 qed "cos_gt_zero_pi";
  
 Goal "[| -(pi/2) <= x; x <= pi/2 |] ==> 0 <= cos x";
-by (auto_tac (claset(),HOL_ss addsimps [real_le_less,
+by (auto_tac (claset(),HOL_ss addsimps [order_le_less,
     cos_gt_zero_pi]));
 by Auto_tac;
 qed "cos_ge_zero";
@@ -1888,7 +1889,7 @@
 qed "sin_gt_zero_pi";
 
 Goal "[| 0 <= x; x <= pi |] ==> 0 <= sin x";
-by (auto_tac (claset(),simpset() addsimps [real_le_less,
+by (auto_tac (claset(),simpset() addsimps [order_le_less,
     sin_gt_zero_pi]));
 qed "sin_ge_zero";
 
@@ -2354,7 +2355,7 @@
 by (cut_inst_tac [("y","x")] arctan_ubound 2);
 by (cut_inst_tac [("y","x")] arctan_lbound 4);
 by (auto_tac (claset(),
-     simpset() addsimps [real_of_nat_Suc, left_distrib,real_le_def, mult_less_0_iff] 
+     simpset() addsimps [real_of_nat_Suc, left_distrib,linorder_not_less RS sym, mult_less_0_iff] 
      delsimps [arctan]));
 qed "cos_arctan_not_zero";
 Addsimps [cos_arctan_not_zero];
@@ -2544,7 +2545,7 @@
 
 Goal "[| 0 <= x; 0 <= y |] ==> (root(Suc n) x <= root(Suc n) y) = (x <= y)";
 by (auto_tac (claset() addIs [real_root_le_mono],simpset()));
-by (simp_tac (simpset() addsimps [real_le_def]) 1);
+by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
 by Auto_tac;
 by (dres_inst_tac [("x","y"),("n","n")] real_root_less_mono 1);
 by Auto_tac;
@@ -2808,11 +2809,11 @@
 by (rotate_tac 2 2);
 by (dtac (real_mult_assoc RS subst) 2);
 by (rotate_tac 2 2);
-by (ftac (real_mult_inv_left RS subst) 2);
+by (ftac (left_inverse RS subst) 2);
 by (assume_tac 2);
 by (thin_tac "(1 - z) * (x + y) = x /(x + y) * (x + y)" 2);
 by (thin_tac "1 - z = x /(x + y)" 2);
-by (auto_tac (claset(),simpset() addsimps [real_mult_assoc]));
+by (auto_tac (claset(),simpset() addsimps [mult_assoc]));
 by (auto_tac (claset(),simpset() addsimps [right_distrib,
     left_diff_distrib]));
 qed "lemma_divide_rearrange";
--- a/src/HOL/Integ/Bin.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Integ/Bin.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -314,17 +314,18 @@
 
 (** Less-than-or-equals (\<le>) **)
 
-lemma le_number_of_eq_not_less: "(number_of x \<le> (number_of y::int)) =
-      (~ number_of y < (number_of x::int))"
-apply (rule linorder_not_less [symmetric])
-done
+text{*Reduces @{term "a\<le>b"} to @{term "~ (b<a)"} for ALL numerals*}
+lemmas le_number_of_eq_not_less =
+       linorder_not_less [of "number_of w" "number_of v", symmetric, standard]
+
+declare le_number_of_eq_not_less [simp]
+
 
 (** Absolute value (abs) **)
 
 lemma zabs_number_of:
  "abs(number_of x::int) =
   (if number_of x < (0::int) then -number_of x else number_of x)"
-
 apply (unfold zabs_def)
 apply (rule refl)
 done
@@ -423,4 +424,11 @@
 (* Numeral0 -> 0 and Numeral1 -> 1 *)
 declare int_numeral_0_eq_0 [simp] int_numeral_1_eq_1 [simp]
 
+
+(*Simplification of  x-y < 0, etc.*)
+declare less_iff_diff_less_0 [symmetric, simp]
+declare eq_iff_diff_eq_0 [symmetric, simp]
+declare le_iff_diff_le_0 [symmetric, simp]
+
+
 end
--- a/src/HOL/Integ/Equiv.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Integ/Equiv.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -118,6 +118,21 @@
   apply blast
   done
 
+lemma quotient_eqI:
+  "[|equiv A r; X \<in> A//r; Y \<in> A//r; x \<in> X; y \<in> Y; (x,y) \<in> r|] ==> X = Y" 
+  apply (clarify elim!: quotientE)
+  apply (rule equiv_class_eq, assumption)
+  apply (unfold equiv_def sym_def trans_def, blast)
+  done
+
+lemma quotient_eq_iff:
+  "[|equiv A r; X \<in> A//r; Y \<in> A//r; x \<in> X; y \<in> Y|] ==> (X = Y) = ((x,y) \<in> r)" 
+  apply (rule iffI)  
+   prefer 2 apply (blast del: equalityI intro: quotient_eqI) 
+  apply (clarify elim!: quotientE)
+  apply (unfold equiv_def sym_def trans_def, blast)
+  done
+
 
 subsection {* Defining unary operations upon equivalence classes *}
 
--- a/src/HOL/Integ/Int.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Integ/Int.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -320,13 +320,17 @@
 by (auto simp add: neg_eq_less_0 zless_iff_Suc_zadd 
                    diff_eq_eq [symmetric] zdiff_def)
 
-lemma int_cases: 
+lemma int_cases [cases type: int, case_names nonneg neg]: 
      "[|!! n. z = int n ==> P;  !! n. z =  - (int (Suc n)) ==> P |] ==> P"
 apply (case_tac "neg z")
 apply (fast dest!: negD)
 apply (drule not_neg_nat [symmetric], auto) 
 done
 
+lemma int_induct [induct type: int, case_names nonneg neg]: 
+     "[|!! n. P (int n);  !!n. P (- (int (Suc n))) |] ==> P z"
+  by (cases z) auto
+
 
 (*Legacy ML bindings, but no longer the structure Int.*)
 ML
--- a/src/HOL/Integ/IntArith.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Integ/IntArith.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -43,6 +43,9 @@
 lemma zadd_left_cancel0 [simp]: "(z = z + w) = (w = (0::int))"
 by arith
 
+lemma int_one_le_iff_zero_less: "((1::int) \<le> z) = (0 < z)"
+by arith
+
 
 subsection{*The Functions @{term nat} and @{term int}*}
 
--- a/src/HOL/Integ/NatBin.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Integ/NatBin.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -269,17 +269,6 @@
 declare less_nat_number_of [simp]
 
 
-(** Less-than-or-equals (<=) **)
-
-lemma le_nat_number_of_eq_not_less:
-     "(number_of x <= (number_of y::nat)) =  
-      (~ number_of y < (number_of x::nat))"
-apply (rule linorder_not_less [symmetric])
-done
-
-declare le_nat_number_of_eq_not_less [simp]
-
-
 (*Maps #n to n for n = 0, 1, 2*)
 lemmas numerals = numeral_0_eq_0 numeral_1_eq_1 numeral_2_eq_2
 
@@ -732,7 +721,6 @@
 val eq_nat_nat_iff = thm"eq_nat_nat_iff";
 val eq_nat_number_of = thm"eq_nat_number_of";
 val less_nat_number_of = thm"less_nat_number_of";
-val le_nat_number_of_eq_not_less = thm"le_nat_number_of_eq_not_less";
 val power2_eq_square = thm "power2_eq_square";
 val zero_le_power2 = thm "zero_le_power2";
 val zero_less_power2 = thm "zero_less_power2";
--- a/src/HOL/Integ/nat_simprocs.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Integ/nat_simprocs.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -67,7 +67,7 @@
 
 val bin_simps = [numeral_0_eq_0 RS sym, numeral_1_eq_1 RS sym,
                  add_nat_number_of, nat_number_of_add_left,
-                 diff_nat_number_of, le_nat_number_of_eq_not_less,
+                 diff_nat_number_of, le_number_of_eq_not_less,
                  less_nat_number_of, mult_nat_number_of,
                  thm "Let_number_of", nat_number_of] @
                 bin_arith_simps @ bin_rel_simps;
@@ -506,7 +506,7 @@
 val add_rules =
   [thm "Let_number_of", thm "Let_0", thm "Let_1", nat_0, nat_1,
    add_nat_number_of, diff_nat_number_of, mult_nat_number_of,
-   eq_nat_number_of, less_nat_number_of, le_nat_number_of_eq_not_less,
+   eq_nat_number_of, less_nat_number_of, le_number_of_eq_not_less,
    le_Suc_number_of,le_number_of_Suc,
    less_Suc_number_of,less_number_of_Suc,
    Suc_eq_number_of,eq_number_of_Suc,
--- a/src/HOL/IsaMakefile	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/IsaMakefile	Tue Jan 27 15:39:51 2004 +0100
@@ -139,16 +139,14 @@
 $(OUT)/HOL-Complex: $(OUT)/HOL Complex/ROOT.ML\
   Library/Zorn.thy\
   Real/Complex_Numbers.thy \
-  Real/Lubs.ML Real/Lubs.thy Real/PNat.ML Real/PNat.thy \
-  Real/PRat.ML Real/PRat.thy \
-  Real/PReal.thy Real/RComplete.ML Real/RComplete.thy \
+  Real/Lubs.ML Real/Lubs.thy Real/rat_arith.ML Real/RatArith.thy\
+  Real/Rational.thy Real/PReal.thy Real/RComplete.thy \
   Real/ROOT.ML Real/Real.thy \
-  Real/RealArith.thy Real/real_arith.ML Real/RealBin.ML \
-  Real/RealBin.thy Real/RealDef.thy \
-  Real/RealInt.thy Real/RealPow.thy Real/document/root.tex Real/real_arith.ML\
+  Real/RealArith.thy Real/real_arith.ML Real/RealDef.thy \
+  Real/RealPow.thy Real/document/root.tex Real/real_arith.ML\
   Hyperreal/EvenOdd.ML Hyperreal/EvenOdd.thy \
   Hyperreal/Fact.ML Hyperreal/Fact.thy\
-  Hyperreal/Filter.ML Hyperreal/Filter.thy Hyperreal/HRealAbs.ML\
+  Hyperreal/Filter.ML Hyperreal/Filter.thy\
   Hyperreal/HRealAbs.thy Hyperreal/HSeries.ML Hyperreal/HSeries.thy\
   Hyperreal/HyperArith.thy Hyperreal/HyperBin.ML Hyperreal/HyperBin.thy \
   Hyperreal/HyperDef.thy Hyperreal/HyperNat.ML Hyperreal/HyperNat.thy\
@@ -197,8 +195,7 @@
   Library/Permutation.thy Library/Primes.thy Library/Quotient.thy \
   Library/Nat_Infinity.thy \
   Library/README.html Library/Continuity.thy \
-  Library/Nested_Environment.thy Library/Rational_Numbers.thy \
-  Library/Zorn.thy\
+  Library/Nested_Environment.thy Library/Zorn.thy\
   Library/Library/ROOT.ML Library/Library/document/root.tex \
   Library/Library/document/root.bib Library/While_Combinator.thy
 	@cd Library; $(ISATOOL) usedir $(OUT)/HOL Library
--- a/src/HOL/Library/Library.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Library/Library.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -2,7 +2,6 @@
 theory Library =
   Quotient +
   Nat_Infinity +
-  Rational_Numbers +
   List_Prefix +
   Nested_Environment +
   Accessible_Part +
--- a/src/HOL/Library/Rational_Numbers.thy	Tue Jan 27 09:44:14 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,672 +0,0 @@
-(*  Title: HOL/Library/Rational_Numbers.thy
-    ID:    $Id$
-    Author: Markus Wenzel, TU Muenchen
-    License: GPL (GNU GENERAL PUBLIC LICENSE)
-*)
-
-header {*
-  \title{Rational numbers}
-  \author{Markus Wenzel}
-*}
-
-theory Rational_Numbers = Quotient + Ring_and_Field:
-
-subsection {* Fractions *}
-
-subsubsection {* The type of fractions *}
-
-typedef fraction = "{(a, b) :: int \<times> int | a b. b \<noteq> 0}"
-proof
-  show "(0, 1) \<in> ?fraction" by simp
-qed
-
-constdefs
-  fract :: "int => int => fraction"
-  "fract a b == Abs_fraction (a, b)"
-  num :: "fraction => int"
-  "num Q == fst (Rep_fraction Q)"
-  den :: "fraction => int"
-  "den Q == snd (Rep_fraction Q)"
-
-lemma fract_num [simp]: "b \<noteq> 0 ==> num (fract a b) = a"
-  by (simp add: fract_def num_def fraction_def Abs_fraction_inverse)
-
-lemma fract_den [simp]: "b \<noteq> 0 ==> den (fract a b) = b"
-  by (simp add: fract_def den_def fraction_def Abs_fraction_inverse)
-
-lemma fraction_cases [case_names fract, cases type: fraction]:
-  "(!!a b. Q = fract a b ==> b \<noteq> 0 ==> C) ==> C"
-proof -
-  assume r: "!!a b. Q = fract a b ==> b \<noteq> 0 ==> C"
-  obtain a b where "Q = fract a b" and "b \<noteq> 0"
-    by (cases Q) (auto simp add: fract_def fraction_def)
-  thus C by (rule r)
-qed
-
-lemma fraction_induct [case_names fract, induct type: fraction]:
-    "(!!a b. b \<noteq> 0 ==> P (fract a b)) ==> P Q"
-  by (cases Q) simp
-
-
-subsubsection {* Equivalence of fractions *}
-
-instance fraction :: eqv ..
-
-defs (overloaded)
-  equiv_fraction_def: "Q \<sim> R == num Q * den R = num R * den Q"
-
-lemma equiv_fraction_iff [iff]:
-    "b \<noteq> 0 ==> b' \<noteq> 0 ==> (fract a b \<sim> fract a' b') = (a * b' = a' * b)"
-  by (simp add: equiv_fraction_def)
-
-instance fraction :: equiv
-proof
-  fix Q R S :: fraction
-  {
-    show "Q \<sim> Q"
-    proof (induct Q)
-      fix a b :: int
-      assume "b \<noteq> 0" and "b \<noteq> 0"
-      with refl show "fract a b \<sim> fract a b" ..
-    qed
-  next
-    assume "Q \<sim> R" and "R \<sim> S"
-    show "Q \<sim> S"
-    proof (insert prems, induct Q, induct R, induct S)
-      fix a b a' b' a'' b'' :: int
-      assume b: "b \<noteq> 0" and b': "b' \<noteq> 0" and b'': "b'' \<noteq> 0"
-      assume "fract a b \<sim> fract a' b'" hence eq1: "a * b' = a' * b" ..
-      assume "fract a' b' \<sim> fract a'' b''" hence eq2: "a' * b'' = a'' * b'" ..
-      have "a * b'' = a'' * b"
-      proof cases
-        assume "a' = 0"
-        with b' eq1 eq2 have "a = 0 \<and> a'' = 0" by auto
-        thus ?thesis by simp
-      next
-        assume a': "a' \<noteq> 0"
-        from eq1 eq2 have "(a * b') * (a' * b'') = (a' * b) * (a'' * b')" by simp
-        hence "(a * b'') * (a' * b') = (a'' * b) * (a' * b')" by (simp only: mult_ac)
-        with a' b' show ?thesis by simp
-      qed
-      thus "fract a b \<sim> fract a'' b''" ..
-    qed
-  next
-    show "Q \<sim> R ==> R \<sim> Q"
-    proof (induct Q, induct R)
-      fix a b a' b' :: int
-      assume b: "b \<noteq> 0" and b': "b' \<noteq> 0"
-      assume "fract a b \<sim> fract a' b'"
-      hence "a * b' = a' * b" ..
-      hence "a' * b = a * b'" ..
-      thus "fract a' b' \<sim> fract a b" ..
-    qed
-  }
-qed
-
-lemma eq_fraction_iff [iff]:
-    "b \<noteq> 0 ==> b' \<noteq> 0 ==> (\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>) = (a * b' = a' * b)"
-  by (simp add: equiv_fraction_iff quot_equality)
-
-
-subsubsection {* Operations on fractions *}
-
-text {*
- We define the basic arithmetic operations on fractions and
- demonstrate their ``well-definedness'', i.e.\ congruence with respect
- to equivalence of fractions.
-*}
-
-instance fraction :: zero ..
-instance fraction :: one ..
-instance fraction :: plus ..
-instance fraction :: minus ..
-instance fraction :: times ..
-instance fraction :: inverse ..
-instance fraction :: ord ..
-
-defs (overloaded)
-  zero_fraction_def: "0 == fract 0 1"
-  one_fraction_def: "1 == fract 1 1"
-  add_fraction_def: "Q + R ==
-    fract (num Q * den R + num R * den Q) (den Q * den R)"
-  minus_fraction_def: "-Q == fract (-(num Q)) (den Q)"
-  mult_fraction_def: "Q * R == fract (num Q * num R) (den Q * den R)"
-  inverse_fraction_def: "inverse Q == fract (den Q) (num Q)"
-  le_fraction_def: "Q \<le> R ==
-    (num Q * den R) * (den Q * den R) \<le> (num R * den Q) * (den Q * den R)"
-
-lemma is_zero_fraction_iff: "b \<noteq> 0 ==> (\<lfloor>fract a b\<rfloor> = \<lfloor>0\<rfloor>) = (a = 0)"
-  by (simp add: zero_fraction_def eq_fraction_iff)
-
-theorem add_fraction_cong:
-  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>
-    ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0
-    ==> \<lfloor>fract a b + fract c d\<rfloor> = \<lfloor>fract a' b' + fract c' d'\<rfloor>"
-proof -
-  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
-  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" ..
-  assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" ..
-  have "\<lfloor>fract (a * d + c * b) (b * d)\<rfloor> = \<lfloor>fract (a' * d' + c' * b') (b' * d')\<rfloor>"
-  proof
-    show "(a * d + c * b) * (b' * d') = (a' * d' + c' * b') * (b * d)"
-      (is "?lhs = ?rhs")
-    proof -
-      have "?lhs = (a * b') * (d * d') + (c * d') * (b * b')"
-        by (simp add: int_distrib mult_ac)
-      also have "... = (a' * b) * (d * d') + (c' * d) * (b * b')"
-        by (simp only: eq1 eq2)
-      also have "... = ?rhs"
-        by (simp add: int_distrib mult_ac)
-      finally show "?lhs = ?rhs" .
-    qed
-    from neq show "b * d \<noteq> 0" by simp
-    from neq show "b' * d' \<noteq> 0" by simp
-  qed
-  with neq show ?thesis by (simp add: add_fraction_def)
-qed
-
-theorem minus_fraction_cong:
-  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> b \<noteq> 0 ==> b' \<noteq> 0
-    ==> \<lfloor>-(fract a b)\<rfloor> = \<lfloor>-(fract a' b')\<rfloor>"
-proof -
-  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"
-  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>"
-  hence "a * b' = a' * b" ..
-  hence "-a * b' = -a' * b" by simp
-  hence "\<lfloor>fract (-a) b\<rfloor> = \<lfloor>fract (-a') b'\<rfloor>" ..
-  with neq show ?thesis by (simp add: minus_fraction_def)
-qed
-
-theorem mult_fraction_cong:
-  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>
-    ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0
-    ==> \<lfloor>fract a b * fract c d\<rfloor> = \<lfloor>fract a' b' * fract c' d'\<rfloor>"
-proof -
-  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
-  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" ..
-  assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" ..
-  have "\<lfloor>fract (a * c) (b * d)\<rfloor> = \<lfloor>fract (a' * c') (b' * d')\<rfloor>"
-  proof
-    from eq1 eq2 have "(a * b') * (c * d') = (a' * b) * (c' * d)" by simp
-    thus "(a * c) * (b' * d') = (a' * c') * (b * d)" by (simp add: mult_ac)
-    from neq show "b * d \<noteq> 0" by simp
-    from neq show "b' * d' \<noteq> 0" by simp
-  qed
-  with neq show "\<lfloor>fract a b * fract c d\<rfloor> = \<lfloor>fract a' b' * fract c' d'\<rfloor>"
-    by (simp add: mult_fraction_def)
-qed
-
-theorem inverse_fraction_cong:
-  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor> ==> \<lfloor>fract a' b'\<rfloor> \<noteq> \<lfloor>0\<rfloor>
-    ==> b \<noteq> 0 ==> b' \<noteq> 0
-    ==> \<lfloor>inverse (fract a b)\<rfloor> = \<lfloor>inverse (fract a' b')\<rfloor>"
-proof -
-  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"
-  assume "\<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor>" and "\<lfloor>fract a' b'\<rfloor> \<noteq> \<lfloor>0\<rfloor>"
-  with neq obtain "a \<noteq> 0" and "a' \<noteq> 0" by (simp add: is_zero_fraction_iff)
-  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>"
-  hence "a * b' = a' * b" ..
-  hence "b * a' = b' * a" by (simp only: mult_ac)
-  hence "\<lfloor>fract b a\<rfloor> = \<lfloor>fract b' a'\<rfloor>" ..
-  with neq show ?thesis by (simp add: inverse_fraction_def)
-qed
-
-theorem le_fraction_cong:
-  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>
-    ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0
-    ==> (fract a b \<le> fract c d) = (fract a' b' \<le> fract c' d')"
-proof -
-  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
-  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" ..
-  assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" ..
-
-  let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))"
-  {
-    fix a b c d x :: int assume x: "x \<noteq> 0"
-    have "?le a b c d = ?le (a * x) (b * x) c d"
-    proof -
-      from x have "0 < x * x" by (auto simp add: zero_less_mult_iff)
-      hence "?le a b c d =
-          ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))"
-        by (simp add: mult_le_cancel_right)
-      also have "... = ?le (a * x) (b * x) c d"
-        by (simp add: mult_ac)
-      finally show ?thesis .
-    qed
-  } note le_factor = this
-
-  let ?D = "b * d" and ?D' = "b' * d'"
-  from neq have D: "?D \<noteq> 0" by simp
-  from neq have "?D' \<noteq> 0" by simp
-  hence "?le a b c d = ?le (a * ?D') (b * ?D') c d"
-    by (rule le_factor)
-  also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')"
-    by (simp add: mult_ac)
-  also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')"
-    by (simp only: eq1 eq2)
-  also have "... = ?le (a' * ?D) (b' * ?D) c' d'"
-    by (simp add: mult_ac)
-  also from D have "... = ?le a' b' c' d'"
-    by (rule le_factor [symmetric])
-  finally have "?le a b c d = ?le a' b' c' d'" .
-  with neq show ?thesis by (simp add: le_fraction_def)
-qed
-
-
-subsection {* Rational numbers *}
-
-subsubsection {* The type of rational numbers *}
-
-typedef (Rat)
-  rat = "UNIV :: fraction quot set" ..
-
-lemma RatI [intro, simp]: "Q \<in> Rat"
-  by (simp add: Rat_def)
-
-constdefs
-  fraction_of :: "rat => fraction"
-  "fraction_of q == pick (Rep_Rat q)"
-  rat_of :: "fraction => rat"
-  "rat_of Q == Abs_Rat \<lfloor>Q\<rfloor>"
-
-theorem rat_of_equality [iff?]: "(rat_of Q = rat_of Q') = (\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor>)"
-  by (simp add: rat_of_def Abs_Rat_inject)
-
-lemma rat_of: "\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor> ==> rat_of Q = rat_of Q'" ..
-
-constdefs
-  Fract :: "int => int => rat"
-  "Fract a b == rat_of (fract a b)"
-
-theorem Fract_inverse: "\<lfloor>fraction_of (Fract a b)\<rfloor> = \<lfloor>fract a b\<rfloor>"
-  by (simp add: fraction_of_def rat_of_def Fract_def Abs_Rat_inverse pick_inverse)
-
-theorem Fract_equality [iff?]:
-    "(Fract a b = Fract c d) = (\<lfloor>fract a b\<rfloor> = \<lfloor>fract c d\<rfloor>)"
-  by (simp add: Fract_def rat_of_equality)
-
-theorem eq_rat:
-    "b \<noteq> 0 ==> d \<noteq> 0 ==> (Fract a b = Fract c d) = (a * d = c * b)"
-  by (simp add: Fract_equality eq_fraction_iff)
-
-theorem Rat_cases [case_names Fract, cases type: rat]:
-  "(!!a b. q = Fract a b ==> b \<noteq> 0 ==> C) ==> C"
-proof -
-  assume r: "!!a b. q = Fract a b ==> b \<noteq> 0 ==> C"
-  obtain x where "q = Abs_Rat x" by (cases q)
-  moreover obtain Q where "x = \<lfloor>Q\<rfloor>" by (cases x)
-  moreover obtain a b where "Q = fract a b" and "b \<noteq> 0" by (cases Q)
-  ultimately have "q = Fract a b" by (simp only: Fract_def rat_of_def)
-  thus ?thesis by (rule r)
-qed
-
-theorem Rat_induct [case_names Fract, induct type: rat]:
-    "(!!a b. b \<noteq> 0 ==> P (Fract a b)) ==> P q"
-  by (cases q) simp
-
-
-subsubsection {* Canonical function definitions *}
-
-text {*
-  Note that the unconditional version below is much easier to read.
-*}
-
-theorem rat_cond_function:
-  "(!!q r. P \<lfloor>fraction_of q\<rfloor> \<lfloor>fraction_of r\<rfloor> ==>
-      f q r == g (fraction_of q) (fraction_of r)) ==>
-    (!!a b a' b' c d c' d'.
-      \<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor> ==>
-      P \<lfloor>fract a b\<rfloor> \<lfloor>fract c d\<rfloor> ==> P \<lfloor>fract a' b'\<rfloor> \<lfloor>fract c' d'\<rfloor> ==>
-      b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0 ==>
-      g (fract a b) (fract c d) = g (fract a' b') (fract c' d')) ==>
-    P \<lfloor>fract a b\<rfloor> \<lfloor>fract c d\<rfloor> ==>
-      f (Fract a b) (Fract c d) = g (fract a b) (fract c d)"
-  (is "PROP ?eq ==> PROP ?cong ==> ?P ==> _")
-proof -
-  assume eq: "PROP ?eq" and cong: "PROP ?cong" and P: ?P
-  have "f (Abs_Rat \<lfloor>fract a b\<rfloor>) (Abs_Rat \<lfloor>fract c d\<rfloor>) = g (fract a b) (fract c d)"
-  proof (rule quot_cond_function)
-    fix X Y assume "P X Y"
-    with eq show "f (Abs_Rat X) (Abs_Rat Y) == g (pick X) (pick Y)"
-      by (simp add: fraction_of_def pick_inverse Abs_Rat_inverse)
-  next
-    fix Q Q' R R' :: fraction
-    show "\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor> ==> \<lfloor>R\<rfloor> = \<lfloor>R'\<rfloor> ==>
-        P \<lfloor>Q\<rfloor> \<lfloor>R\<rfloor> ==> P \<lfloor>Q'\<rfloor> \<lfloor>R'\<rfloor> ==> g Q R = g Q' R'"
-      by (induct Q, induct Q', induct R, induct R') (rule cong)
-  qed
-  thus ?thesis by (unfold Fract_def rat_of_def)
-qed
-
-theorem rat_function:
-  "(!!q r. f q r == g (fraction_of q) (fraction_of r)) ==>
-    (!!a b a' b' c d c' d'.
-      \<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor> ==>
-      b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0 ==>
-      g (fract a b) (fract c d) = g (fract a' b') (fract c' d')) ==>
-    f (Fract a b) (Fract c d) = g (fract a b) (fract c d)"
-proof -
-  case rule_context from this TrueI
-  show ?thesis by (rule rat_cond_function)
-qed
-
-
-subsubsection {* Standard operations on rational numbers *}
-
-instance rat :: zero ..
-instance rat :: one ..
-instance rat :: plus ..
-instance rat :: minus ..
-instance rat :: times ..
-instance rat :: inverse ..
-instance rat :: ord ..
-instance rat :: number ..
-
-defs (overloaded)
-  zero_rat_def: "0 == rat_of 0"
-  one_rat_def: "1 == rat_of 1"
-  add_rat_def: "q + r == rat_of (fraction_of q + fraction_of r)"
-  minus_rat_def: "-q == rat_of (-(fraction_of q))"
-  diff_rat_def: "q - r == q + (-(r::rat))"
-  mult_rat_def: "q * r == rat_of (fraction_of q * fraction_of r)"
-  inverse_rat_def: "q \<noteq> 0 ==> inverse q == rat_of (inverse (fraction_of q))"
-  divide_rat_def: "r \<noteq> 0 ==> q / r == q * inverse (r::rat)"
-  le_rat_def: "q \<le> r == fraction_of q \<le> fraction_of r"
-  less_rat_def: "q < r == q \<le> r \<and> q \<noteq> (r::rat)"
-  abs_rat_def: "\<bar>q\<bar> == if q < 0 then -q else (q::rat)"
-  number_of_rat_def: "number_of b == Fract (number_of b) 1"
-
-theorem zero_rat: "0 = Fract 0 1"
-  by (simp add: zero_rat_def zero_fraction_def rat_of_def Fract_def)        
-
-theorem one_rat: "1 = Fract 1 1"
-  by (simp add: one_rat_def one_fraction_def rat_of_def Fract_def)
-
-theorem add_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
-  Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
-proof -
-  have "Fract a b + Fract c d = rat_of (fract a b + fract c d)"
-    by (rule rat_function, rule add_rat_def, rule rat_of, rule add_fraction_cong)
-  also
-  assume "b \<noteq> 0"  "d \<noteq> 0"
-  hence "fract a b + fract c d = fract (a * d + c * b) (b * d)"
-    by (simp add: add_fraction_def)
-  finally show ?thesis by (unfold Fract_def)
-qed
-
-theorem minus_rat: "b \<noteq> 0 ==> -(Fract a b) = Fract (-a) b"
-proof -
-  have "-(Fract a b) = rat_of (-(fract a b))"
-    by (rule rat_function, rule minus_rat_def, rule rat_of, rule minus_fraction_cong)
-  also assume "b \<noteq> 0" hence "-(fract a b) = fract (-a) b"
-    by (simp add: minus_fraction_def)
-  finally show ?thesis by (unfold Fract_def)
-qed
-
-theorem diff_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
-    Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
-  by (simp add: diff_rat_def add_rat minus_rat)
-
-theorem mult_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
-  Fract a b * Fract c d = Fract (a * c) (b * d)"
-proof -
-  have "Fract a b * Fract c d = rat_of (fract a b * fract c d)"
-    by (rule rat_function, rule mult_rat_def, rule rat_of, rule mult_fraction_cong)
-  also
-  assume "b \<noteq> 0"  "d \<noteq> 0"
-  hence "fract a b * fract c d = fract (a * c) (b * d)"
-    by (simp add: mult_fraction_def)
-  finally show ?thesis by (unfold Fract_def)
-qed
-
-theorem inverse_rat: "Fract a b \<noteq> 0 ==> b \<noteq> 0 ==>
-  inverse (Fract a b) = Fract b a"
-proof -
-  assume neq: "b \<noteq> 0" and nonzero: "Fract a b \<noteq> 0"
-  hence "\<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor>"
-    by (simp add: zero_rat eq_rat is_zero_fraction_iff)
-  with _ inverse_fraction_cong [THEN rat_of]
-  have "inverse (Fract a b) = rat_of (inverse (fract a b))"
-  proof (rule rat_cond_function)
-    fix q assume cond: "\<lfloor>fraction_of q\<rfloor> \<noteq> \<lfloor>0\<rfloor>"
-    have "q \<noteq> 0"
-    proof (cases q)
-      fix a b assume "b \<noteq> 0" and "q = Fract a b"
-      from this cond show ?thesis
-        by (simp add: Fract_inverse is_zero_fraction_iff zero_rat eq_rat)
-    qed
-    thus "inverse q == rat_of (inverse (fraction_of q))"
-      by (rule inverse_rat_def)
-  qed
-  also from neq nonzero have "inverse (fract a b) = fract b a"
-    by (simp add: inverse_fraction_def)
-  finally show ?thesis by (unfold Fract_def)
-qed
-
-theorem divide_rat: "Fract c d \<noteq> 0 ==> b \<noteq> 0 ==> d \<noteq> 0 ==>
-  Fract a b / Fract c d = Fract (a * d) (b * c)"
-proof -
-  assume neq: "b \<noteq> 0"  "d \<noteq> 0" and nonzero: "Fract c d \<noteq> 0"
-  hence "c \<noteq> 0" by (simp add: zero_rat eq_rat)
-  with neq nonzero show ?thesis
-    by (simp add: divide_rat_def inverse_rat mult_rat)
-qed
-
-theorem le_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
-  (Fract a b \<le> Fract c d) = ((a * d) * (b * d) \<le> (c * b) * (b * d))"
-proof -
-  have "(Fract a b \<le> Fract c d) = (fract a b \<le> fract c d)"
-    by (rule rat_function, rule le_rat_def, rule le_fraction_cong)
-  also
-  assume "b \<noteq> 0"  "d \<noteq> 0"
-  hence "(fract a b \<le> fract c d) = ((a * d) * (b * d) \<le> (c * b) * (b * d))"
-    by (simp add: le_fraction_def)
-  finally show ?thesis .
-qed
-
-theorem less_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
-    (Fract a b < Fract c d) = ((a * d) * (b * d) < (c * b) * (b * d))"
-  by (simp add: less_rat_def le_rat eq_rat int_less_le)
-
-theorem abs_rat: "b \<noteq> 0 ==> \<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
-  by (simp add: abs_rat_def minus_rat zero_rat less_rat eq_rat)
-     (auto simp add: mult_less_0_iff zero_less_mult_iff int_le_less 
-                split: abs_split)
-
-
-subsubsection {* The ordered field of rational numbers *}
-
-lemma rat_add_assoc: "(q + r) + s = q + (r + (s::rat))"
-  by (induct q, induct r, induct s) 
-     (simp add: add_rat add_ac mult_ac int_distrib)
-
-lemma rat_add_0: "0 + q = (q::rat)"
-  by (induct q) (simp add: zero_rat add_rat)
-
-lemma rat_left_minus: "(-q) + q = (0::rat)"
-  by (induct q) (simp add: zero_rat minus_rat add_rat eq_rat)
-
-
-instance rat :: field
-proof
-  fix q r s :: rat
-  show "(q + r) + s = q + (r + s)"
-    by (rule rat_add_assoc)
-  show "q + r = r + q"
-    by (induct q, induct r) (simp add: add_rat add_ac mult_ac)
-  show "0 + q = q"
-    by (induct q) (simp add: zero_rat add_rat)
-  show "(-q) + q = 0"
-    by (rule rat_left_minus)
-  show "q - r = q + (-r)"
-    by (induct q, induct r) (simp add: add_rat minus_rat diff_rat)
-  show "(q * r) * s = q * (r * s)"
-    by (induct q, induct r, induct s) (simp add: mult_rat mult_ac)
-  show "q * r = r * q"
-    by (induct q, induct r) (simp add: mult_rat mult_ac)
-  show "1 * q = q"
-    by (induct q) (simp add: one_rat mult_rat)
-  show "(q + r) * s = q * s + r * s"
-    by (induct q, induct r, induct s) 
-       (simp add: add_rat mult_rat eq_rat int_distrib)
-  show "q \<noteq> 0 ==> inverse q * q = 1"
-    by (induct q) (simp add: inverse_rat mult_rat one_rat zero_rat eq_rat)
-  show "r \<noteq> 0 ==> q / r = q * inverse r"
-    by (induct q, induct r)
-       (simp add: mult_rat divide_rat inverse_rat zero_rat eq_rat)
-  show "0 \<noteq> (1::rat)"
-    by (simp add: zero_rat one_rat eq_rat) 
-  assume eq: "s+q = s+r" 
-    hence "(-s + s) + q = (-s + s) + r" by (simp only: eq rat_add_assoc)
-    thus "q = r" by (simp add: rat_left_minus rat_add_0)
-qed
-
-instance rat :: linorder
-proof
-  fix q r s :: rat
-  {
-    assume "q \<le> r" and "r \<le> s"
-    show "q \<le> s"
-    proof (insert prems, induct q, induct r, induct s)
-      fix a b c d e f :: int
-      assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
-      assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract e f"
-      show "Fract a b \<le> Fract e f"
-      proof -
-        from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f"
-          by (auto simp add: zero_less_mult_iff linorder_neq_iff)
-        have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)"
-        proof -
-          from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
-            by (simp add: le_rat)
-          with ff show ?thesis by (simp add: mult_le_cancel_right)
-        qed
-        also have "... = (c * f) * (d * f) * (b * b)"
-          by (simp only: mult_ac)
-        also have "... \<le> (e * d) * (d * f) * (b * b)"
-        proof -
-          from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)"
-            by (simp add: le_rat)
-          with bb show ?thesis by (simp add: mult_le_cancel_right)
-        qed
-        finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)"
-          by (simp only: mult_ac)
-        with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)"
-          by (simp add: mult_le_cancel_right)
-        with neq show ?thesis by (simp add: le_rat)
-      qed
-    qed
-  next
-    assume "q \<le> r" and "r \<le> q"
-    show "q = r"
-    proof (insert prems, induct q, induct r)
-      fix a b c d :: int
-      assume neq: "b \<noteq> 0"  "d \<noteq> 0"
-      assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract a b"
-      show "Fract a b = Fract c d"
-      proof -
-        from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
-          by (simp add: le_rat)
-        also have "... \<le> (a * d) * (b * d)"
-        proof -
-          from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)"
-            by (simp add: le_rat)
-          thus ?thesis by (simp only: mult_ac)
-        qed
-        finally have "(a * d) * (b * d) = (c * b) * (b * d)" .
-        moreover from neq have "b * d \<noteq> 0" by simp
-        ultimately have "a * d = c * b" by simp
-        with neq show ?thesis by (simp add: eq_rat)
-      qed
-    qed
-  next
-    show "q \<le> q"
-      by (induct q) (simp add: le_rat)
-    show "(q < r) = (q \<le> r \<and> q \<noteq> r)"
-      by (simp only: less_rat_def)
-    show "q \<le> r \<or> r \<le> q"
-      by (induct q, induct r) (simp add: le_rat mult_ac, arith)
-  }
-qed
-
-instance rat :: ordered_field
-proof
-  fix q r s :: rat
-  show "0 < (1::rat)" 
-    by (simp add: zero_rat one_rat less_rat) 
-  show "q \<le> r ==> s + q \<le> s + r"
-  proof (induct q, induct r, induct s)
-    fix a b c d e f :: int
-    assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
-    assume le: "Fract a b \<le> Fract c d"
-    show "Fract e f + Fract a b \<le> Fract e f + Fract c d"
-    proof -
-      let ?F = "f * f" from neq have F: "0 < ?F"
-        by (auto simp add: zero_less_mult_iff)
-      from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
-        by (simp add: le_rat)
-      with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F"
-        by (simp add: mult_le_cancel_right)
-      with neq show ?thesis by (simp add: add_rat le_rat mult_ac int_distrib)
-    qed
-  qed
-  show "q < r ==> 0 < s ==> s * q < s * r"
-  proof (induct q, induct r, induct s)
-    fix a b c d e f :: int
-    assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
-    assume le: "Fract a b < Fract c d"
-    assume gt: "0 < Fract e f"
-    show "Fract e f * Fract a b < Fract e f * Fract c d"
-    proof -
-      let ?E = "e * f" and ?F = "f * f"
-      from neq gt have "0 < ?E"
-        by (auto simp add: zero_rat less_rat le_rat int_less_le eq_rat)
-      moreover from neq have "0 < ?F"
-        by (auto simp add: zero_less_mult_iff)
-      moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)"
-        by (simp add: less_rat)
-      ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F"
-        by (simp add: mult_less_cancel_right)
-      with neq show ?thesis
-        by (simp add: less_rat mult_rat mult_ac)
-    qed
-  qed
-  show "\<bar>q\<bar> = (if q < 0 then -q else q)"
-    by (simp only: abs_rat_def)
-qed
-
-
-subsection {* Embedding integers *}
-
-constdefs
-  rat :: "int => rat"    (* FIXME generalize int to any numeric subtype (?) *)
-  "rat z == Fract z 1"
-  int_set :: "rat set"    ("\<int>")    (* FIXME generalize rat to any numeric supertype (?) *)
-  "\<int> == range rat"
-
-lemma rat_inject: "(rat z = rat w) = (z = w)"
-proof
-  assume "rat z = rat w"
-  hence "Fract z 1 = Fract w 1" by (unfold rat_def)
-  hence "\<lfloor>fract z 1\<rfloor> = \<lfloor>fract w 1\<rfloor>" ..
-  thus "z = w" by auto
-next
-  assume "z = w"
-  thus "rat z = rat w" by simp
-qed
-
-lemma int_set_cases [case_names rat, cases set: int_set]:
-  "q \<in> \<int> ==> (!!z. q = rat z ==> C) ==> C"
-proof (unfold int_set_def)
-  assume "!!z. q = rat z ==> C"
-  assume "q \<in> range rat" thus C ..
-qed
-
-lemma int_set_induct [case_names rat, induct set: int_set]:
-  "q \<in> \<int> ==> (!!z. P (rat z)) ==> P q"
-  by (rule int_set_cases) auto
-
-theorem number_of_rat: "number_of b = rat (number_of b)"
-  by (simp add: number_of_rat_def rat_def)
-
-end
--- a/src/HOL/MicroJava/BV/Kildall.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/MicroJava/BV/Kildall.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -367,17 +367,17 @@
   prefer 3
   apply assumption
   apply (erule listE_nth_in)
-  apply blast
- apply blast
+  apply simp
+ apply simp
 apply (subst decomp_propa)
- apply blast
+ apply fast
 apply simp
 apply (rule conjI)
  apply (rule merges_preserves_type)
  apply blast
  apply clarify
  apply (rule conjI)
-  apply(clarsimp, blast dest!: boundedD)
+  apply(clarsimp, fast dest!: boundedD)
  apply (erule pres_typeD)
   prefer 3
   apply assumption
--- a/src/HOL/Real/PNat.ML	Tue Jan 27 09:44:14 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,540 +0,0 @@
-(*  Title       : HOL/Real/PNat.ML
-    ID          : $Id$
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-
-The positive naturals -- proofs mainly as in theory Nat.
-*)
-
-Goal "mono(%X. {Suc 0} Un Suc`X)";
-by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1));
-qed "pnat_fun_mono";
-
-bind_thm ("pnat_unfold", pnat_fun_mono RS (pnat_def RS def_lfp_unfold));
-
-Goal "Suc 0 : pnat";
-by (stac pnat_unfold 1);
-by (rtac (singletonI RS UnI1) 1);
-qed "one_RepI";
-
-Addsimps [one_RepI];
-
-Goal "i: pnat ==> Suc(i) : pnat";
-by (stac pnat_unfold 1);
-by (etac (imageI RS UnI2) 1);
-qed "pnat_Suc_RepI";
-
-Goal "Suc (Suc 0) : pnat";
-by (rtac (one_RepI RS pnat_Suc_RepI) 1);
-qed "two_RepI";
-
-(*** Induction ***)
-
-val major::prems = Goal
-    "[| i: pnat;  P(Suc 0);   \
-\       !!j. [| j: pnat; P(j) |] ==> P(Suc(j)) |]  ==> P(i)";
-by (rtac ([pnat_def, pnat_fun_mono, major] MRS def_lfp_induct) 1);
-by (blast_tac (claset() addIs prems) 1);
-qed "PNat_induct";
-
-val prems = Goalw [pnat_one_def,pnat_Suc_def]
-    "[| P(1);   \
-\       !!n. P(n) ==> P(pSuc n) |]  ==> P(n)";
-by (rtac (Rep_pnat_inverse RS subst) 1);   
-by (rtac (Rep_pnat RS PNat_induct) 1);
-by (REPEAT (ares_tac prems 1
-     ORELSE eresolve_tac [Abs_pnat_inverse RS subst] 1));
-qed "pnat_induct";
-
-(*Perform induction on n. *)
-fun pnat_ind_tac a i = 
-  induct_thm_tac pnat_induct a i  THEN  rename_last_tac a [""] (i+1);
-
-val prems = Goal
-    "[| !!x. P x 1;  \
-\       !!y. P 1 (pSuc y);  \
-\       !!x y. [| P x y |] ==> P (pSuc x) (pSuc y)  \
-\    |] ==> P m n";
-by (res_inst_tac [("x","m")] spec 1);
-by (pnat_ind_tac "n" 1);
-by (rtac allI 2);
-by (pnat_ind_tac "x" 2);
-by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1));
-qed "pnat_diff_induct";
-
-(*Case analysis on the natural numbers*)
-val prems = Goal
-    "[| n=1 ==> P;  !!x. n = pSuc(x) ==> P |] ==> P";
-by (subgoal_tac "n=1 | (EX x. n = pSuc(x))" 1);
-by (fast_tac (claset() addSEs prems) 1);
-by (pnat_ind_tac "n" 1);
-by (rtac (refl RS disjI1) 1);
-by (Blast_tac 1);
-qed "pnatE";
-
-(*** Isomorphisms: Abs_Nat and Rep_Nat ***)
-
-Goal "inj_on Abs_pnat pnat";
-by (rtac inj_on_inverseI 1);
-by (etac Abs_pnat_inverse 1);
-qed "inj_on_Abs_pnat";
-
-Addsimps [inj_on_Abs_pnat RS inj_on_iff];
-
-Goal "inj(Rep_pnat)";
-by (rtac inj_inverseI 1);
-by (rtac Rep_pnat_inverse 1);
-qed "inj_Rep_pnat";
-
-Goal "0 ~: pnat";
-by (stac pnat_unfold 1);
-by Auto_tac;
-qed "zero_not_mem_pnat";
-
-(* 0 : pnat ==> P *)
-bind_thm ("zero_not_mem_pnatE", zero_not_mem_pnat RS notE);
-
-Addsimps [zero_not_mem_pnat];
-
-Goal "x : pnat ==> 0 < x";
-by (dtac (pnat_unfold RS subst) 1);
-by Auto_tac;
-qed "mem_pnat_gt_zero";
-
-Goal "0 < x ==> x: pnat";
-by (stac pnat_unfold 1);
-by (dtac (gr_implies_not0 RS not0_implies_Suc) 1); 
-by (etac exE 1 THEN Asm_simp_tac 1);
-by (induct_tac "m" 1);
-by (auto_tac (claset(),simpset() 
-    addsimps [one_RepI]) THEN dtac pnat_Suc_RepI 1);
-by (Blast_tac 1);
-qed "gt_0_mem_pnat";
-
-Goal "(x: pnat) = (0 < x)";
-by (blast_tac (claset() addDs [mem_pnat_gt_zero,gt_0_mem_pnat]) 1);
-qed "mem_pnat_gt_0_iff";
-
-Goal "0 < Rep_pnat x";
-by (rtac (Rep_pnat RS mem_pnat_gt_zero) 1);
-qed "Rep_pnat_gt_zero";
-
-Goalw [pnat_add_def] "(x::pnat) + y = y + x";
-by (simp_tac (simpset() addsimps [add_commute]) 1);
-qed "pnat_add_commute";
-
-(** alternative definition for pnat **)
-(** order isomorphism **)
-Goal "pnat = {x::nat. 0 < x}";
-by (auto_tac (claset(), simpset() addsimps [mem_pnat_gt_0_iff]));  
-qed "Collect_pnat_gt_0";
-
-(*** Distinctness of constructors ***)
-
-Goalw [pnat_one_def,pnat_Suc_def] "pSuc(m) ~= 1";
-by (rtac (inj_on_Abs_pnat RS inj_on_contraD) 1);
-by (rtac (Rep_pnat_gt_zero RS Suc_mono RS less_not_refl2) 1);
-by (REPEAT (resolve_tac [Rep_pnat RS  pnat_Suc_RepI, one_RepI] 1));
-qed "pSuc_not_one";
-
-bind_thm ("one_not_pSuc", pSuc_not_one RS not_sym);
-
-AddIffs [pSuc_not_one,one_not_pSuc];
-
-bind_thm ("pSuc_neq_one", (pSuc_not_one RS notE));
-bind_thm ("one_neq_pSuc", pSuc_neq_one RS pSuc_neq_one);
-
-(** Injectiveness of pSuc **)
-
-Goalw [pnat_Suc_def] "inj(pSuc)";
-by (rtac injI 1);
-by (dtac (inj_on_Abs_pnat RS inj_onD) 1);
-by (REPEAT (resolve_tac [Rep_pnat, pnat_Suc_RepI] 1));
-by (dtac (inj_Suc RS injD) 1);
-by (etac (inj_Rep_pnat RS injD) 1);
-qed "inj_pSuc"; 
-
-bind_thm ("pSuc_inject", inj_pSuc RS injD);
-
-Goal "(pSuc(m)=pSuc(n)) = (m=n)";
-by (EVERY1 [rtac iffI, etac pSuc_inject, etac arg_cong]); 
-qed "pSuc_pSuc_eq";
-
-AddIffs [pSuc_pSuc_eq];
-
-Goal "n ~= pSuc(n)";
-by (pnat_ind_tac "n" 1);
-by (ALLGOALS Asm_simp_tac);
-qed "n_not_pSuc_n";
-
-bind_thm ("pSuc_n_not_n", n_not_pSuc_n RS not_sym);
-
-Goal "n ~= 1 ==> EX m. n = pSuc m";
-by (rtac pnatE 1);
-by (REPEAT (Blast_tac 1));
-qed "not1_implies_pSuc";
-
-Goal "pSuc m = m + 1";
-by (auto_tac (claset(),simpset() addsimps [pnat_Suc_def,
-    pnat_one_def,Abs_pnat_inverse,pnat_add_def]));
-qed "pSuc_is_plus_one";
-
-Goal
-      "(Rep_pnat x + Rep_pnat y): pnat";
-by (cut_facts_tac [[Rep_pnat_gt_zero,
-    Rep_pnat_gt_zero] MRS add_less_mono,Collect_pnat_gt_0] 1);
-by (etac ssubst 1);
-by Auto_tac;
-qed "sum_Rep_pnat";
-
-Goalw [pnat_add_def] 
-      "Rep_pnat x + Rep_pnat y = Rep_pnat (x + y)";
-by (simp_tac (simpset() addsimps [sum_Rep_pnat RS 
-                          Abs_pnat_inverse]) 1);
-qed "sum_Rep_pnat_sum";
-
-Goalw [pnat_add_def] 
-      "(x + y) + z = x + (y + (z::pnat))";
-by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
-by (simp_tac (simpset() addsimps [sum_Rep_pnat RS 
-                Abs_pnat_inverse,add_assoc]) 1);
-qed "pnat_add_assoc";
-
-Goalw [pnat_add_def] "x + (y + z) = y + (x + (z::pnat))";
-by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
-by (simp_tac (simpset() addsimps [sum_Rep_pnat RS 
-          Abs_pnat_inverse,add_left_commute]) 1);
-qed "pnat_add_left_commute";
-
-(*Addition is an AC-operator*)
-bind_thms ("pnat_add_ac", [pnat_add_assoc, pnat_add_commute, pnat_add_left_commute]);
-
-Goalw [pnat_add_def] "((x::pnat) + y = x + z) = (y = z)";
-by (auto_tac (claset() addDs [inj_on_Abs_pnat RS inj_onD,
-     inj_Rep_pnat RS injD],simpset() addsimps [sum_Rep_pnat]));
-qed "pnat_add_left_cancel";
-
-Goalw [pnat_add_def] "(y + (x::pnat) = z + x) = (y = z)";
-by (auto_tac (claset() addDs [inj_on_Abs_pnat RS inj_onD,
-     inj_Rep_pnat RS injD],simpset() addsimps [sum_Rep_pnat]));
-qed "pnat_add_right_cancel";
-
-Goalw [pnat_add_def] "!(y::pnat). x + y ~= x";
-by (rtac (Rep_pnat_inverse RS subst) 1);
-by (auto_tac (claset() addDs [inj_on_Abs_pnat RS inj_onD] 
-  	               addSDs [add_eq_self_zero],
-	      simpset() addsimps [sum_Rep_pnat, Rep_pnat,Abs_pnat_inverse,
-				  Rep_pnat_gt_zero RS less_not_refl2]));
-qed "pnat_no_add_ident";
-
-
-(***) (***) (***) (***) (***) (***) (***) (***) (***)
-
-  (*** pnat_less ***)
-
-Goalw [pnat_less_def] "~ y < (y::pnat)";
-by Auto_tac;
-qed "pnat_less_not_refl";
-
-bind_thm ("pnat_less_irrefl",pnat_less_not_refl RS notE);
-
-Goalw [pnat_less_def] 
-     "x < (y::pnat) ==> x ~= y";
-by Auto_tac;
-qed "pnat_less_not_refl2";
-
-Goal "~ Rep_pnat y < 0";
-by Auto_tac;
-qed "Rep_pnat_not_less0";
-
-(*** Rep_pnat < 0 ==> P ***)
-bind_thm ("Rep_pnat_less_zeroE",Rep_pnat_not_less0 RS notE);
-
-Goal "~ Rep_pnat y < Suc 0";
-by (auto_tac (claset(),simpset() addsimps [less_Suc_eq,
-                  Rep_pnat_gt_zero,less_not_refl2]));
-qed "Rep_pnat_not_less_one";
-
-(*** Rep_pnat < 1 ==> P ***)
-bind_thm ("Rep_pnat_less_oneE",Rep_pnat_not_less_one RS notE);
-
-Goalw [pnat_less_def] 
-     "x < (y::pnat) ==> Rep_pnat y ~= Suc 0";
-by (auto_tac (claset(),simpset() 
-    addsimps [Rep_pnat_not_less_one] delsimps [less_Suc0]));
-qed "Rep_pnat_gt_implies_not0";
-
-Goalw [pnat_less_def] 
-      "(x::pnat) < y | x = y | y < x";
-by (cut_facts_tac [less_linear] 1);
-by (fast_tac (claset() addIs [inj_Rep_pnat RS injD]) 1);
-qed "pnat_less_linear";
-
-Goalw [le_def] "Suc 0 <= Rep_pnat x";
-by (rtac Rep_pnat_not_less_one 1);
-qed "Rep_pnat_le_one";
-
-Goalw [pnat_less_def]
-     "!! (z1::nat). z1 < z2  ==> EX z3. z1 + Rep_pnat z3 = z2";
-by (dtac less_imp_add_positive 1);
-by (force_tac (claset() addSIs [Abs_pnat_inverse],
-	       simpset() addsimps [Collect_pnat_gt_0]) 1);
-qed "lemma_less_ex_sum_Rep_pnat";
-
-
-   (*** pnat_le ***)
-
-(*** alternative definition for pnat_le ***)
-Goalw [pnat_le_def,pnat_less_def] 
-      "((m::pnat) <= n) = (Rep_pnat m <= Rep_pnat n)";
-by (auto_tac (claset() addSIs [leI] addSEs [leD],simpset()));
-qed "pnat_le_iff_Rep_pnat_le";
-
-Goal "!!k::pnat. (k + m <= k + n) = (m<=n)";
-by (simp_tac (simpset() addsimps [pnat_le_iff_Rep_pnat_le,
-                           sum_Rep_pnat_sum RS sym]) 1);
-qed "pnat_add_left_cancel_le";
-
-Goalw [pnat_less_def] "!!k::pnat. (k + m < k + n) = (m<n)";
-by (simp_tac (simpset() addsimps [sum_Rep_pnat_sum RS sym]) 1);
-qed "pnat_add_left_cancel_less";
-
-Addsimps [pnat_add_left_cancel, pnat_add_right_cancel,
-  pnat_add_left_cancel_le, pnat_add_left_cancel_less];
-
-Goalw [pnat_less_def] "i+j < (k::pnat) ==> i<k";
-by (auto_tac (claset() addEs [add_lessD1],
-    simpset() addsimps [sum_Rep_pnat_sum RS sym]));
-qed "pnat_add_lessD1";
-
-Goal "!!i::pnat. ~ (i+j < i)";
-by (rtac  notI 1);
-by (etac (pnat_add_lessD1 RS pnat_less_irrefl) 1);
-qed "pnat_not_add_less1";
-
-Goal "!!i::pnat. ~ (j+i < i)";
-by (simp_tac (simpset() addsimps [pnat_add_commute, pnat_not_add_less1]) 1);
-qed "pnat_not_add_less2";
-
-AddIffs [pnat_not_add_less1, pnat_not_add_less2];
-
-Goal "m + k <= n --> m <= (n::pnat)";
-by (simp_tac (simpset() addsimps [pnat_le_iff_Rep_pnat_le,
-                                  sum_Rep_pnat_sum RS sym]) 1);
-qed_spec_mp "pnat_add_leD1";
-
-Goal "!!n::pnat. m + k <= n ==> k <= n";
-by (full_simp_tac (simpset() addsimps [pnat_add_commute]) 1);
-by (etac pnat_add_leD1 1);
-qed_spec_mp "pnat_add_leD2";
-
-Goal "!!n::pnat. m + k <= n ==> m <= n & k <= n";
-by (blast_tac (claset() addDs [pnat_add_leD1, pnat_add_leD2]) 1);
-bind_thm ("pnat_add_leE", result() RS conjE);
-
-Goalw [pnat_less_def] 
-      "!!k l::pnat. [| k < l; m + l = k + n |] ==> m < n";
-by (rtac less_add_eq_less 1 THEN assume_tac 1);
-by (auto_tac (claset(),simpset() addsimps [sum_Rep_pnat_sum]));
-qed "pnat_less_add_eq_less";
-
-(* ordering on positive naturals in terms of existence of sum *)
-(* could provide alternative definition -- Gleason *)
-Goalw [pnat_less_def,pnat_add_def] 
-      "((z1::pnat) < z2) = (EX z3. z1 + z3 = z2)";
-by (rtac iffI 1);
-by (res_inst_tac [("t","z2")] (Rep_pnat_inverse RS subst) 1);
-by (dtac lemma_less_ex_sum_Rep_pnat 1);
-by (etac exE 1 THEN res_inst_tac [("x","z3")] exI 1);
-by (auto_tac (claset(),simpset() addsimps [sum_Rep_pnat_sum,Rep_pnat_inverse]));
-by (res_inst_tac [("t","Rep_pnat z1")] (add_0_right RS subst) 1);
-by (auto_tac (claset(),simpset() addsimps [sum_Rep_pnat_sum RS sym,
-               Rep_pnat_gt_zero] delsimps [add_0_right]));
-qed "pnat_less_iff";
-
-Goal "(EX (x::pnat). z1 + x = z2) | z1 = z2 \
-\          |(EX x. z2 + x = z1)";
-by (cut_facts_tac [pnat_less_linear] 1);
-by (asm_full_simp_tac (simpset() addsimps [pnat_less_iff]) 1);
-qed "pnat_linear_Ex_eq";
-
-Goal "!!(x::pnat). x + y = z ==> x < z";
-by (rtac (pnat_less_iff RS iffD2) 1);
-by (Blast_tac 1);
-qed "pnat_eq_lessI";
-
-(*** Monotonicity of Addition ***)
-
-Goal "1 * Rep_pnat n = Rep_pnat n";
-by (Asm_simp_tac 1);
-qed "Rep_pnat_mult_1";
-
-Goal "Rep_pnat n * 1 = Rep_pnat n";
-by (Asm_simp_tac 1);
-qed "Rep_pnat_mult_1_right";
-
-Goal "(Rep_pnat x * Rep_pnat y): pnat";
-by (cut_facts_tac [[Rep_pnat_gt_zero,
-    Rep_pnat_gt_zero] MRS mult_less_mono1,Collect_pnat_gt_0] 1);
-by (etac ssubst 1);
-by Auto_tac;
-qed "mult_Rep_pnat";
-
-Goalw [pnat_mult_def] 
-      "Rep_pnat x * Rep_pnat y = Rep_pnat (x * y)";
-by (simp_tac (simpset() addsimps [mult_Rep_pnat RS Abs_pnat_inverse]) 1);
-qed "mult_Rep_pnat_mult";
-
-Goalw [pnat_mult_def] "m * n = n * (m::pnat)";
-by (full_simp_tac (simpset() addsimps [mult_commute]) 1);
-qed "pnat_mult_commute";
-
-Goalw [pnat_mult_def,pnat_add_def] "(m + n)*k = (m*k) + ((n*k)::pnat)";
-by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
-by (simp_tac (simpset() addsimps [mult_Rep_pnat RS 
-                Abs_pnat_inverse,sum_Rep_pnat RS 
-             Abs_pnat_inverse, add_mult_distrib]) 1);
-qed "pnat_add_mult_distrib";
-
-Goalw [pnat_mult_def,pnat_add_def] "k*(m + n) = (k*m) + ((k*n)::pnat)";
-by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
-by (simp_tac (simpset() addsimps [mult_Rep_pnat RS 
-                Abs_pnat_inverse,sum_Rep_pnat RS 
-             Abs_pnat_inverse, add_mult_distrib2]) 1);
-qed "pnat_add_mult_distrib2";
-
-Goalw [pnat_mult_def] 
-      "(x * y) * z = x * (y * (z::pnat))";
-by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
-by (simp_tac (simpset() addsimps [mult_Rep_pnat RS 
-                Abs_pnat_inverse,mult_assoc]) 1);
-qed "pnat_mult_assoc";
-
-Goalw [pnat_mult_def] "x * (y * z) = y * (x * (z::pnat))";
-by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
-by (simp_tac (simpset() addsimps [mult_Rep_pnat RS 
-          Abs_pnat_inverse,mult_left_commute]) 1);
-qed "pnat_mult_left_commute";
-
-Goalw [pnat_mult_def] "x * (Abs_pnat (Suc 0)) = x";
-by (full_simp_tac (simpset() addsimps [one_RepI RS Abs_pnat_inverse,
-                   Rep_pnat_inverse]) 1);
-qed "pnat_mult_1";
-
-Goal "Abs_pnat (Suc 0) * x = x";
-by (full_simp_tac (simpset() addsimps [pnat_mult_1,
-                   pnat_mult_commute]) 1);
-qed "pnat_mult_1_left";
-
-(*Multiplication is an AC-operator*)
-bind_thms ("pnat_mult_ac", 
-	   [pnat_mult_assoc, pnat_mult_commute, pnat_mult_left_commute]);
-
-
-Goal "!!i::pnat. i<j ==> k*i < k*j";
-by (asm_full_simp_tac (simpset() addsimps [pnat_less_def,
-    mult_Rep_pnat_mult RS sym,Rep_pnat_gt_zero,mult_less_mono2]) 1);
-qed "pnat_mult_less_mono2";
-
-Goal "!!i::pnat. i<j ==> i*k < j*k";
-by (dtac pnat_mult_less_mono2 1);
-by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [pnat_mult_commute])));
-qed "pnat_mult_less_mono1";
-
-Goalw [pnat_less_def] "(m*(k::pnat) < n*k) = (m<n)";
-by (asm_full_simp_tac (simpset() addsimps [mult_Rep_pnat_mult 
-              RS sym,Rep_pnat_gt_zero]) 1);
-qed "pnat_mult_less_cancel2";
-
-Goalw [pnat_less_def] "((k::pnat)*m < k*n) = (m<n)";
-by (asm_full_simp_tac (simpset() addsimps [mult_Rep_pnat_mult 
-              RS sym,Rep_pnat_gt_zero]) 1);
-qed "pnat_mult_less_cancel1";
-
-Addsimps [pnat_mult_less_cancel1, pnat_mult_less_cancel2];
-
-Goalw [pnat_mult_def]  "(m*(k::pnat) = n*k) = (m=n)";
-by (cut_inst_tac [("x","k")] Rep_pnat_gt_zero 1);
-by (auto_tac (claset() addSDs [inj_on_Abs_pnat RS inj_onD,
-                               inj_Rep_pnat RS injD] 
-                       addIs [mult_Rep_pnat], 
-    simpset() addsimps [mult_cancel2]));
-qed "pnat_mult_cancel2";
-
-Goal "((k::pnat)*m = k*n) = (m=n)";
-by (rtac (pnat_mult_cancel2 RS subst) 1);
-by (auto_tac (claset () addIs [pnat_mult_commute RS subst],simpset()));
-qed "pnat_mult_cancel1";
-
-Addsimps [pnat_mult_cancel1, pnat_mult_cancel2];
-
-Goal "!!(z1::pnat). z2*z3 = z4*z5  ==> z2*(z1*z3) = z4*(z1*z5)";
-by (auto_tac (claset() addIs [pnat_mult_cancel1 RS iffD2],
-              simpset() addsimps [pnat_mult_left_commute]));
-qed "pnat_same_multI2";
-
-val [prem] = Goal
-    "(!!u. z = Abs_pnat(u) ==> P) ==> P";
-by (cut_inst_tac [("x1","z")] 
-    (rewrite_rule [pnat_def] (Rep_pnat RS Abs_pnat_inverse)) 1);
-by (res_inst_tac [("u","Rep_pnat z")] prem 1);
-by (dtac (inj_Rep_pnat RS injD) 1);
-by (Asm_simp_tac 1);
-qed "eq_Abs_pnat";
-
-(** embedding of naturals in positive naturals **)
-
-(* pnat_one_eq! *)
-Goalw [pnat_of_nat_def,pnat_one_def]"1 = pnat_of_nat 0";
-by (Full_simp_tac 1);
-qed "pnat_one_iff";
-
-Goalw [pnat_of_nat_def,pnat_one_def,
-       pnat_add_def] "1 + 1 = pnat_of_nat 1";
-by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
-by (auto_tac (claset() addIs [(gt_0_mem_pnat RS Abs_pnat_inverse RS ssubst)],
-    simpset()));
-qed "pnat_two_eq";
-
-Goal "inj(pnat_of_nat)";
-by (rtac injI 1);
-by (rewtac pnat_of_nat_def);
-by (dtac (inj_on_Abs_pnat RS inj_onD) 1);
-by (auto_tac (claset() addSIs [gt_0_mem_pnat],simpset()));
-qed "inj_pnat_of_nat";
-
-Goal "0 < n + (1::nat)";
-by Auto_tac;
-qed "nat_add_one_less";
-
-Goal "0 < n1 + n2 + (1::nat)";
-by Auto_tac;
-qed "nat_add_one_less1";
-
-(* this worked with one call to auto_tac before! *)
-Goalw [pnat_add_def,pnat_of_nat_def,pnat_one_def] 
-      "pnat_of_nat n1 + pnat_of_nat n2 = \
-\      pnat_of_nat (n1 + n2) + 1";
-by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
-by (rtac (gt_0_mem_pnat RS Abs_pnat_inverse RS ssubst) 1);
-by (rtac (gt_0_mem_pnat RS Abs_pnat_inverse RS ssubst) 2);
-by (rtac (gt_0_mem_pnat RS Abs_pnat_inverse RS ssubst) 3);
-by (rtac (gt_0_mem_pnat RS Abs_pnat_inverse RS ssubst) 4);
-by (auto_tac (claset(),
-	      simpset() addsimps [sum_Rep_pnat_sum,
-				  nat_add_one_less,nat_add_one_less1]));
-qed "pnat_of_nat_add";
-
-Goalw [pnat_of_nat_def,pnat_less_def] 
-       "(n < m) = (pnat_of_nat n < pnat_of_nat m)";
-by (auto_tac (claset(),simpset() 
-    addsimps [Abs_pnat_inverse,Collect_pnat_gt_0]));
-qed "pnat_of_nat_less_iff";
-Addsimps [pnat_of_nat_less_iff RS sym];
-
-Goalw [pnat_mult_def,pnat_of_nat_def] 
-      "pnat_of_nat n1 * pnat_of_nat n2 = \
-\      pnat_of_nat (n1 * n2 + n1 + n2)";
-by (auto_tac (claset(),simpset() addsimps [mult_Rep_pnat_mult,
-    pnat_add_def,Abs_pnat_inverse,gt_0_mem_pnat]));
-qed "pnat_of_nat_mult";
--- a/src/HOL/Real/PNat.thy	Tue Jan 27 09:44:14 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,45 +0,0 @@
-(*  Title       : PNat.thy
-    ID          : $Id$
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : The positive naturals
-*) 
-
-
-PNat = Main +
-
-typedef
-  pnat = "lfp(%X. {Suc 0} Un Suc`X)"   (lfp_def)
-
-instance
-   pnat :: {ord, one, plus, times}
-
-consts
-
-  pSuc       :: pnat => pnat
-
-constdefs
-  
-  pnat_of_nat  :: nat => pnat           
-  "pnat_of_nat n     == Abs_pnat(n + 1)"
- 
-defs
-
-  pnat_one_def      
-       "1 == Abs_pnat(Suc 0)"
-  pnat_Suc_def      
-       "pSuc == (%n. Abs_pnat(Suc(Rep_pnat(n))))"
-
-  pnat_add_def
-       "x + y == Abs_pnat(Rep_pnat(x) +  Rep_pnat(y))"
-
-  pnat_mult_def
-       "x * y == Abs_pnat(Rep_pnat(x) * Rep_pnat(y))"
-
-  pnat_less_def
-       "x < (y::pnat) == Rep_pnat(x) < Rep_pnat(y)"
-
-  pnat_le_def
-       "x <= (y::pnat) ==  ~(y < x)"
-
-end
--- a/src/HOL/Real/PRat.ML	Tue Jan 27 09:44:14 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,805 +0,0 @@
-(*  Title       : PRat.ML
-    ID          : $Id$
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : The positive rationals
-*) 
-
-(*** Many theorems similar to those in theory Integ ***)
-(*** Proving that ratrel is an equivalence relation ***)
-
-Goal "[| (x1::pnat) * y2 = x2 * y1; x2 * y3 = x3 * y2 |] \
-\            ==> x1 * y3 = x3 * y1";        
-by (res_inst_tac [("k1","y2")] (pnat_mult_cancel1 RS iffD1) 1);
-by (auto_tac (claset(), simpset() addsimps [pnat_mult_assoc RS sym]));
-by (auto_tac (claset(),simpset() addsimps [pnat_mult_commute]));
-by (dres_inst_tac [("s","x2 * y3")] sym 1);
-by (asm_simp_tac (simpset() addsimps [pnat_mult_left_commute,
-    pnat_mult_commute]) 1);
-qed "prat_trans_lemma";
-
-(** Natural deduction for ratrel **)
-
-Goalw [ratrel_def]
-    "(((x1,y1),(x2,y2)): ratrel) = (x1 * y2 = x2 * y1)";
-by (Fast_tac 1);
-qed "ratrel_iff";
-
-Goalw [ratrel_def]
-    "[| x1 * y2 = x2 * y1 |] ==> ((x1,y1),(x2,y2)): ratrel";
-by (Fast_tac  1);
-qed "ratrelI";
-
-Goalw [ratrel_def]
-  "p: ratrel --> (EX x1 y1 x2 y2. \
-\                  p = ((x1,y1),(x2,y2)) & x1 *y2 = x2 *y1)";
-by (Fast_tac 1);
-qed "ratrelE_lemma";
-
-val [major,minor] = Goal
-  "[| p: ratrel;  \
-\     !!x1 y1 x2 y2. [| p = ((x1,y1),(x2,y2));  x1*y2 = x2*y1 \
-\                    |] ==> Q |] ==> Q";
-by (cut_facts_tac [major RS (ratrelE_lemma RS mp)] 1);
-by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
-qed "ratrelE";
-
-AddSIs [ratrelI];
-AddSEs [ratrelE];
-
-Goal "(x,x): ratrel";
-by (pair_tac "x" 1);
-by (rtac ratrelI 1);
-by (rtac refl 1);
-qed "ratrel_refl";
-
-Goalw [equiv_def, refl_def, sym_def, trans_def]
-    "equiv UNIV ratrel";
-by (fast_tac (claset() addSIs [ratrel_refl] 
-                       addSEs [sym, prat_trans_lemma]) 1);
-qed "equiv_ratrel";
-
-bind_thm ("equiv_ratrel_iff", [equiv_ratrel, UNIV_I, UNIV_I] MRS eq_equiv_class_iff);
-
-Goalw  [prat_def,ratrel_def,quotient_def] "ratrel``{(x,y)}:prat";
-by (Blast_tac 1);
-qed "ratrel_in_prat";
-
-Goal "inj_on Abs_prat prat";
-by (rtac inj_on_inverseI 1);
-by (etac Abs_prat_inverse 1);
-qed "inj_on_Abs_prat";
-
-Addsimps [equiv_ratrel_iff,inj_on_Abs_prat RS inj_on_iff,
-          ratrel_iff, ratrel_in_prat, Abs_prat_inverse];
-
-Addsimps [equiv_ratrel RS eq_equiv_class_iff];
-bind_thm ("eq_ratrelD", equiv_ratrel RSN (2,eq_equiv_class));
-
-Goal "inj(Rep_prat)";
-by (rtac inj_inverseI 1);
-by (rtac Rep_prat_inverse 1);
-qed "inj_Rep_prat";
-
-(** prat_of_pnat: the injection from pnat to prat **)
-Goal "inj(prat_of_pnat)";
-by (rtac injI 1);
-by (rewtac prat_of_pnat_def);
-by (dtac (inj_on_Abs_prat RS inj_onD) 1);
-by (REPEAT (rtac ratrel_in_prat 1));
-by (dtac eq_equiv_class 1);
-by (rtac equiv_ratrel 1);
-by (Fast_tac 1);
-by Safe_tac;
-by (Asm_full_simp_tac 1);
-qed "inj_prat_of_pnat";
-
-val [prem] = Goal
-    "(!!x y. z = Abs_prat(ratrel``{(x,y)}) ==> P) ==> P";
-by (res_inst_tac [("x1","z")] 
-    (rewrite_rule [prat_def] Rep_prat RS quotientE) 1);
-by (dres_inst_tac [("f","Abs_prat")] arg_cong 1);
-by (res_inst_tac [("p","x")] PairE 1);
-by (rtac prem 1);
-by (asm_full_simp_tac (simpset() addsimps [Rep_prat_inverse]) 1);
-qed "eq_Abs_prat";
-
-(**** qinv: inverse on prat ****)
-
-Goalw [congruent_def] "congruent ratrel (%(x,y). ratrel``{(y,x)})";
-by (auto_tac (claset(), simpset() addsimps [pnat_mult_commute]));  
-qed "qinv_congruent";
-
-Goalw [qinv_def]
-      "qinv (Abs_prat(ratrel``{(x,y)})) = Abs_prat(ratrel `` {(y,x)})";
-by (simp_tac (simpset() addsimps 
-	      [equiv_ratrel RS UN_equiv_class, qinv_congruent]) 1);
-qed "qinv";
-
-Goal "qinv (qinv z) = z";
-by (res_inst_tac [("z","z")] eq_Abs_prat 1);
-by (asm_simp_tac (simpset() addsimps [qinv]) 1);
-qed "qinv_qinv";
-
-Goal "inj(qinv)";
-by (rtac injI 1);
-by (dres_inst_tac [("f","qinv")] arg_cong 1);
-by (asm_full_simp_tac (simpset() addsimps [qinv_qinv]) 1);
-qed "inj_qinv";
-
-Goalw [prat_of_pnat_def] 
-      "qinv(prat_of_pnat  (Abs_pnat (Suc 0))) = prat_of_pnat (Abs_pnat (Suc 0))";
-by (simp_tac (simpset() addsimps [qinv]) 1);
-qed "qinv_1";
-
-Goal "!!(x1::pnat). [| x1 * y2 = x2 * y1 |] ==> \
-\     (x * y1 + x1 * ya) * (ya * y2) = (x * y2 + x2 * ya) * (ya * y1)";
-by (auto_tac (claset() addSIs [pnat_same_multI2],
-       simpset() addsimps [pnat_add_mult_distrib,
-       pnat_mult_assoc]));
-by (res_inst_tac [("n1","y2")] (pnat_mult_commute RS subst) 1);
-by (auto_tac (claset() addIs [pnat_add_left_cancel RS iffD2],simpset() addsimps pnat_mult_ac));
-by (res_inst_tac [("y1","x1")] (pnat_mult_left_commute RS subst) 1);
-by (res_inst_tac [("y1","x1")] (pnat_mult_left_commute RS ssubst) 1);
-by (auto_tac (claset(),simpset() addsimps [pnat_mult_assoc RS sym]));
-qed "prat_add_congruent2_lemma";
-
-Goal "congruent2 ratrel (%p1 p2.                  \
-\        (%(x1,y1). (%(x2,y2). ratrel``{(x1*y2 + x2*y1, y1*y2)}) p2) p1)";
-by (rtac (equiv_ratrel RS congruent2_commuteI) 1);
-by (auto_tac (claset() delrules [equalityI],
-              simpset() addsimps [prat_add_congruent2_lemma]));
-by (asm_simp_tac (simpset() addsimps [pnat_mult_commute,pnat_add_commute]) 1);
-qed "prat_add_congruent2";
-
-Goalw [prat_add_def]
-   "Abs_prat((ratrel``{(x1,y1)})) + Abs_prat((ratrel``{(x2,y2)})) =   \
-\   Abs_prat(ratrel `` {(x1*y2 + x2*y1, y1*y2)})";
-by (simp_tac (simpset() addsimps [UN_UN_split_split_eq, prat_add_congruent2, 
-				  equiv_ratrel RS UN_equiv_class2]) 1);
-qed "prat_add";
-
-Goal "(z::prat) + w = w + z";
-by (res_inst_tac [("z","z")] eq_Abs_prat 1);
-by (res_inst_tac [("z","w")] eq_Abs_prat 1);
-by (asm_simp_tac
-    (simpset() addsimps [prat_add] @ pnat_add_ac @ pnat_mult_ac) 1);
-qed "prat_add_commute";
-
-Goal "((z1::prat) + z2) + z3 = z1 + (z2 + z3)";
-by (res_inst_tac [("z","z1")] eq_Abs_prat 1);
-by (res_inst_tac [("z","z2")] eq_Abs_prat 1);
-by (res_inst_tac [("z","z3")] eq_Abs_prat 1);
-by (asm_simp_tac (simpset() addsimps [pnat_add_mult_distrib2,prat_add] @ 
-		                     pnat_add_ac @ pnat_mult_ac) 1);
-qed "prat_add_assoc";
-
-(*For AC rewriting*)
-Goal "(z1::prat) + (z2 + z3) = z2 + (z1 + z3)";
-by(rtac ([prat_add_assoc,prat_add_commute] MRS
-         read_instantiate[("f","op +")](thm"mk_left_commute")) 1);
-qed "prat_add_left_commute";
-
-(* Positive Rational addition is an AC operator *)
-bind_thms ("prat_add_ac", [prat_add_assoc, prat_add_commute, prat_add_left_commute]);
-
-
-(*** Congruence property for multiplication ***)
-
-Goalw [congruent2_def]
-    "congruent2 ratrel (%p1 p2.                  \
-\         (%(x1,y1). (%(x2,y2). ratrel``{(x1*x2, y1*y2)}) p2) p1)";
-(*Proof via congruent2_commuteI seems longer*)
-by (Clarify_tac 1);
-by (asm_simp_tac (simpset() addsimps [pnat_mult_assoc]) 1);
-(*The rest should be trivial, but rearranging terms is hard*)
-by (res_inst_tac [("x1","x1a")] (pnat_mult_left_commute RS ssubst) 1);
-by (asm_simp_tac (simpset() addsimps [pnat_mult_assoc RS sym]) 1);
-by (asm_simp_tac (simpset() addsimps pnat_mult_ac) 1);
-qed "pnat_mult_congruent2";
-
-Goalw [prat_mult_def]
-  "Abs_prat(ratrel``{(x1,y1)}) * Abs_prat(ratrel``{(x2,y2)}) = \
-\  Abs_prat(ratrel``{(x1*x2, y1*y2)})";
-by (asm_simp_tac 
-    (simpset() addsimps [UN_UN_split_split_eq, pnat_mult_congruent2,
-			 equiv_ratrel RS UN_equiv_class2]) 1);
-qed "prat_mult";
-
-Goal "(z::prat) * w = w * z";
-by (res_inst_tac [("z","z")] eq_Abs_prat 1);
-by (res_inst_tac [("z","w")] eq_Abs_prat 1);
-by (asm_simp_tac (simpset() addsimps pnat_mult_ac @ [prat_mult]) 1);
-qed "prat_mult_commute";
-
-Goal "((z1::prat) * z2) * z3 = z1 * (z2 * z3)";
-by (res_inst_tac [("z","z1")] eq_Abs_prat 1);
-by (res_inst_tac [("z","z2")] eq_Abs_prat 1);
-by (res_inst_tac [("z","z3")] eq_Abs_prat 1);
-by (asm_simp_tac (simpset() addsimps [prat_mult, pnat_mult_assoc]) 1);
-qed "prat_mult_assoc";
-
-(*For AC rewriting*)
-Goal "(x::prat)*(y*z)=y*(x*z)";
-by(rtac ([prat_mult_assoc,prat_mult_commute] MRS
-         read_instantiate[("f","op *")](thm"mk_left_commute")) 1);
-qed "prat_mult_left_commute";
-
-(*Positive Rational multiplication is an AC operator*)
-bind_thms ("prat_mult_ac", [prat_mult_assoc,
-                    prat_mult_commute,prat_mult_left_commute]);
-
-Goalw [prat_of_pnat_def] 
-      "(prat_of_pnat (Abs_pnat (Suc 0))) * z = z";
-by (res_inst_tac [("z","z")] eq_Abs_prat 1);
-by (asm_full_simp_tac (simpset() addsimps [prat_mult] @ pnat_mult_ac) 1);
-qed "prat_mult_1";
-
-Goalw [prat_of_pnat_def] 
-      "z * (prat_of_pnat (Abs_pnat (Suc 0))) = z";
-by (res_inst_tac [("z","z")] eq_Abs_prat 1);
-by (asm_full_simp_tac (simpset() addsimps [prat_mult] @ pnat_mult_ac) 1);
-qed "prat_mult_1_right";
-
-Goalw [prat_of_pnat_def] 
-      "prat_of_pnat ((z1::pnat) + z2) = \
-\      prat_of_pnat z1 + prat_of_pnat z2";
-by (asm_simp_tac (simpset() addsimps [prat_add,
-				      pnat_add_mult_distrib,pnat_mult_1]) 1);
-qed "prat_of_pnat_add";
-
-Goalw [prat_of_pnat_def] 
-      "prat_of_pnat ((z1::pnat) * z2) = \
-\      prat_of_pnat z1 * prat_of_pnat z2";
-by (asm_simp_tac (simpset() addsimps [prat_mult, pnat_mult_1]) 1);
-qed "prat_of_pnat_mult";
-
-(*** prat_mult and qinv ***)
-
-Goalw [prat_def,prat_of_pnat_def] 
-      "qinv (q) * q = prat_of_pnat (Abs_pnat (Suc 0))";
-by (res_inst_tac [("z","q")] eq_Abs_prat 1);
-by (asm_full_simp_tac (simpset() addsimps [qinv,
-        prat_mult,pnat_mult_1,pnat_mult_1_left, pnat_mult_commute]) 1);
-qed "prat_mult_qinv";
-
-Goal "q * qinv (q) = prat_of_pnat (Abs_pnat (Suc 0))";
-by (rtac (prat_mult_commute RS subst) 1);
-by (simp_tac (simpset() addsimps [prat_mult_qinv]) 1);
-qed "prat_mult_qinv_right";
-
-Goal "EX y. (x::prat) * y = prat_of_pnat (Abs_pnat (Suc 0))";
-by (fast_tac (claset() addIs [prat_mult_qinv_right]) 1);
-qed "prat_qinv_ex";
-
-Goal "EX! y. (x::prat) * y = prat_of_pnat (Abs_pnat (Suc 0))";
-by (auto_tac (claset() addIs [prat_mult_qinv_right],simpset()));
-by (dres_inst_tac [("f","%x. ya*x")] arg_cong 1);
-by (asm_full_simp_tac (simpset() addsimps [prat_mult_assoc RS sym]) 1);
-by (asm_full_simp_tac (simpset() addsimps [prat_mult_commute,
-    prat_mult_1,prat_mult_1_right]) 1);
-qed "prat_qinv_ex1";
-
-Goal "EX! y. y * (x::prat) = prat_of_pnat (Abs_pnat (Suc 0))";
-by (auto_tac (claset() addIs [prat_mult_qinv],simpset()));
-by (dres_inst_tac [("f","%x. x*ya")] arg_cong 1);
-by (asm_full_simp_tac (simpset() addsimps [prat_mult_assoc]) 1);
-by (asm_full_simp_tac (simpset() addsimps [prat_mult_commute,
-    prat_mult_1,prat_mult_1_right]) 1);
-qed "prat_qinv_left_ex1";
-
-Goal "x * y = prat_of_pnat (Abs_pnat (Suc 0)) ==> x = qinv y";
-by (cut_inst_tac [("q","y")] prat_mult_qinv 1);
-by (res_inst_tac [("x1","y")] (prat_qinv_left_ex1 RS ex1E) 1);
-by (Blast_tac 1);
-qed "prat_mult_inv_qinv";
-
-Goal "EX y. x = qinv y";
-by (cut_inst_tac [("x","x")] prat_qinv_ex 1);
-by (etac exE 1 THEN dtac prat_mult_inv_qinv 1);
-by (Fast_tac 1);
-qed "prat_as_inverse_ex";
-
-Goal "qinv(x*y) = qinv(x)*qinv(y)";
-by (res_inst_tac [("z","x")] eq_Abs_prat 1);
-by (res_inst_tac [("z","y")] eq_Abs_prat 1);
-by (auto_tac (claset(),simpset() addsimps [qinv,prat_mult]));
-qed "qinv_mult_eq";
-
-(** Lemmas **)
-
-Goal "((z1::prat) + z2) * w = (z1 * w) + (z2 * w)";
-by (res_inst_tac [("z","z1")] eq_Abs_prat 1);
-by (res_inst_tac [("z","z2")] eq_Abs_prat 1);
-by (res_inst_tac [("z","w")] eq_Abs_prat 1);
-by (asm_simp_tac 
-    (simpset() addsimps [pnat_add_mult_distrib2, prat_add, prat_mult] @ 
-                        pnat_add_ac @ pnat_mult_ac) 1);
-qed "prat_add_mult_distrib";
-
-val prat_mult_commute'= read_instantiate [("z","w")] prat_mult_commute;
-
-Goal "(w::prat) * (z1 + z2) = (w * z1) + (w * z2)";
-by (simp_tac (simpset() addsimps [prat_mult_commute',prat_add_mult_distrib]) 1);
-qed "prat_add_mult_distrib2";
-
-Addsimps [prat_mult_1, prat_mult_1_right, 
-	  prat_mult_qinv, prat_mult_qinv_right];
-
-      (*** theorems for ordering ***)
-(* prove introduction and elimination rules for prat_less *)
-
-Goalw [prat_less_def]
-    "(Q1 < (Q2::prat)) = (EX Q3. Q1 + Q3 = Q2)";
-by (Fast_tac 1);
-qed "prat_less_iff";
-
-Goalw [prat_less_def]
-      "!!(Q1::prat). Q1 + Q3 = Q2 ==> Q1 < Q2";
-by (Fast_tac  1);
-qed "prat_lessI";
-
-(* ordering on positive fractions in terms of existence of sum *)
-Goalw [prat_less_def]
-      "Q1 < (Q2::prat) --> (EX Q3. Q1 + Q3 = Q2)";
-by (Fast_tac 1);
-qed "prat_lessE_lemma";
-
-Goal "!!P. [| Q1 < (Q2::prat); \
-\             !! (Q3::prat). Q1 + Q3 = Q2 ==> P |] \
-\          ==> P";
-by (dtac (prat_lessE_lemma RS mp) 1);
-by Auto_tac;
-qed "prat_lessE";
-
-(* qless is a strong order i.e nonreflexive and transitive *)
-Goal "!!(q1::prat). [| q1 < q2; q2 < q3 |] ==> q1 < q3";
-by (REPEAT(dtac (prat_lessE_lemma RS mp) 1));
-by (REPEAT(etac exE 1));
-by (hyp_subst_tac 1);
-by (res_inst_tac [("Q3.0","Q3 + Q3a")] prat_lessI 1);
-by (auto_tac (claset(),simpset() addsimps [prat_add_assoc]));
-qed "prat_less_trans";
-
-Goal "~q < (q::prat)";
-by (EVERY1[rtac notI, dtac (prat_lessE_lemma RS mp)]);
-by (res_inst_tac [("z","q")] eq_Abs_prat 1);
-by (res_inst_tac [("z","Q3")] eq_Abs_prat 1);
-by (etac exE 1 THEN res_inst_tac [("z","Q3a")] eq_Abs_prat 1);
-by (REPEAT(hyp_subst_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [prat_add,
-    pnat_no_add_ident,pnat_add_mult_distrib2] @ pnat_mult_ac) 1);
-qed "prat_less_not_refl";
-
-(*** y < y ==> P ***)
-bind_thm("prat_less_irrefl",prat_less_not_refl RS notE);
-
-Goal "!! (q1::prat). q1 < q2 ==> ~ q2 < q1";
-by (rtac notI 1);
-by (dtac prat_less_trans 1 THEN assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [prat_less_not_refl]) 1);
-qed "prat_less_not_sym";
-
-(* [| x < y;  ~P ==> y < x |] ==> P *)
-bind_thm ("prat_less_asym", prat_less_not_sym RS contrapos_np);
-
-(* half of positive fraction exists- Gleason p. 120- Proposition 9-2.6(i)*)
-Goal "!(q::prat). EX x. x + x = q";
-by (rtac allI 1);
-by (res_inst_tac [("z","q")] eq_Abs_prat 1);
-by (res_inst_tac [("x","Abs_prat (ratrel `` {(x, y+y)})")] exI 1);
-by (auto_tac (claset(),
-	      simpset() addsimps 
-              [prat_add,pnat_mult_assoc RS sym,pnat_add_mult_distrib,
-               pnat_add_mult_distrib2]));
-qed "lemma_prat_dense";
-
-Goal "EX (x::prat). x + x = q";
-by (res_inst_tac [("z","q")] eq_Abs_prat 1);
-by (res_inst_tac [("x","Abs_prat (ratrel `` {(x, y+y)})")] exI 1);
-by (auto_tac (claset(),simpset() addsimps 
-              [prat_add,pnat_mult_assoc RS sym,pnat_add_mult_distrib,
-               pnat_add_mult_distrib2]));
-qed "prat_lemma_dense";
-
-(* there exists a number between any two positive fractions *)
-(* Gleason p. 120- Proposition 9-2.6(iv) *)
-Goalw [prat_less_def] 
-      "!! (q1::prat). q1 < q2 ==> EX x. q1 < x & x < q2";
-by (auto_tac (claset() addIs [lemma_prat_dense],simpset()));
-by (res_inst_tac [("x","T")] (lemma_prat_dense RS allE) 1);
-by (etac exE 1);
-by (res_inst_tac [("x","q1 + x")] exI 1);
-by (auto_tac (claset() addIs [prat_lemma_dense],
-	      simpset() addsimps [prat_add_assoc]));
-qed "prat_dense";
-
-(* ordering of addition for positive fractions *)
-Goalw [prat_less_def] "!!(q1::prat). q1 < q2 ==> q1 + x < q2 + x";
-by (Step_tac 1);
-by (res_inst_tac [("x","T")] exI 1);
-by (auto_tac (claset(),simpset() addsimps prat_add_ac));
-qed "prat_add_less2_mono1";
-
-Goal "!!(q1::prat). q1 < q2 ==> x + q1 < x + q2";
-by (auto_tac (claset() addIs [prat_add_less2_mono1],
-    simpset() addsimps [prat_add_commute]));
-qed "prat_add_less2_mono2";
-
-(* ordering of multiplication for positive fractions *)
-Goalw [prat_less_def] 
-      "!!(q1::prat). q1 < q2 ==> q1 * x < q2 * x";
-by (Step_tac 1);
-by (res_inst_tac [("x","T*x")] exI 1);
-by (auto_tac (claset(),simpset() addsimps [prat_add_mult_distrib]));
-qed "prat_mult_less2_mono1";
-
-Goal "!!(q1::prat). q1 < q2  ==> x * q1 < x * q2";
-by (auto_tac (claset() addDs [prat_mult_less2_mono1],
-    simpset() addsimps [prat_mult_commute]));
-qed "prat_mult_left_less2_mono1";
-
-Goal "!!(a1::prat). a1 < a2 ==> a1 * b + a2 * c < a2 * (b + c)";
-by (auto_tac (claset() addSIs [prat_add_less2_mono1,prat_mult_less2_mono1],
-              simpset() addsimps [prat_add_mult_distrib2]));
-qed "lemma_prat_add_mult_mono";
-
-(* there is no smallest positive fraction *)
-Goalw [prat_less_def] "EX (x::prat). x < y";
-by (cut_facts_tac [lemma_prat_dense] 1);
-by (Fast_tac 1);
-qed "qless_Ex";
-
-(* lemma for proving $< is linear *)
-Goalw [prat_def,prat_less_def] 
-      "ratrel `` {(x, y * ya)} : {p::(pnat*pnat).True}//ratrel";
-by (asm_full_simp_tac (simpset() addsimps [ratrel_def,quotient_def]) 1);
-by (Blast_tac 1);
-qed "lemma_prat_less_linear";
-
-(* linearity of < -- Gleason p. 120 - Proposition 9-2.6 *)
-(*** FIXME Proof long ***)
-Goalw [prat_less_def] 
-      "(q1::prat) < q2 | q1 = q2 | q2 < q1";
-by (res_inst_tac [("z","q1")] eq_Abs_prat 1);
-by (res_inst_tac [("z","q2")] eq_Abs_prat 1);
-by (Step_tac 1 THEN REPEAT(dtac (not_ex RS iffD1) 1) 
-               THEN Auto_tac);
-by (cut_inst_tac  [("z1.0","x*ya"), ("z2.0","xa*y")] pnat_linear_Ex_eq 1);
-by (EVERY1[etac disjE,etac exE]);
-by (eres_inst_tac 
-    [("x","Abs_prat(ratrel``{(xb,ya*y)})")] allE 1);
-by (asm_full_simp_tac 
-    (simpset() addsimps [prat_add, pnat_mult_assoc 
-     RS sym,pnat_add_mult_distrib RS sym]) 1);
-by (EVERY1[asm_full_simp_tac (simpset() addsimps pnat_mult_ac),
-    etac disjE, assume_tac, etac exE]);
-by (thin_tac "!T. Abs_prat (ratrel `` {(x, y)}) + T ~= \
-\     Abs_prat (ratrel `` {(xa, ya)})" 1);
-by (eres_inst_tac [("x","Abs_prat(ratrel``{(xb,y*ya)})")] allE 1);
-by (asm_full_simp_tac (simpset() addsimps [prat_add,
-      pnat_mult_assoc RS sym,pnat_add_mult_distrib RS sym]) 1);
-by (asm_full_simp_tac (simpset() addsimps pnat_mult_ac) 1);
-qed "prat_linear";
-
-Goal "!!(x::prat). [| x < y ==> P;  x = y ==> P; \
-\          y < x ==> P |] ==> P";
-by (cut_inst_tac [("q1.0","x"),("q2.0","y")] prat_linear 1);
-by Auto_tac;
-qed "prat_linear_less2";
-
-(* Gleason p. 120 -- 9-2.6 (iv) *)
-Goal "[| q1 < q2; qinv(q1) = qinv(q2) |] ==> P";
-by (cut_inst_tac [("x","qinv (q2)"),("q1.0","q1"), ("q2.0","q2")] 
-    prat_mult_less2_mono1 1);
-by (assume_tac 1);
-by (Asm_full_simp_tac 1 THEN dtac sym 1);
-by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl]));
-qed "lemma1_qinv_prat_less";
-
-Goal "[| q1 < q2; qinv(q1) < qinv(q2) |] ==> P";
-by (cut_inst_tac [("x","qinv (q2)"),("q1.0","q1"), ("q2.0","q2")] 
-    prat_mult_less2_mono1 1);
-by (assume_tac 1);
-by (cut_inst_tac [("x","q1"),("q1.0","qinv (q1)"), ("q2.0","qinv (q2)")] 
-    prat_mult_left_less2_mono1 1);
-by Auto_tac;
-by (dres_inst_tac [("q2.0","prat_of_pnat (Abs_pnat (Suc 0))")] prat_less_trans 1);
-by (auto_tac (claset(),simpset() addsimps 
-    [prat_less_not_refl]));
-qed "lemma2_qinv_prat_less";
-
-Goal "q1 < q2  ==> qinv (q2) < qinv (q1)";
-by (res_inst_tac [("y","qinv q1"), ("x","qinv q2")] prat_linear_less2 1);
-by (auto_tac (claset() addEs [lemma1_qinv_prat_less,
-                 lemma2_qinv_prat_less],simpset()));
-qed "qinv_prat_less";
-
-Goal "q1 < prat_of_pnat (Abs_pnat (Suc 0)) \
-\     ==> prat_of_pnat (Abs_pnat (Suc 0)) < qinv(q1)";
-by (dtac qinv_prat_less 1);
-by (full_simp_tac (simpset() addsimps [qinv_1]) 1);
-qed "prat_qinv_gt_1";
-
-Goalw [pnat_one_def] 
-     "q1 < prat_of_pnat 1 ==> prat_of_pnat 1 < qinv(q1)";
-by (etac prat_qinv_gt_1 1);
-qed "prat_qinv_is_gt_1";
-
-Goalw [prat_less_def] 
-      "prat_of_pnat (Abs_pnat (Suc 0)) < prat_of_pnat (Abs_pnat (Suc 0)) \
-\                   + prat_of_pnat (Abs_pnat (Suc 0))";
-by (Fast_tac 1); 
-qed "prat_less_1_2";
-
-Goal "qinv(prat_of_pnat (Abs_pnat (Suc 0)) + \
-\     prat_of_pnat (Abs_pnat (Suc 0))) < prat_of_pnat (Abs_pnat (Suc 0))";
-by (cut_facts_tac [prat_less_1_2 RS qinv_prat_less] 1);
-by (asm_full_simp_tac (simpset() addsimps [qinv_1]) 1);
-qed "prat_less_qinv_2_1";
-
-Goal "!!(x::prat). x < y ==> x*qinv(y) < prat_of_pnat (Abs_pnat (Suc 0))";
-by (dres_inst_tac [("x","qinv(y)")] prat_mult_less2_mono1 1);
-by (Asm_full_simp_tac 1);
-qed "prat_mult_qinv_less_1";
-
-Goal "(x::prat) < x + x";
-by (cut_inst_tac [("x","x")] 
-    (prat_less_1_2 RS prat_mult_left_less2_mono1) 1);
-by (asm_full_simp_tac (simpset() addsimps 
-    [prat_add_mult_distrib2]) 1);
-qed "prat_self_less_add_self";
-
-Goalw [prat_less_def] "(x::prat) < y + x";
-by (res_inst_tac [("x","y")] exI 1);
-by (simp_tac (simpset() addsimps [prat_add_commute]) 1);
-qed "prat_self_less_add_right";
-
-Goal "(x::prat) < x + y";
-by (rtac (prat_add_commute RS subst) 1);
-by (simp_tac (simpset() addsimps [prat_self_less_add_right]) 1);
-qed "prat_self_less_add_left";
-
-Goalw [prat_less_def] "prat_of_pnat 1 < y ==> (x::prat) < x * y";
-by (auto_tac (claset(),simpset() addsimps [pnat_one_def,
-    prat_add_mult_distrib2]));
-qed "prat_self_less_mult_right";
-
-(*** Properties of <= ***)
-
-Goalw [prat_le_def] "~(w < z) ==> z <= (w::prat)";
-by (assume_tac 1);
-qed "prat_leI";
-
-Goalw [prat_le_def] "z<=w ==> ~(w<(z::prat))";
-by (assume_tac 1);
-qed "prat_leD";
-
-bind_thm ("prat_leE", make_elim prat_leD);
-
-Goal "(~(w < z)) = (z <= (w::prat))";
-by (fast_tac (claset() addSIs [prat_leI,prat_leD]) 1);
-qed "prat_less_le_iff";
-
-Goalw [prat_le_def] "~ z <= w ==> w<(z::prat)";
-by (Fast_tac 1);
-qed "not_prat_leE";
-
-Goalw [prat_le_def] "z < w ==> z <= (w::prat)";
-by (fast_tac (claset() addEs [prat_less_asym]) 1);
-qed "prat_less_imp_le";
-
-Goalw [prat_le_def] "!!(x::prat). x <= y ==> x < y | x = y";
-by (cut_facts_tac [prat_linear] 1);
-by (fast_tac (claset() addEs [prat_less_irrefl,prat_less_asym]) 1);
-qed "prat_le_imp_less_or_eq";
-
-Goalw [prat_le_def] "z<w | z=w ==> z <=(w::prat)";
-by (cut_facts_tac [prat_linear] 1);
-by (fast_tac (claset() addEs [prat_less_irrefl,prat_less_asym]) 1);
-qed "prat_less_or_eq_imp_le";
-
-Goal "(x <= (y::prat)) = (x < y | x=y)";
-by (REPEAT(ares_tac [iffI, prat_less_or_eq_imp_le, prat_le_imp_less_or_eq] 1));
-qed "prat_le_eq_less_or_eq";
-
-Goal "w <= (w::prat)";
-by (simp_tac (simpset() addsimps [prat_le_eq_less_or_eq]) 1);
-qed "prat_le_refl";
-
-Goal "[| i <= j; j < k |] ==> i < (k::prat)";
-by (dtac prat_le_imp_less_or_eq 1);
-by (fast_tac (claset() addIs [prat_less_trans]) 1);
-qed "prat_le_less_trans";
-
-Goal "[| i <= j; j <= k |] ==> i <= (k::prat)";
-by (EVERY1 [dtac prat_le_imp_less_or_eq, dtac prat_le_imp_less_or_eq,
-            rtac prat_less_or_eq_imp_le, fast_tac (claset() addIs [prat_less_trans])]);
-qed "prat_le_trans";
-
-Goal "[| ~ y < x; y ~= x |] ==> x < (y::prat)";
-by (rtac not_prat_leE 1);
-by (fast_tac (claset() addDs [prat_le_imp_less_or_eq]) 1);
-qed "not_less_not_eq_prat_less";
-
-Goalw [prat_less_def] 
-      "[| x1 < y1; x2 < y2 |] ==> x1 + x2 < y1 + (y2::prat)";
-by (REPEAT(etac exE 1));
-by (res_inst_tac [("x","T+Ta")] exI 1);
-by (auto_tac (claset(),simpset() addsimps prat_add_ac));
-qed "prat_add_less_mono";
-
-Goalw [prat_less_def] 
-      "[| x1 < y1; x2 < y2 |] ==> x1 * x2 < y1 * (y2::prat)";
-by (REPEAT(etac exE 1));
-by (res_inst_tac [("x","T*Ta+T*x2+x1*Ta")] exI 1);
-by (auto_tac (claset(),
-	      simpset() addsimps prat_add_ac @ 
-	                      [prat_add_mult_distrib,prat_add_mult_distrib2]));
-qed "prat_mult_less_mono";
-
-(* more prat_le *)
-Goal "!!(q1::prat). q1 <= q2  ==> x * q1 <= x * q2";
-by (dtac prat_le_imp_less_or_eq 1);
-by (Step_tac 1);
-by (auto_tac (claset() addSIs [prat_le_refl, prat_less_imp_le,
-			       prat_mult_left_less2_mono1],
-	      simpset()));
-qed "prat_mult_left_le2_mono1";
-
-Goal "!!(q1::prat). q1 <= q2  ==> q1 * x <= q2 * x";
-by (auto_tac (claset() addDs [prat_mult_left_le2_mono1],
-	      simpset() addsimps [prat_mult_commute]));
-qed "prat_mult_le2_mono1";
-
-Goal "q1 <= q2  ==> qinv (q2) <= qinv (q1)";
-by (dtac prat_le_imp_less_or_eq 1);
-by (Step_tac 1);
-by (auto_tac (claset() addSIs [prat_le_refl, prat_less_imp_le,qinv_prat_less],
-	      simpset()));
-qed "qinv_prat_le";
-
-Goal "!!(q1::prat). q1 <= q2  ==> x + q1 <= x + q2";
-by (dtac prat_le_imp_less_or_eq 1);
-by (Step_tac 1);
-by (auto_tac (claset() addSIs [prat_le_refl,
-			       prat_less_imp_le,prat_add_less2_mono1],
-	      simpset() addsimps [prat_add_commute]));
-qed "prat_add_left_le2_mono1";
-
-Goal "!!(q1::prat). q1 <= q2  ==> q1 + x <= q2 + x";
-by (auto_tac (claset() addDs [prat_add_left_le2_mono1],
-	      simpset() addsimps [prat_add_commute]));
-qed "prat_add_le2_mono1";
-
-Goal "!!k l::prat. [|i<=j;  k<=l |] ==> i + k <= j + l";
-by (etac (prat_add_le2_mono1 RS prat_le_trans) 1);
-by (simp_tac (simpset() addsimps [prat_add_commute]) 1);
-(*j moves to the end because it is free while k, l are bound*)
-by (etac prat_add_le2_mono1 1);
-qed "prat_add_le_mono";
-
-Goal "!!(x::prat). x + y < z + y ==> x < z";
-by (rtac ccontr 1);
-by (etac (prat_leI RS prat_le_imp_less_or_eq RS disjE) 1);
-by (dres_inst_tac [("x","y"),("q1.0","z")] prat_add_less2_mono1 1);
-by (auto_tac (claset() addIs [prat_less_asym],
-    simpset() addsimps [prat_less_not_refl]));
-qed "prat_add_right_less_cancel";
-
-Goal "!!(x::prat). y + x < y + z ==> x < z";
-by (res_inst_tac [("y","y")] prat_add_right_less_cancel 1);
-by (asm_full_simp_tac (simpset() addsimps [prat_add_commute]) 1);
-qed "prat_add_left_less_cancel";
-
-(*** lemmas required for lemma_gleason9_34 in PReal : w*y > y/z ***)
-Goalw [prat_of_pnat_def] 
-      "Abs_prat(ratrel``{(x,y)}) = (prat_of_pnat x)*qinv(prat_of_pnat y)";
-by (auto_tac (claset(),simpset() addsimps [prat_mult,qinv,pnat_mult_1_left,
-    pnat_mult_1]));
-qed "Abs_prat_mult_qinv";
-
-Goal "Abs_prat(ratrel``{(x,y)}) <= Abs_prat(ratrel``{(x,Abs_pnat (Suc 0))})";
-by (simp_tac (simpset() addsimps [Abs_prat_mult_qinv]) 1);
-by (rtac prat_mult_left_le2_mono1 1);
-by (rtac qinv_prat_le 1);
-by (pnat_ind_tac "y" 1);
-by (dres_inst_tac [("x","prat_of_pnat (Abs_pnat (Suc 0))")] prat_add_le2_mono1 2);
-by (cut_facts_tac [prat_less_1_2 RS prat_less_imp_le] 2);
-by (auto_tac (claset() addIs [prat_le_trans],
-    simpset() addsimps [prat_le_refl,
-    pSuc_is_plus_one,pnat_one_def,prat_of_pnat_add]));
-qed "lemma_Abs_prat_le1";
-
-Goal "Abs_prat(ratrel``{(x,Abs_pnat (Suc 0))}) <= Abs_prat(ratrel``{(x*y,Abs_pnat (Suc 0))})";
-by (simp_tac (simpset() addsimps [Abs_prat_mult_qinv]) 1);
-by (rtac prat_mult_le2_mono1 1);
-by (pnat_ind_tac "y" 1);
-by (dres_inst_tac [("x","prat_of_pnat x")] prat_add_le2_mono1 2);
-by (cut_inst_tac [("z","prat_of_pnat x")] (prat_self_less_add_self 
-    RS prat_less_imp_le) 2);
-by (auto_tac (claset() addIs [prat_le_trans],
-    simpset() addsimps [prat_le_refl,
-			pSuc_is_plus_one,pnat_one_def,prat_add_mult_distrib2,
-			prat_of_pnat_add,prat_of_pnat_mult]));
-qed "lemma_Abs_prat_le2";
-
-Goal "Abs_prat(ratrel``{(x,z)}) <= Abs_prat(ratrel``{(x*y,Abs_pnat (Suc 0))})";
-by (fast_tac (claset() addIs [prat_le_trans,
-			      lemma_Abs_prat_le1,lemma_Abs_prat_le2]) 1);
-qed "lemma_Abs_prat_le3";
-
-Goal "Abs_prat(ratrel``{(x*y,Abs_pnat (Suc 0))}) * Abs_prat(ratrel``{(w,x)}) = \
-\         Abs_prat(ratrel``{(w*y,Abs_pnat (Suc 0))})";
-by (full_simp_tac (simpset() addsimps [prat_mult,
-    pnat_mult_1,pnat_mult_1_left] @ pnat_mult_ac) 1);
-qed "pre_lemma_gleason9_34";
-
-Goal "Abs_prat(ratrel``{(y*x,Abs_pnat (Suc 0)*y)}) = \
-\         Abs_prat(ratrel``{(x,Abs_pnat (Suc 0))})";
-by (auto_tac (claset(),
-	      simpset() addsimps [pnat_mult_1,pnat_mult_1_left] @ pnat_mult_ac));
-qed "pre_lemma_gleason9_34b";
-
-Goal "(prat_of_pnat n < prat_of_pnat m) = (n < m)";
-by (auto_tac (claset(),simpset() addsimps [prat_less_def,
-    pnat_less_iff,prat_of_pnat_add]));
-by (res_inst_tac [("z","T")] eq_Abs_prat 1);
-by (auto_tac (claset() addDs [pnat_eq_lessI],
-    simpset() addsimps [prat_add,pnat_mult_1,
-    pnat_mult_1_left,prat_of_pnat_def,pnat_less_iff RS sym]));
-qed "prat_of_pnat_less_iff";
-
-Addsimps [prat_of_pnat_less_iff];
-
-(*------------------------------------------------------------------*)
-
-(*** prove witness that will be required to prove non-emptiness ***)
-(*** of preal type as defined using Dedekind Sections in PReal  ***)
-(*** Show that exists positive real `one' ***)
-
-Goal "EX q. q: {x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))}";
-by (fast_tac (claset() addIs [prat_less_qinv_2_1]) 1);
-qed "lemma_prat_less_1_memEx";
-
-Goal "{x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))} ~= {}";
-by (rtac notI 1);
-by (cut_facts_tac [lemma_prat_less_1_memEx] 1);
-by (Asm_full_simp_tac 1);
-qed "lemma_prat_less_1_set_non_empty";
-
-Goalw [psubset_def] "{} < {x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))}";
-by (asm_full_simp_tac (simpset() addsimps 
-         [lemma_prat_less_1_set_non_empty RS not_sym]) 1);
-qed "empty_set_psubset_lemma_prat_less_1_set";
-
-(*** exists rational not in set --- prat_of_pnat (Abs_pnat 1) itself ***)
-Goal "EX q. q ~: {x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))}";
-by (res_inst_tac [("x","prat_of_pnat (Abs_pnat (Suc 0))")] exI 1);
-by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl]));
-qed "lemma_prat_less_1_not_memEx";
-
-Goal "{x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))} ~= UNIV";
-by (rtac notI 1);
-by (cut_facts_tac [lemma_prat_less_1_not_memEx] 1);
-by (Asm_full_simp_tac 1);
-qed "lemma_prat_less_1_set_not_rat_set";
-
-Goalw [psubset_def,subset_def] 
-      "{x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))} < UNIV";
-by (asm_full_simp_tac
-    (simpset() addsimps [lemma_prat_less_1_set_not_rat_set,
-			 lemma_prat_less_1_not_memEx]) 1);
-qed "lemma_prat_less_1_set_psubset_rat_set";
-
-(*** prove non_emptiness of type ***)
-Goal "{x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))} : {A. {} < A & \
-\               A < UNIV & \
-\               (!y: A. ((!z. z < y --> z: A) & \
-\               (EX u: A. y < u)))}";
-by (auto_tac (claset() addDs [prat_less_trans],
-    simpset() addsimps [empty_set_psubset_lemma_prat_less_1_set,
-                       lemma_prat_less_1_set_psubset_rat_set]));
-by (dtac prat_dense 1);
-by (Fast_tac 1);
-qed "preal_1";
--- a/src/HOL/Real/PRat.thy	Tue Jan 27 09:44:14 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,50 +0,0 @@
-(*  Title       : PRat.thy
-    ID          : $Id$
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : The positive rationals
-*) 
-
-PRat = PNat +
-
-constdefs
-    ratrel   ::  "((pnat * pnat) * (pnat * pnat)) set"
-    "ratrel  ==  {p. ? x1 y1 x2 y2. p=((x1::pnat,y1),(x2,y2)) & x1*y2 = x2*y1}" 
-
-typedef prat = "UNIV//ratrel"          (quotient_def)
-
-instance
-   prat  :: {ord,plus,times}
-
-
-constdefs
-
-  prat_of_pnat :: pnat => prat           
-  "prat_of_pnat m == Abs_prat(ratrel``{(m,Abs_pnat 1)})"
-
-  qinv      :: prat => prat
-  "qinv(Q)  == Abs_prat(UN (x,y):Rep_prat(Q). ratrel``{(y,x)})" 
-
-defs
-
-  prat_add_def  
-  "P + Q == Abs_prat(UN (x1,y1):Rep_prat(P). UN (x2,y2):Rep_prat(Q).
-		     ratrel``{(x1*y2 + x2*y1, y1*y2)})"
-
-  prat_mult_def  
-  "P * Q == Abs_prat(UN (x1,y1):Rep_prat(P). UN (x2,y2):Rep_prat(Q).
-		     ratrel``{(x1*x2, y1*y2)})"
- 
-  (*** Gleason p. 119 ***)
-  prat_less_def
-  "P < (Q::prat) == EX T. P + T = Q"
-
-  prat_le_def
-  "P <= (Q::prat) == ~(Q < P)" 
-
-end
-  
-
-
-
-
--- a/src/HOL/Real/PReal.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Real/PReal.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -7,44 +7,95 @@
                   provides some of the definitions.
 *)
 
-theory PReal = PRat:
+theory PReal = RatArith:
+
+text{*Could be generalized and moved to @{text Ring_and_Field}*}
+lemma add_eq_exists: "\<exists>x. a+x = (b::rat)"
+by (rule_tac x="b-a" in exI, simp)
 
-typedef preal = "{A::prat set. {} < A & A < UNIV &
-                               (\<forall>y \<in> A. ((\<forall>z. z < y --> z \<in> A) &
-                                        (\<exists>u \<in> A. y < u)))}"
-apply (rule exI) 
-apply (rule preal_1) 
-done
+text{*As a special case, the sum of two positives is positive.  One of the
+premises could be weakened to the relation @{text "\<le>"}.*}
+lemma pos_add_strict: "[|0<a; b<c|] ==> b < a + (c::'a::ordered_semiring)"
+by (insert add_strict_mono [of 0 a b c], simp)
 
-
-instance preal :: ord ..
-instance preal :: plus ..
-instance preal :: times ..
+lemma interval_empty_iff:
+     "({y::'a::ordered_field. x < y & y < z} = {}) = (~(x < z))"
+by (blast dest: dense intro: order_less_trans)
 
 
 constdefs
-  preal_of_prat :: "prat => preal"
-   "preal_of_prat q     == Abs_preal({x::prat. x < q})"
+  cut :: "rat set => bool"
+    "cut A == {} \<subset> A &
+              A < {r. 0 < r} &
+              (\<forall>y \<in> A. ((\<forall>z. 0<z & z < y --> z \<in> A) & (\<exists>u \<in> A. y < u)))"
+
 
-  pinv       :: "preal => preal"
-  "pinv(R)   == Abs_preal({w. \<exists>y. w < y & qinv y \<notin> Rep_preal(R)})"
+lemma cut_of_rat: 
+  assumes q: "0 < q" shows "cut {r::rat. 0 < r & r < q}"
+proof -
+  let ?A = "{r::rat. 0 < r & r < q}"
+  from q have pos: "?A < {r. 0 < r}" by force
+  have nonempty: "{} \<subset> ?A"
+  proof
+    show "{} \<subseteq> ?A" by simp
+    show "{} \<noteq> ?A"
+      by (force simp only: q eq_commute [of "{}"] interval_empty_iff)
+  qed
+  show ?thesis
+    by (simp add: cut_def pos nonempty,
+        blast dest: dense intro: order_less_trans)
+qed
+
+
+typedef preal = "{A. cut A}"
+  by (blast intro: cut_of_rat [OF zero_less_one])
+
+instance preal :: ord ..
+instance preal :: plus ..
+instance preal :: minus ..
+instance preal :: times ..
+instance preal :: inverse ..
+
+
+constdefs
+  preal_of_rat :: "rat => preal"
+     "preal_of_rat q == Abs_preal({x::rat. 0 < x & x < q})"
 
   psup       :: "preal set => preal"
-  "psup(P)   == Abs_preal({w. \<exists>X \<in> P. w \<in> Rep_preal(X)})"
+    "psup(P)   == Abs_preal(\<Union>X \<in> P. Rep_preal(X))"
+
+  add_set :: "[rat set,rat set] => rat set"
+    "add_set A B == {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x + y}"
+
+  diff_set :: "[rat set,rat set] => rat set"
+    "diff_set A B == {w. \<exists>x. 0 < w & 0 < x & x \<notin> B & x + w \<in> A}"
+
+  mult_set :: "[rat set,rat set] => rat set"
+    "mult_set A B == {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x * y}"
+
+  inverse_set :: "rat set => rat set"
+    "inverse_set A == {x. \<exists>y. 0 < x & x < y & inverse y \<notin> A}"
+
 
 defs (overloaded)
 
+  preal_less_def:
+    "R < (S::preal) == Rep_preal R < Rep_preal S"
+
+  preal_le_def:
+    "R \<le> (S::preal) == Rep_preal R \<subseteq> Rep_preal S"
+
   preal_add_def:
-    "R + S == Abs_preal({w. \<exists>x \<in> Rep_preal(R). \<exists>y \<in> Rep_preal(S). w = x + y})"
+    "R + S == Abs_preal (add_set (Rep_preal R) (Rep_preal S))"
+
+  preal_diff_def:
+    "R - S == Abs_preal (diff_set (Rep_preal R) (Rep_preal S))"
 
   preal_mult_def:
-    "R * S == Abs_preal({w. \<exists>x \<in> Rep_preal(R). \<exists>y \<in> Rep_preal(S). w = x * y})"
+    "R * S == Abs_preal(mult_set (Rep_preal R) (Rep_preal S))"
 
-  preal_less_def:
-    "R < (S::preal) == Rep_preal(R) < Rep_preal(S)"
-
-  preal_le_def:
-    "R \<le> (S::preal) == Rep_preal(R) \<subseteq> Rep_preal(S)"
+  preal_inverse_def:
+    "inverse R == Abs_preal(inverse_set (Rep_preal R))"
 
 
 lemma inj_on_Abs_preal: "inj_on Abs_preal preal"
@@ -59,108 +110,61 @@
 apply (rule Rep_preal_inverse)
 done
 
-lemma empty_not_mem_preal [simp]: "{} \<notin> preal"
-by (unfold preal_def, fast)
+lemma preal_nonempty: "A \<in> preal ==> \<exists>x\<in>A. 0 < x"
+by (unfold preal_def cut_def, blast)
 
-lemma one_set_mem_preal: "{x::prat. x < prat_of_pnat (Abs_pnat (Suc 0))} \<in> preal"
-apply (unfold preal_def)
-apply (rule preal_1)
-done
+lemma preal_imp_psubset_positives: "A \<in> preal ==> A < {r. 0 < r}"
+by (force simp add: preal_def cut_def)
 
-declare one_set_mem_preal [simp]
+lemma preal_exists_bound: "A \<in> preal ==> \<exists>x. 0 < x & x \<notin> A"
+by (drule preal_imp_psubset_positives, auto)
 
-lemma preal_psubset_empty: "x \<in> preal ==> {} < x"
-by (unfold preal_def, fast)
-
-lemma Rep_preal_psubset_empty: "{} < Rep_preal x"
-by (rule Rep_preal [THEN preal_psubset_empty])
+lemma preal_exists_greater: "[| A \<in> preal; y \<in> A |] ==> \<exists>u \<in> A. y < u"
+by (unfold preal_def cut_def, blast)
 
 lemma mem_Rep_preal_Ex: "\<exists>x. x \<in> Rep_preal X"
-apply (cut_tac x = X in Rep_preal_psubset_empty)
-apply (auto intro: equals0I [symmetric] simp add: psubset_def)
-done
-
-lemma prealI1:
-      "[| {} < A; A < UNIV;
-               (\<forall>y \<in> A. ((\<forall>z. z < y --> z \<in> A) &
-                         (\<exists>u \<in> A. y < u))) |] ==> A \<in> preal"
-apply (unfold preal_def, fast)
+apply (insert Rep_preal [of X])
+apply (unfold preal_def cut_def, blast)
 done
 
-lemma prealI2:
-      "[| {} < A; A < UNIV;
-               \<forall>y \<in> A. (\<forall>z. z < y --> z \<in> A);
-               \<forall>y \<in> A. (\<exists>u \<in> A. y < u) |] ==> A \<in> preal"
-
-apply (unfold preal_def, best)
-done
-
-lemma prealE_lemma:
-      "A \<in> preal ==> {} < A & A < UNIV &
-                          (\<forall>y \<in> A. ((\<forall>z. z < y --> z \<in> A) &
-                                   (\<exists>u \<in> A. y < u)))"
-apply (unfold preal_def, fast)
-done
-
-declare prealI1 [intro!] prealI2 [intro!]
-
 declare Abs_preal_inverse [simp]
 
-
-lemma prealE_lemma1: "A \<in> preal ==> {} < A"
-by (unfold preal_def, fast)
-
-lemma prealE_lemma2: "A \<in> preal ==> A < UNIV"
-by (unfold preal_def, fast)
-
-lemma prealE_lemma3: "A \<in> preal ==> \<forall>y \<in> A. (\<forall>z. z < y --> z \<in> A)"
-by (unfold preal_def, fast)
-
-lemma prealE_lemma3a: "[| A \<in> preal; y \<in> A |] ==> (\<forall>z. z < y --> z \<in> A)"
-by (fast dest!: prealE_lemma3)
+lemma preal_downwards_closed: "[| A \<in> preal; y \<in> A; 0 < z; z < y |] ==> z \<in> A"
+by (unfold preal_def cut_def, blast)
 
-lemma prealE_lemma3b: "[| A \<in> preal; y \<in> A; z < y |] ==> z \<in> A"
-by (fast dest!: prealE_lemma3a)
-
-lemma prealE_lemma4: "A \<in> preal ==> \<forall>y \<in> A. (\<exists>u \<in> A. y < u)"
-by (unfold preal_def, fast)
+text{*Relaxing the final premise*}
+lemma preal_downwards_closed':
+     "[| A \<in> preal; y \<in> A; 0 < z; z \<le> y |] ==> z \<in> A"
+apply (simp add: order_le_less)
+apply (blast intro: preal_downwards_closed)
+done
 
-lemma prealE_lemma4a: "[| A \<in> preal; y \<in> A |] ==> \<exists>u \<in> A. y < u"
-by (fast dest!: prealE_lemma4)
-
-lemma not_mem_Rep_preal_Ex: "\<exists>x. x\<notin> Rep_preal X"
+lemma Rep_preal_exists_bound: "\<exists>x. 0 < x & x \<notin> Rep_preal X"
 apply (cut_tac x = X in Rep_preal)
-apply (drule prealE_lemma2)
+apply (drule preal_imp_psubset_positives)
 apply (auto simp add: psubset_def)
 done
 
 
 subsection{*@{term preal_of_prat}: the Injection from prat to preal*}
 
-text{*A few lemmas*}
-
-lemma lemma_prat_less_set_mem_preal: "{u::prat. u < y} \<in> preal"
-apply (cut_tac qless_Ex)
-apply (auto intro: prat_less_trans elim!: prat_less_irrefl)
-apply (blast dest: prat_dense)
+lemma rat_less_set_mem_preal: "0 < y ==> {u::rat. 0 < u & u < y} \<in> preal"
+apply (auto simp add: preal_def cut_def intro: order_less_trans)
+apply (force simp only: eq_commute [of "{}"] interval_empty_iff)
+apply (blast dest: dense intro: order_less_trans)
 done
 
-lemma lemma_prat_set_eq: "{u::prat. u < x} = {x. x < y} ==> x = y"
-apply (insert prat_linear [of x y], safe)
-apply (drule_tac [2] prat_dense, erule_tac [2] exE)
-apply (drule prat_dense, erule exE)
-apply (blast dest: prat_less_not_sym)
-apply (blast dest: prat_less_not_sym)
+lemma rat_subset_imp_le:
+     "[|{u::rat. 0 < u & u < x} \<subseteq> {u. 0 < u & u < y}; 0<x|] ==> x \<le> y"
+apply (simp add: linorder_not_less [symmetric])
+apply (blast dest: dense intro: order_less_trans)
 done
 
-lemma inj_preal_of_prat: "inj(preal_of_prat)"
-apply (rule inj_onI)
-apply (unfold preal_of_prat_def)
-apply (drule inj_on_Abs_preal [THEN inj_onD])
-apply (rule lemma_prat_less_set_mem_preal)
-apply (rule lemma_prat_less_set_mem_preal)
-apply (erule lemma_prat_set_eq)
-done
+lemma rat_set_eq_imp_eq:
+     "[|{u::rat. 0 < u & u < x} = {u. 0 < u & u < y};
+        0 < x; 0 < y|] ==> x = y"
+by (blast intro: rat_subset_imp_le order_antisym)
+
 
 
 subsection{*Theorems for Ordering*}
@@ -168,127 +172,173 @@
 text{*A positive fraction not in a positive real is an upper bound.
  Gleason p. 122 - Remark (1)*}
 
-lemma not_in_preal_ub: "x \<notin> Rep_preal(R) ==> \<forall>y \<in> Rep_preal(R). y < x"
-apply (cut_tac x1 = R in Rep_preal [THEN prealE_lemma]) 
-apply (blast intro: not_less_not_eq_prat_less)
-done
+lemma not_in_preal_ub:
+     assumes A: "A \<in> preal"
+         and notx: "x \<notin> A"
+         and y: "y \<in> A"
+         and pos: "0 < x"
+        shows "y < x"
+proof (cases rule: linorder_cases)
+  assume "x<y"
+  with notx show ?thesis
+    by (simp add:  preal_downwards_closed [OF A y] pos)
+next
+  assume "x=y"
+  with notx and y show ?thesis by simp
+next
+  assume "y<x"
+  thus ?thesis by assumption
+qed
+
+lemmas not_in_Rep_preal_ub = not_in_preal_ub [OF Rep_preal]
 
 
-text{*@{text preal_less} is a strict order: nonreflexive and transitive *}
+subsection{*The @{text "\<le>"} Ordering*}
+
+lemma preal_le_refl: "w \<le> (w::preal)"
+by (simp add: preal_le_def)
 
-lemma preal_less_not_refl: "~ (x::preal) < x"
-apply (unfold preal_less_def)
-apply (simp (no_asm) add: psubset_def)
+lemma preal_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::preal)"
+by (force simp add: preal_le_def)
+
+lemma preal_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::preal)"
+apply (simp add: preal_le_def)
+apply (rule Rep_preal_inject [THEN iffD1], blast)
 done
 
-lemmas preal_less_irrefl = preal_less_not_refl [THEN notE, standard]
+(* Axiom 'order_less_le' of class 'order': *)
+lemma preal_less_le: "((w::preal) < z) = (w \<le> z & w \<noteq> z)"
+by (simp add: preal_le_def preal_less_def Rep_preal_inject psubset_def)
+
+instance preal :: order
+proof qed
+ (assumption |
+  rule preal_le_refl preal_le_trans preal_le_anti_sym preal_less_le)+
 
-lemma preal_not_refl2: "!!(x::preal). x < y ==> x \<noteq> y"
-by (auto simp add: preal_less_not_refl)
+lemma preal_imp_pos: "[|A \<in> preal; r \<in> A|] ==> 0 < r"
+by (insert preal_imp_psubset_positives, blast)
 
-lemma preal_less_trans: "!!(x::preal). [| x < y; y < z |] ==> x < z"
-apply (unfold preal_less_def)
-apply (auto dest: subsetD equalityI simp add: psubset_def)
+lemma preal_le_linear: "x <= y | y <= (x::preal)"
+apply (auto simp add: preal_le_def)
+apply (rule ccontr)
+apply (blast dest: not_in_Rep_preal_ub intro: preal_imp_pos [OF Rep_preal]
+             elim: order_less_asym)
 done
 
-lemma preal_less_not_sym: "!! (q1::preal). q1 < q2 ==> ~ q2 < q1"
-apply (rule notI)
-apply (drule preal_less_trans, assumption)
-apply (simp add: preal_less_not_refl)
-done
+instance preal :: linorder
+  by (intro_classes, rule preal_le_linear)
 
-(* [| x < y;  ~P ==> y < x |] ==> P *)
-lemmas preal_less_asym = preal_less_not_sym [THEN contrapos_np, standard]
-
-lemma preal_linear:
-      "(x::preal) < y | x = y | y < x"
-apply (unfold preal_less_def)
-apply (auto dest!: inj_Rep_preal [THEN injD] simp add: psubset_def)
-apply (rule prealE_lemma3b, rule Rep_preal, assumption)
-apply (fast dest: not_in_preal_ub)
-done
 
 
 subsection{*Properties of Addition*}
 
 lemma preal_add_commute: "(x::preal) + y = y + x"
-apply (unfold preal_add_def)
+apply (unfold preal_add_def add_set_def)
 apply (rule_tac f = Abs_preal in arg_cong)
-apply (blast intro: prat_add_commute [THEN subst])
-done
-
-text{*Addition of two positive reals gives a positive real*}
-
-text{*Lemmas for proving positive reals addition set in @{typ preal}*}
-
-text{*Part 1 of Dedekind sections definition*}
-lemma preal_add_set_not_empty:
-     "{} < {w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x + y}"
-apply (cut_tac mem_Rep_preal_Ex mem_Rep_preal_Ex)
-apply (auto intro!: psubsetI)
+apply (force simp add: add_commute)
 done
 
-text{*Part 2 of Dedekind sections definition*}
-lemma preal_not_mem_add_set_Ex:
-     "\<exists>q. q  \<notin> {w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x + y}"
-apply (cut_tac X = R in not_mem_Rep_preal_Ex)
-apply (cut_tac X = S in not_mem_Rep_preal_Ex, clarify) 
-apply (drule not_in_preal_ub)+
-apply (rule_tac x = "x+xa" in exI)
-apply (auto dest!: bspec) 
-apply (drule prat_add_less_mono)
-apply (auto simp add: prat_less_not_refl)
+text{*Lemmas for proving that addition of two positive reals gives
+ a positive real*}
+
+lemma empty_psubset_nonempty: "a \<in> A ==> {} \<subset> A"
+by blast
+
+text{*Part 1 of Dedekind sections definition*}
+lemma add_set_not_empty:
+     "[|A \<in> preal; B \<in> preal|] ==> {} \<subset> add_set A B"
+apply (insert preal_nonempty [of A] preal_nonempty [of B]) 
+apply (auto simp add: add_set_def)
 done
 
-lemma preal_add_set_not_prat_set:
-     "{w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x + y} < UNIV"
-apply (auto intro!: psubsetI)
-apply (cut_tac R = R and S = S in preal_not_mem_add_set_Ex, auto)
+text{*Part 2 of Dedekind sections definition.  A structured version of
+this proof is @{text preal_not_mem_mult_set_Ex} below.*}
+lemma preal_not_mem_add_set_Ex:
+     "[|A \<in> preal; B \<in> preal|] ==> \<exists>q. 0 < q & q \<notin> add_set A B"
+apply (insert preal_exists_bound [of A] preal_exists_bound [of B], auto) 
+apply (rule_tac x = "x+xa" in exI)
+apply (simp add: add_set_def, clarify)
+apply (drule not_in_preal_ub, assumption+)+
+apply (force dest: add_strict_mono)
 done
 
+lemma add_set_not_rat_set:
+   assumes A: "A \<in> preal" 
+       and B: "B \<in> preal"
+     shows "add_set A B < {r. 0 < r}"
+proof
+  from preal_imp_pos [OF A] preal_imp_pos [OF B]
+  show "add_set A B \<subseteq> {r. 0 < r}" by (force simp add: add_set_def) 
+next
+  show "add_set A B \<noteq> {r. 0 < r}"
+    by (insert preal_not_mem_add_set_Ex [OF A B], blast) 
+qed
+
 text{*Part 3 of Dedekind sections definition*}
-lemma preal_add_set_lemma3:
-     "\<forall>y \<in> {w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x + y}.
-         \<forall>z. z < y --> z \<in> {w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x+y}"
-apply auto
-apply (frule prat_mult_qinv_less_1)
-apply (frule_tac x = x 
-       in prat_mult_less2_mono1 [of _ "prat_of_pnat (Abs_pnat (Suc 0))"])
-apply (frule_tac x = ya 
-       in prat_mult_less2_mono1 [of _ "prat_of_pnat (Abs_pnat (Suc 0))"])
-apply simp
-apply (drule Rep_preal [THEN prealE_lemma3a])+
-apply (erule allE)+
-apply auto
-apply (rule bexI)+
-apply (auto simp add: prat_add_mult_distrib2 [symmetric] 
-      prat_add_assoc [symmetric] prat_mult_assoc)
+lemma add_set_lemma3:
+     "[|A \<in> preal; B \<in> preal; u \<in> add_set A B; 0 < z; z < u|] 
+      ==> z \<in> add_set A B"
+proof (unfold add_set_def, clarify)
+  fix x::rat and y::rat
+  assume A: "A \<in> preal" 
+     and B: "B \<in> preal"
+     and [simp]: "0 < z"
+     and zless: "z < x + y"
+     and x:  "x \<in> A"
+     and y:  "y \<in> B"
+  have xpos [simp]: "0<x" by (rule preal_imp_pos [OF A x])
+  have ypos [simp]: "0<y" by (rule preal_imp_pos [OF B y])
+  have xypos [simp]: "0 < x+y" by (simp add: pos_add_strict)
+  let ?f = "z/(x+y)"
+  have fless: "?f < 1" by (simp add: zless pos_divide_less_eq)
+  show "\<exists>x' \<in> A. \<exists>y'\<in>B. z = x' + y'"
+  proof
+    show "\<exists>y' \<in> B. z = x*?f + y'"
+    proof
+      show "z = x*?f + y*?f"
+	by (simp add: left_distrib [symmetric] divide_inverse_zero mult_ac
+		      order_less_imp_not_eq2)
+    next
+      show "y * ?f \<in> B"
+      proof (rule preal_downwards_closed [OF B y])
+        show "0 < y * ?f"
+          by (simp add: divide_inverse_zero zero_less_mult_iff)
+      next
+        show "y * ?f < y"
+          by (insert mult_strict_left_mono [OF fless ypos], simp)
+      qed
+    qed
+  next
+    show "x * ?f \<in> A"
+    proof (rule preal_downwards_closed [OF A x])
+      show "0 < x * ?f"
+	by (simp add: divide_inverse_zero zero_less_mult_iff)
+    next
+      show "x * ?f < x"
+	by (insert mult_strict_left_mono [OF fless xpos], simp)
+    qed
+  qed
+qed
+
+text{*Part 4 of Dedekind sections definition*}
+lemma add_set_lemma4:
+     "[|A \<in> preal; B \<in> preal; y \<in> add_set A B|] ==> \<exists>u \<in> add_set A B. y < u"
+apply (auto simp add: add_set_def)
+apply (frule preal_exists_greater [of A], auto) 
+apply (rule_tac x="u + y" in exI)
+apply (auto intro: add_strict_left_mono)
 done
 
-lemma preal_add_set_lemma4:
-     "\<forall>y \<in> {w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x + y}.
-          \<exists>u \<in> {w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x + y}. y < u"
-apply auto
-apply (drule Rep_preal [THEN prealE_lemma4a])
-apply (auto intro: prat_add_less2_mono1)
-done
-
-lemma preal_mem_add_set:
-     "{w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x + y} \<in> preal"
-apply (rule prealI2)
-apply (rule preal_add_set_not_empty)
-apply (rule preal_add_set_not_prat_set)
-apply (rule preal_add_set_lemma3)
-apply (rule preal_add_set_lemma4)
+lemma mem_add_set:
+     "[|A \<in> preal; B \<in> preal|] ==> add_set A B \<in> preal"
+apply (simp (no_asm_simp) add: preal_def cut_def)
+apply (blast intro!: add_set_not_empty add_set_not_rat_set
+                     add_set_lemma3 add_set_lemma4)
 done
 
 lemma preal_add_assoc: "((x::preal) + y) + z = x + (y + z)"
-apply (unfold preal_add_def)
-apply (rule_tac f = Abs_preal in arg_cong)
-apply (simp (no_asm) add: preal_mem_add_set [THEN Abs_preal_inverse])
-apply (auto simp add: prat_add_ac)
-apply (rule bexI)
-apply (auto intro!: exI simp add: prat_add_ac)
+apply (simp add: preal_add_def mem_add_set Rep_preal)
+apply (force simp add: add_set_def add_ac)
 done
 
 lemma preal_add_left_commute: "x + (y + z) = y + ((x + z)::preal)"
@@ -297,7 +347,7 @@
   apply (rule preal_add_commute)
   done
 
-(* Positive Reals addition is an AC operator *)
+text{* Positive Real addition is an AC operator *}
 lemmas preal_add_ac = preal_add_assoc preal_add_commute preal_add_left_commute
 
 
@@ -306,9 +356,9 @@
 text{*Proofs essentially same as for addition*}
 
 lemma preal_mult_commute: "(x::preal) * y = y * x"
-apply (unfold preal_mult_def)
+apply (unfold preal_mult_def mult_set_def)
 apply (rule_tac f = Abs_preal in arg_cong)
-apply (blast intro: prat_mult_commute [THEN subst])
+apply (force simp add: mult_commute)
 done
 
 text{*Multiplication of two positive reals gives a positive real.}
@@ -316,68 +366,109 @@
 text{*Lemmas for proving positive reals multiplication set in @{typ preal}*}
 
 text{*Part 1 of Dedekind sections definition*}
-lemma preal_mult_set_not_empty:
-     "{} < {w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x * y}"
-apply (cut_tac mem_Rep_preal_Ex mem_Rep_preal_Ex)
-apply (auto intro!: psubsetI)
+lemma mult_set_not_empty:
+     "[|A \<in> preal; B \<in> preal|] ==> {} \<subset> mult_set A B"
+apply (insert preal_nonempty [of A] preal_nonempty [of B]) 
+apply (auto simp add: mult_set_def)
 done
 
 text{*Part 2 of Dedekind sections definition*}
 lemma preal_not_mem_mult_set_Ex:
-     "\<exists>q. q  \<notin> {w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x * y}"
-apply (cut_tac X = R in not_mem_Rep_preal_Ex)
-apply (cut_tac X = S in not_mem_Rep_preal_Ex)
-apply (erule exE)+
-apply (drule not_in_preal_ub)+
-apply (rule_tac x = "x*xa" in exI)
-apply (auto, (erule ballE)+, auto)
-apply (drule prat_mult_less_mono)
-apply (auto simp add: prat_less_not_refl)
-done
+   assumes A: "A \<in> preal" 
+       and B: "B \<in> preal"
+     shows "\<exists>q. 0 < q & q \<notin> mult_set A B"
+proof -
+  from preal_exists_bound [OF A]
+  obtain x where [simp]: "0 < x" "x \<notin> A" by blast
+  from preal_exists_bound [OF B]
+  obtain y where [simp]: "0 < y" "y \<notin> B" by blast
+  show ?thesis
+  proof (intro exI conjI)
+    show "0 < x*y" by (simp add: mult_pos)
+    show "x * y \<notin> mult_set A B"
+    proof (auto simp add: mult_set_def)
+      fix u::rat and v::rat
+      assume "u \<in> A" and "v \<in> B" and "x*y = u*v"
+      moreover
+      with prems have "u<x" and "v<y" by (blast dest: not_in_preal_ub)+
+      moreover
+      with prems have "0\<le>v"
+        by (blast intro: preal_imp_pos [OF B]  order_less_imp_le prems)
+      moreover
+      hence "u*v < x*y" by (blast intro: mult_strict_mono prems)
+      ultimately show False by force
+    qed
+  qed
+qed
 
-lemma preal_mult_set_not_prat_set:
-     "{w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x * y} < UNIV"
-apply (auto intro!: psubsetI)
-apply (cut_tac R = R and S = S in preal_not_mem_mult_set_Ex, auto)
-done
+lemma mult_set_not_rat_set:
+   assumes A: "A \<in> preal" 
+       and B: "B \<in> preal"
+     shows "mult_set A B < {r. 0 < r}"
+proof
+  show "mult_set A B \<subseteq> {r. 0 < r}"
+    by (force simp add: mult_set_def
+              intro: preal_imp_pos [OF A] preal_imp_pos [OF B] mult_pos)
+next
+  show "mult_set A B \<noteq> {r. 0 < r}"
+    by (insert preal_not_mem_mult_set_Ex [OF A B], blast)
+qed
+
+
 
 text{*Part 3 of Dedekind sections definition*}
-lemma preal_mult_set_lemma3:
-     "\<forall>y \<in> {w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x * y}.
-         \<forall>z. z < y --> z \<in> {w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x*y}"
-apply auto
-apply (frule_tac x = "qinv (ya)" in prat_mult_left_less2_mono1)
-apply (simp add: prat_mult_ac)
-apply (drule Rep_preal [THEN prealE_lemma3a])
-apply (erule allE)
-apply (rule bexI)+
-apply (auto simp add: prat_mult_assoc)
+lemma mult_set_lemma3:
+     "[|A \<in> preal; B \<in> preal; u \<in> mult_set A B; 0 < z; z < u|] 
+      ==> z \<in> mult_set A B"
+proof (unfold mult_set_def, clarify)
+  fix x::rat and y::rat
+  assume A: "A \<in> preal" 
+     and B: "B \<in> preal"
+     and [simp]: "0 < z"
+     and zless: "z < x * y"
+     and x:  "x \<in> A"
+     and y:  "y \<in> B"
+  have [simp]: "0<y" by (rule preal_imp_pos [OF B y])
+  show "\<exists>x' \<in> A. \<exists>y' \<in> B. z = x' * y'"
+  proof
+    show "\<exists>y'\<in>B. z = (z/y) * y'"
+    proof
+      show "z = (z/y)*y"
+	by (simp add: divide_inverse_zero mult_commute [of y] mult_assoc
+		      order_less_imp_not_eq2)
+      show "y \<in> B" .
+    qed
+  next
+    show "z/y \<in> A"
+    proof (rule preal_downwards_closed [OF A x])
+      show "0 < z/y"
+	by (simp add: zero_less_divide_iff)
+      show "z/y < x" by (simp add: pos_divide_less_eq zless)
+    qed
+  qed
+qed
+
+text{*Part 4 of Dedekind sections definition*}
+lemma mult_set_lemma4:
+     "[|A \<in> preal; B \<in> preal; y \<in> mult_set A B|] ==> \<exists>u \<in> mult_set A B. y < u"
+apply (auto simp add: mult_set_def)
+apply (frule preal_exists_greater [of A], auto) 
+apply (rule_tac x="u * y" in exI)
+apply (auto intro: preal_imp_pos [of A] preal_imp_pos [of B] 
+                   mult_strict_right_mono)
 done
 
-lemma preal_mult_set_lemma4:
-     "\<forall>y \<in> {w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x * y}.
-          \<exists>u \<in> {w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x * y}. y < u"
-apply auto
-apply (drule Rep_preal [THEN prealE_lemma4a])
-apply (auto intro: prat_mult_less2_mono1)
-done
 
-lemma preal_mem_mult_set:
-     "{w. \<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. w = x * y} \<in> preal"
-apply (rule prealI2)
-apply (rule preal_mult_set_not_empty)
-apply (rule preal_mult_set_not_prat_set)
-apply (rule preal_mult_set_lemma3)
-apply (rule preal_mult_set_lemma4)
+lemma mem_mult_set:
+     "[|A \<in> preal; B \<in> preal|] ==> mult_set A B \<in> preal"
+apply (simp (no_asm_simp) add: preal_def cut_def)
+apply (blast intro!: mult_set_not_empty mult_set_not_rat_set
+                     mult_set_lemma3 mult_set_lemma4)
 done
 
 lemma preal_mult_assoc: "((x::preal) * y) * z = x * (y * z)"
-apply (unfold preal_mult_def)
-apply (rule_tac f = Abs_preal in arg_cong)
-apply (simp (no_asm) add: preal_mem_mult_set [THEN Abs_preal_inverse])
-apply (auto simp add: prat_mult_ac)
-apply (rule bexI)
-apply (auto intro!: exI simp add: prat_mult_ac)
+apply (simp add: preal_mult_def mem_mult_set Rep_preal)
+apply (force simp add: mult_set_def mult_ac)
 done
 
 lemma preal_mult_left_commute: "x * (y * z) = y * ((x * z)::preal)"
@@ -386,32 +477,64 @@
   apply (rule preal_mult_commute)
   done
 
-(* Positive Reals multiplication is an AC operator *)
+
+text{* Positive Real multiplication is an AC operator *}
 lemmas preal_mult_ac =
        preal_mult_assoc preal_mult_commute preal_mult_left_commute
 
-(* Positive Real 1 is the multiplicative identity element *)
-(* long *)
-lemma preal_mult_1:
-      "(preal_of_prat (prat_of_pnat (Abs_pnat (Suc 0)))) * z = z"
-apply (unfold preal_of_prat_def preal_mult_def)
-apply (rule Rep_preal_inverse [THEN subst])
-apply (rule_tac f = Abs_preal in arg_cong)
-apply (rule one_set_mem_preal [THEN Abs_preal_inverse, THEN ssubst])
-apply (auto simp add: Rep_preal_inverse)
-apply (drule Rep_preal [THEN prealE_lemma4a]) 
-apply (erule bexE) 
-apply (drule prat_mult_less_mono)
-apply (auto dest: Rep_preal [THEN prealE_lemma3a])
-apply (frule Rep_preal [THEN prealE_lemma4a]) 
-apply (erule bexE) 
-apply (frule_tac x = "qinv (u)" in prat_mult_less2_mono1)
-apply (rule exI, auto, rule_tac x = u in bexI)
-apply (auto simp add: prat_mult_assoc)
-done
+
+text{* Positive real 1 is the multiplicative identity element *}
+
+lemma rat_mem_preal: "0 < q ==> {r::rat. 0 < r & r < q} \<in> preal"
+by (simp add: preal_def cut_of_rat)
 
-lemma preal_mult_1_right:
-     "z * (preal_of_prat (prat_of_pnat (Abs_pnat (Suc 0)))) = z"
+lemma preal_mult_1: "(preal_of_rat 1) * z = z"
+proof (induct z)
+  fix A :: "rat set"
+  assume A: "A \<in> preal"
+  have "{w. \<exists>u. 0 < u \<and> u < 1 & (\<exists>v \<in> A. w = u * v)} = A" (is "?lhs = A")
+  proof
+    show "?lhs \<subseteq> A"
+    proof clarify
+      fix x::rat and u::rat and v::rat
+      assume upos: "0<u" and "u<1" and v: "v \<in> A"
+      have vpos: "0<v" by (rule preal_imp_pos [OF A v])
+      hence "u*v < 1*v" by (simp only: mult_strict_right_mono prems)
+      thus "u * v \<in> A"
+        by (force intro: preal_downwards_closed [OF A v] mult_pos upos vpos)
+    qed
+  next
+    show "A \<subseteq> ?lhs"
+    proof clarify
+      fix x::rat
+      assume x: "x \<in> A"
+      have xpos: "0<x" by (rule preal_imp_pos [OF A x])
+      from preal_exists_greater [OF A x]
+      obtain v where v: "v \<in> A" and xlessv: "x < v" ..
+      have vpos: "0<v" by (rule preal_imp_pos [OF A v])
+      show "\<exists>u. 0 < u \<and> u < 1 \<and> (\<exists>v\<in>A. x = u * v)"
+      proof (intro exI conjI)
+        show "0 < x/v"
+          by (simp add: zero_less_divide_iff xpos vpos)
+	show "x / v < 1"
+          by (simp add: pos_divide_less_eq vpos xlessv)
+        show "\<exists>v'\<in>A. x = (x / v) * v'"
+        proof
+          show "x = (x/v)*v"
+	    by (simp add: divide_inverse_zero mult_assoc vpos
+                          order_less_imp_not_eq2)
+          show "v \<in> A" .
+        qed
+      qed
+    qed
+  qed
+  thus "preal_of_rat 1 * Abs_preal A = Abs_preal A"
+    by (simp add: preal_of_rat_def preal_mult_def mult_set_def 
+                  rat_mem_preal A)
+qed
+
+
+lemma preal_mult_1_right: "z * (preal_of_rat 1) = z"
 apply (rule preal_mult_commute [THEN subst])
 apply (rule preal_mult_1)
 done
@@ -419,884 +542,821 @@
 
 subsection{*Distribution of Multiplication across Addition*}
 
-lemma mem_Rep_preal_addD:
-      "z \<in> Rep_preal(R+S) ==>
-            \<exists>x \<in> Rep_preal(R). \<exists>y \<in> Rep_preal(S). z = x + y"
-apply (unfold preal_add_def)
-apply (drule preal_mem_add_set [THEN Abs_preal_inverse, THEN subst], fast)
-done
-
-lemma mem_Rep_preal_addI:
-      "\<exists>x \<in> Rep_preal(R). \<exists>y \<in> Rep_preal(S). z = x + y
-       ==> z \<in> Rep_preal(R+S)"
-apply (unfold preal_add_def)
-apply (rule preal_mem_add_set [THEN Abs_preal_inverse, THEN ssubst], fast)
-done
-
 lemma mem_Rep_preal_add_iff:
-     "(z \<in> Rep_preal(R+S)) = (\<exists>x \<in> Rep_preal(R).
-                                  \<exists>y \<in> Rep_preal(S). z = x + y)"
-apply (fast intro!: mem_Rep_preal_addD mem_Rep_preal_addI)
-done
-
-lemma mem_Rep_preal_multD:
-      "z \<in> Rep_preal(R*S) ==>
-            \<exists>x \<in> Rep_preal(R). \<exists>y \<in> Rep_preal(S). z = x * y"
-apply (unfold preal_mult_def)
-apply (drule preal_mem_mult_set [THEN Abs_preal_inverse, THEN subst], fast)
-done
-
-lemma mem_Rep_preal_multI:
-      "\<exists>x \<in> Rep_preal(R). \<exists>y \<in> Rep_preal(S). z = x * y
-       ==> z \<in> Rep_preal(R*S)"
-apply (unfold preal_mult_def)
-apply (rule preal_mem_mult_set [THEN Abs_preal_inverse, THEN ssubst], fast)
+      "(z \<in> Rep_preal(R+S)) = (\<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. z = x + y)"
+apply (simp add: preal_add_def mem_add_set Rep_preal)
+apply (simp add: add_set_def) 
 done
 
 lemma mem_Rep_preal_mult_iff:
-     "(z \<in> Rep_preal(R*S)) =
-      (\<exists>x \<in> Rep_preal(R). \<exists>y \<in> Rep_preal(S). z = x * y)"
-by (fast intro!: mem_Rep_preal_multD mem_Rep_preal_multI)
-
-lemma lemma_add_mult_mem_Rep_preal:
-     "[| xb \<in> Rep_preal z1; xc \<in> Rep_preal z2; ya:
-                   Rep_preal w; yb \<in> Rep_preal w |] ==>
-                   xb * ya + xc * yb \<in> Rep_preal (z1 * w + z2 * w)"
-by (fast intro: mem_Rep_preal_addI mem_Rep_preal_multI)
+      "(z \<in> Rep_preal(R*S)) = (\<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. z = x * y)"
+apply (simp add: preal_mult_def mem_mult_set Rep_preal)
+apply (simp add: mult_set_def) 
+done
 
-lemma lemma_add_mult_mem_Rep_preal1:
-     "[| xb \<in> Rep_preal z1; xc \<in> Rep_preal z2; ya:
-                   Rep_preal w; yb \<in> Rep_preal w |] ==>
-                   yb*(xb + xc) \<in> Rep_preal (w*(z1 + z2))"
-by (fast intro: mem_Rep_preal_addI mem_Rep_preal_multI)
-
-lemma lemma_preal_add_mult_distrib:
-     "x \<in> Rep_preal (w * z1 + w * z2) ==>
-               x \<in> Rep_preal (w * (z1 + z2))"
-apply (auto dest!: mem_Rep_preal_addD mem_Rep_preal_multD)
-apply (frule_tac ya = xa and yb = xb and xb = ya and xc = yb in lemma_add_mult_mem_Rep_preal1, auto)
-apply (rule_tac x = xa and y = xb in prat_linear_less2)
-apply (drule_tac b = ya and c = yb in lemma_prat_add_mult_mono)
-apply (rule Rep_preal [THEN prealE_lemma3b])
-apply (auto simp add: prat_add_mult_distrib2)
-apply (drule_tac ya = xb and yb = xa and xc = ya and xb = yb in lemma_add_mult_mem_Rep_preal1, auto)
-apply (drule_tac b = yb and c = ya in lemma_prat_add_mult_mono)
-apply (rule Rep_preal [THEN prealE_lemma3b])
-apply (erule_tac V = "xb * ya + xb * yb \<in> Rep_preal (w * (z1 + z2))" in thin_rl)
-apply (auto simp add: prat_add_mult_distrib prat_add_commute preal_add_ac)
+lemma distrib_subset1:
+     "Rep_preal (w * (x + y)) \<subseteq> Rep_preal (w * x + w * y)"
+apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff)
+apply (force simp add: right_distrib)
 done
 
-lemma lemma_preal_add_mult_distrib2:
-     "x \<in> Rep_preal (w * (z1 + z2)) ==>
-               x \<in> Rep_preal (w * z1 + w * z2)"
-by (auto dest!: mem_Rep_preal_addD mem_Rep_preal_multD
-         intro!: bexI mem_Rep_preal_addI mem_Rep_preal_multI 
-         simp add: prat_add_mult_distrib2)
+lemma linorder_le_cases [case_names le ge]:
+    "((x::'a::linorder) <= y ==> P) ==> (y <= x ==> P) ==> P"
+  apply (insert linorder_linear, blast)
+  done
 
-lemma preal_add_mult_distrib2: "(w * ((z1::preal) + z2)) = (w * z1) + (w * z2)"
-apply (rule inj_Rep_preal [THEN injD])
-apply (fast intro: lemma_preal_add_mult_distrib lemma_preal_add_mult_distrib2)
+lemma preal_add_mult_distrib_mean:
+  assumes a: "a \<in> Rep_preal w"
+      and b: "b \<in> Rep_preal w"
+      and d: "d \<in> Rep_preal x"
+      and e: "e \<in> Rep_preal y"
+     shows "\<exists>c \<in> Rep_preal w. a * d + b * e = c * (d + e)"
+proof
+  let ?c = "(a*d + b*e)/(d+e)"
+  have [simp]: "0<a" "0<b" "0<d" "0<e" "0<d+e"
+    by (blast intro: preal_imp_pos [OF Rep_preal] a b d e pos_add_strict)+
+  have cpos: "0 < ?c"
+    by (simp add: zero_less_divide_iff zero_less_mult_iff pos_add_strict)
+  show "a * d + b * e = ?c * (d + e)"
+    by (simp add: divide_inverse_zero mult_assoc order_less_imp_not_eq2)
+  show "?c \<in> Rep_preal w"
+    proof (cases rule: linorder_le_cases)
+      assume "a \<le> b"
+      hence "?c \<le> b"
+	by (simp add: pos_divide_le_eq right_distrib mult_right_mono
+                      order_less_imp_le)
+      thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal b cpos])
+    next
+      assume "b \<le> a"
+      hence "?c \<le> a"
+	by (simp add: pos_divide_le_eq right_distrib mult_right_mono
+                      order_less_imp_le)
+      thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal a cpos])
+    qed
+  qed
+
+lemma distrib_subset2:
+     "Rep_preal (w * x + w * y) \<subseteq> Rep_preal (w * (x + y))"
+apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff)
+apply (drule_tac w=w and x=x and y=y in preal_add_mult_distrib_mean, auto)
 done
 
-lemma preal_add_mult_distrib: "(((z1::preal) + z2) * w) = (z1 * w) + (z2 * w)"
-apply (simp (no_asm) add: preal_mult_commute preal_add_mult_distrib2)
+lemma preal_add_mult_distrib2: "(w * ((x::preal) + y)) = (w * x) + (w * y)"
+apply (rule inj_Rep_preal [THEN injD])
+apply (rule equalityI [OF distrib_subset1 distrib_subset2])
 done
 
+lemma preal_add_mult_distrib: "(((x::preal) + y) * w) = (x * w) + (y * w)"
+by (simp add: preal_mult_commute preal_add_mult_distrib2)
+
 
 subsection{*Existence of Inverse, a Positive Real*}
 
-lemma qinv_not_mem_Rep_preal_Ex: "\<exists>y. qinv(y) \<notin>  Rep_preal X"
-apply (cut_tac X = X in not_mem_Rep_preal_Ex)
-apply (erule exE, cut_tac x = x in prat_as_inverse_ex, auto)
-done
-
-lemma lemma_preal_mem_inv_set_ex:
-     "\<exists>q. q \<in> {x. \<exists>y. x < y & qinv y \<notin>  Rep_preal A}"
-apply (cut_tac X = A in qinv_not_mem_Rep_preal_Ex, auto)
-apply (cut_tac y = y in qless_Ex, fast)
-done
+lemma mem_inv_set_ex:
+  assumes A: "A \<in> preal" shows "\<exists>x y. 0 < x & x < y & inverse y \<notin> A"
+proof -
+  from preal_exists_bound [OF A]
+  obtain x where [simp]: "0<x" "x \<notin> A" by blast
+  show ?thesis
+  proof (intro exI conjI)
+    show "0 < inverse (x+1)"
+      by (simp add: order_less_trans [OF _ less_add_one]) 
+    show "inverse(x+1) < inverse x"
+      by (simp add: less_imp_inverse_less less_add_one)
+    show "inverse (inverse x) \<notin> A"
+      by (simp add: order_less_imp_not_eq2)
+  qed
+qed
 
 text{*Part 1 of Dedekind sections definition*}
-lemma preal_inv_set_not_empty: "{} < {x. \<exists>y. x < y & qinv y \<notin>  Rep_preal A}"
-apply (cut_tac lemma_preal_mem_inv_set_ex)
-apply (auto intro!: psubsetI)
+lemma inverse_set_not_empty:
+     "A \<in> preal ==> {} \<subset> inverse_set A"
+apply (insert mem_inv_set_ex [of A])
+apply (auto simp add: inverse_set_def)
 done
 
 text{*Part 2 of Dedekind sections definition*}
-lemma qinv_mem_Rep_preal_Ex: "\<exists>y. qinv(y) \<in>  Rep_preal X"
-apply (cut_tac X = X in mem_Rep_preal_Ex)
-apply (erule exE, cut_tac x = x in prat_as_inverse_ex, auto)
-done
 
-lemma preal_not_mem_inv_set_Ex:
-     "\<exists>x. x \<notin> {x. \<exists>y. x < y & qinv y \<notin>  Rep_preal A}"
-apply (rule ccontr)
-apply (cut_tac X = A in qinv_mem_Rep_preal_Ex, auto)
-apply (erule allE, clarify) 
-apply (drule qinv_prat_less, drule not_in_preal_ub)
-apply (erule_tac x = "qinv y" in ballE)
-apply (drule prat_less_trans)
-apply (auto simp add: prat_less_not_refl)
-done
+lemma preal_not_mem_inverse_set_Ex:
+   assumes A: "A \<in> preal"  shows "\<exists>q. 0 < q & q \<notin> inverse_set A"
+proof -
+  from preal_nonempty [OF A]
+  obtain x where x: "x \<in> A" and  xpos [simp]: "0<x" ..
+  show ?thesis
+  proof (intro exI conjI)
+    show "0 < inverse x" by simp
+    show "inverse x \<notin> inverse_set A"
+    proof (auto simp add: inverse_set_def)
+      fix y::rat 
+      assume ygt: "inverse x < y"
+      have [simp]: "0 < y" by (simp add: order_less_trans [OF _ ygt])
+      have iyless: "inverse y < x" 
+        by (simp add: inverse_less_imp_less [of x] ygt)
+      show "inverse y \<in> A"
+        by (simp add: preal_downwards_closed [OF A x] iyless) 
+    qed
+  qed
+qed
 
-lemma preal_inv_set_not_prat_set:
-     "{x. \<exists>y. x < y & qinv y \<notin>  Rep_preal A} < UNIV"
-apply (auto intro!: psubsetI)
-apply (cut_tac A = A in preal_not_mem_inv_set_Ex, auto)
-done
+lemma inverse_set_not_rat_set:
+   assumes A: "A \<in> preal"  shows "inverse_set A < {r. 0 < r}"
+proof
+  show "inverse_set A \<subseteq> {r. 0 < r}"  by (force simp add: inverse_set_def)
+next
+  show "inverse_set A \<noteq> {r. 0 < r}"
+    by (insert preal_not_mem_inverse_set_Ex [OF A], blast)
+qed
 
 text{*Part 3 of Dedekind sections definition*}
-lemma preal_inv_set_lemma3:
-     "\<forall>y \<in> {x. \<exists>y. x < y & qinv y \<notin> Rep_preal A}.
-        \<forall>z. z < y --> z \<in> {x. \<exists>y. x < y & qinv y \<notin> Rep_preal A}"
-apply auto
-apply (rule_tac x = ya in exI)
-apply (auto intro: prat_less_trans)
+lemma inverse_set_lemma3:
+     "[|A \<in> preal; u \<in> inverse_set A; 0 < z; z < u|] 
+      ==> z \<in> inverse_set A"
+apply (auto simp add: inverse_set_def)
+apply (auto intro: order_less_trans)
 done
 
-lemma preal_inv_set_lemma4:
-     "\<forall>y \<in> {x. \<exists>y. x < y & qinv y \<notin> Rep_preal A}.
-        Bex {x. \<exists>y. x < y & qinv y \<notin> Rep_preal A} (op < y)"
-by (blast dest: prat_dense)
-
-lemma preal_mem_inv_set: "{x. \<exists>y. x < y & qinv(y) \<notin> Rep_preal(A)} \<in> preal"
-apply (rule prealI2)
-apply (rule preal_inv_set_not_empty)
-apply (rule preal_inv_set_not_prat_set)
-apply (rule preal_inv_set_lemma3)
-apply (rule preal_inv_set_lemma4)
+text{*Part 4 of Dedekind sections definition*}
+lemma inverse_set_lemma4:
+     "[|A \<in> preal; y \<in> inverse_set A|] ==> \<exists>u \<in> inverse_set A. y < u"
+apply (auto simp add: inverse_set_def)
+apply (drule dense [of y]) 
+apply (blast intro: order_less_trans)
 done
 
-(*more lemmas for inverse *)
-lemma preal_mem_mult_invD:
-     "x \<in> Rep_preal(pinv(A)*A) ==>
-      x \<in> Rep_preal(preal_of_prat (prat_of_pnat (Abs_pnat (Suc 0))))"
-apply (auto dest!: mem_Rep_preal_multD simp add: pinv_def preal_of_prat_def)
-apply (drule preal_mem_inv_set [THEN Abs_preal_inverse, THEN subst])
-apply (auto dest!: not_in_preal_ub)
-apply (drule prat_mult_less_mono, blast, auto)
+
+lemma mem_inverse_set:
+     "A \<in> preal ==> inverse_set A \<in> preal"
+apply (simp (no_asm_simp) add: preal_def cut_def)
+apply (blast intro!: inverse_set_not_empty inverse_set_not_rat_set
+                     inverse_set_lemma3 inverse_set_lemma4)
 done
 
+
 subsection{*Gleason's Lemma 9-3.4, page 122*}
 
-lemma lemma1_gleason9_34:
-     "\<forall>xa \<in> Rep_preal(A). xa + x \<in> Rep_preal(A) ==>
-             \<exists>xb \<in> Rep_preal(A). xb + (prat_of_pnat p)*x \<in> Rep_preal(A)"
-apply (cut_tac mem_Rep_preal_Ex)
-apply (induct_tac "p" rule: pnat_induct)
-apply (auto simp add: pnat_one_def pSuc_is_plus_one prat_add_mult_distrib 
-                      prat_of_pnat_add prat_add_assoc [symmetric])
-done
+(*????Why can't I get case_names like nonneg to work?*)
+lemma Gleason9_34_exists:
+  assumes A: "A \<in> preal"
+      and closed: "\<forall>x\<in>A. x + u \<in> A"
+      and nonneg: "0 \<le> z"
+     shows "\<exists>b\<in>A. b + (rat z) * u \<in> A"
+proof (cases z)
+  case (1 n)
+  show ?thesis
+  proof (simp add: prems, induct n)
+    case 0
+      from preal_nonempty [OF A]
+      show ?case  by force 
+    case (Suc k)
+      from this obtain b where "b \<in> A" "b + rat (int k) * u \<in> A" ..
+      hence "b + rat (int k)*u + u \<in> A" by (simp add: closed)
+      thus ?case by (force simp add: left_distrib add_ac prems) 
+  qed
+next
+  case (2 n)
+  with nonneg show ?thesis by simp
+qed
+
 
-lemma lemma1b_gleason9_34:
-     "Abs_prat (ratrel `` {(y, z)}) < 
-      xb +
-      Abs_prat (ratrel `` {(x*y, Abs_pnat (Suc 0))}) * 
-      Abs_prat (ratrel `` {(w, x)})"
-apply (rule_tac j =
-        "Abs_prat (ratrel `` 
-           { (x * y, Abs_pnat (Suc 0))}) * Abs_prat (ratrel `` {(w, x)})" 
-       in prat_le_less_trans)
-apply (rule_tac [2] prat_self_less_add_right)
-apply (auto intro: lemma_Abs_prat_le3 
-            simp add: prat_mult pre_lemma_gleason9_34b pnat_mult_assoc)
-done
+lemma Gleason9_34_contra:
+  assumes A: "A \<in> preal"
+    shows "[|\<forall>x\<in>A. x + u \<in> A; 0 < u; 0 < y; y \<notin> A|] ==> False"
+proof (induct u, induct y)
+  fix a::int and b::int
+  fix c::int and d::int
+  assume bpos [simp]: "0 < b"
+     and dpos [simp]: "0 < d"
+     and closed: "\<forall>x\<in>A. x + (Fract c d) \<in> A"
+     and upos: "0 < Fract c d"
+     and ypos: "0 < Fract a b"
+     and notin: "Fract a b \<notin> A"
+  have cpos [simp]: "0 < c" 
+    by (simp add: zero_less_Fract_iff [OF dpos, symmetric] upos) 
+  have apos [simp]: "0 < a" 
+    by (simp add: zero_less_Fract_iff [OF bpos, symmetric] ypos) 
+  let ?k = "a*d"
+  have frle: "Fract a b \<le> rat ?k * (Fract c d)" 
+  proof -
+    have "?thesis = ((a * d * b * d) \<le> c * b * (a * d * b * d))"
+      by (simp add: rat_def mult_rat le_rat order_less_imp_not_eq2 mult_ac) 
+    moreover
+    have "(1 * (a * d * b * d)) \<le> c * b * (a * d * b * d)"
+      by (rule mult_mono, 
+          simp_all add: int_one_le_iff_zero_less zero_less_mult_iff 
+                        order_less_imp_le)
+    ultimately
+    show ?thesis by simp
+  qed
+  have k: "0 \<le> ?k" by (simp add: order_less_imp_le zero_less_mult_iff)  
+  from Gleason9_34_exists [OF A closed k]
+  obtain z where z: "z \<in> A" 
+             and mem: "z + rat ?k * Fract c d \<in> A" ..
+  have less: "z + rat ?k * Fract c d < Fract a b"
+    by (rule not_in_preal_ub [OF A notin mem ypos])
+  have "0<z" by (rule preal_imp_pos [OF A z])
+  with frle and less show False by arith 
+qed
 
-lemma lemma_gleason9_34a:
-     "\<forall>xa \<in> Rep_preal(A). xa + x \<in> Rep_preal(A) ==> False"
-apply (cut_tac X = A in not_mem_Rep_preal_Ex)
-apply (erule exE)
-apply (drule not_in_preal_ub)
-apply (rule_tac z = x in eq_Abs_prat)
-apply (rule_tac z = xa in eq_Abs_prat)
-apply (drule_tac p = "y*xb" in lemma1_gleason9_34)
-apply (erule bexE)
-apply (cut_tac x = y and y = xb and w = xaa and z = ya and xb = xba in lemma1b_gleason9_34)
-apply (drule_tac x = "xba + prat_of_pnat (y * xb) * x" in bspec)
-apply (auto intro: prat_less_asym simp add: prat_of_pnat_def)
-done
 
-lemma lemma_gleason9_34: "\<exists>r \<in> Rep_preal(R). r + x \<notin> Rep_preal(R)"
-apply (rule ccontr)
-apply (blast intro: lemma_gleason9_34a)
-done
+lemma Gleason9_34:
+  assumes A: "A \<in> preal"
+      and upos: "0 < u"
+    shows "\<exists>r \<in> A. r + u \<notin> A"
+proof (rule ccontr, simp)
+  assume closed: "\<forall>r\<in>A. r + u \<in> A"
+  from preal_exists_bound [OF A]
+  obtain y where y: "y \<notin> A" and ypos: "0 < y" by blast
+  show False
+    by (rule Gleason9_34_contra [OF A closed upos ypos y])
+qed
+
 
 
 subsection{*Gleason's Lemma 9-3.6*}
 
-lemma lemma1_gleason9_36: "r + r*qinv(xa)*Q3 = r*qinv(xa)*(xa + Q3)"
-apply (simp (no_asm_use) add: prat_add_mult_distrib2 prat_mult_assoc)
-done
+lemma lemma_gleason9_36:
+  assumes A: "A \<in> preal"
+      and x: "1 < x"
+    shows "\<exists>r \<in> A. r*x \<notin> A"
+proof -
+  from preal_nonempty [OF A]
+  obtain y where y: "y \<in> A" and  ypos: "0<y" ..
+  show ?thesis 
+  proof (rule classical)
+    assume "~(\<exists>r\<in>A. r * x \<notin> A)"
+    with y have ymem: "y * x \<in> A" by blast 
+    from ypos mult_strict_left_mono [OF x]
+    have yless: "y < y*x" by simp 
+    let ?d = "y*x - y"
+    from yless have dpos: "0 < ?d" and eq: "y + ?d = y*x" by auto
+    from Gleason9_34 [OF A dpos]
+    obtain r where r: "r\<in>A" and notin: "r + ?d \<notin> A" ..
+    have rpos: "0<r" by (rule preal_imp_pos [OF A r])
+    with dpos have rdpos: "0 < r + ?d" by arith
+    have "~ (r + ?d \<le> y + ?d)"
+    proof
+      assume le: "r + ?d \<le> y + ?d" 
+      from ymem have yd: "y + ?d \<in> A" by (simp add: eq)
+      have "r + ?d \<in> A" by (rule preal_downwards_closed' [OF A yd rdpos le])
+      with notin show False by simp
+    qed
+    hence "y < r" by simp
+    with ypos have  dless: "?d < (r * ?d)/y"
+      by (simp add: pos_less_divide_eq mult_commute [of ?d]
+                    mult_strict_right_mono dpos)
+    have "r + ?d < r*x"
+    proof -
+      have "r + ?d < r + (r * ?d)/y" by (simp add: dless)
+      also with ypos have "... = (r/y) * (y + ?d)"
+	by (simp only: right_distrib divide_inverse_zero mult_ac, simp)
+      also have "... = r*x" using ypos
+	by simp
+      finally show "r + ?d < r*x" .
+    qed
+    with r notin rdpos
+    show "\<exists>r\<in>A. r * x \<notin> A" by (blast dest:  preal_downwards_closed [OF A])
+  qed  
+qed
 
-lemma lemma2_gleason9_36: "r*qinv(xa)*(xa*x) = r*x"
-apply (simp (no_asm_use) add: prat_mult_ac)
+subsection{*Existence of Inverse: Part 2*}
+
+lemma mem_Rep_preal_inverse_iff:
+      "(z \<in> Rep_preal(inverse R)) = 
+       (0 < z \<and> (\<exists>y. z < y \<and> inverse y \<notin> Rep_preal R))"
+apply (simp add: preal_inverse_def mem_inverse_set Rep_preal)
+apply (simp add: inverse_set_def) 
 done
 
-(*** FIXME: long! ***)
-lemma lemma_gleason9_36:
-     "prat_of_pnat 1 < x ==> \<exists>r \<in> Rep_preal(A). r*x \<notin> Rep_preal(A)"
-apply (rule_tac X1 = A in mem_Rep_preal_Ex [THEN exE])
-apply (rule_tac Q = "xa*x \<in> Rep_preal (A) " in excluded_middle [THEN disjE])
-apply fast
-apply (drule_tac x = xa in prat_self_less_mult_right)
-apply (erule prat_lessE)
-apply (cut_tac R = A and x = Q3 in lemma_gleason9_34)
-apply (drule sym, auto)
-apply (frule not_in_preal_ub)
-apply (drule_tac x = "xa + Q3" in bspec, assumption)
-apply (drule prat_add_right_less_cancel)
-apply (drule_tac x = "qinv (xa) *Q3" in prat_mult_less2_mono1)
-apply (drule_tac x = r in prat_add_less2_mono2)
-apply (simp add: prat_mult_assoc [symmetric] lemma1_gleason9_36)
-apply (drule sym)
-apply (auto simp add: lemma2_gleason9_36)
-apply (rule_tac x = r in bexI)
-apply (rule notI)
-apply (drule_tac y = "r*x" in Rep_preal [THEN prealE_lemma3b], auto)
+lemma Rep_preal_of_rat:
+     "0 < q ==> Rep_preal (preal_of_rat q) = {x. 0 < x \<and> x < q}"
+by (simp add: preal_of_rat_def rat_mem_preal) 
+
+lemma subset_inverse_mult_lemma:
+      assumes xpos: "0 < x" and xless: "x < 1"
+         shows "\<exists>r u y. 0 < r & r < y & inverse y \<notin> Rep_preal R & 
+                        u \<in> Rep_preal R & x = r * u"
+proof -
+  from xpos and xless have "1 < inverse x" by (simp add: one_less_inverse_iff)
+  from lemma_gleason9_36 [OF Rep_preal this]
+  obtain r where r: "r \<in> Rep_preal R" 
+             and notin: "r * (inverse x) \<notin> Rep_preal R" ..
+  have rpos: "0<r" by (rule preal_imp_pos [OF Rep_preal r])
+  from preal_exists_greater [OF Rep_preal r]
+  obtain u where u: "u \<in> Rep_preal R" and rless: "r < u" ..
+  have upos: "0<u" by (rule preal_imp_pos [OF Rep_preal u])
+  show ?thesis
+  proof (intro exI conjI)
+    show "0 < x/u" using xpos upos
+      by (simp add: zero_less_divide_iff)  
+    show "x/u < x/r" using xpos upos rpos
+      by (simp add: divide_inverse_zero mult_less_cancel_left rless) 
+    show "inverse (x / r) \<notin> Rep_preal R" using notin
+      by (simp add: divide_inverse_zero mult_commute) 
+    show "u \<in> Rep_preal R" by (rule u) 
+    show "x = x / u * u" using upos 
+      by (simp add: divide_inverse_zero mult_commute) 
+  qed
+qed
+
+lemma subset_inverse_mult: 
+     "Rep_preal(preal_of_rat 1) \<subseteq> Rep_preal(inverse R * R)"
+apply (auto simp add: Bex_def Rep_preal_of_rat mem_Rep_preal_inverse_iff 
+                      mem_Rep_preal_mult_iff)
+apply (blast dest: subset_inverse_mult_lemma) 
 done
 
-lemma lemma_gleason9_36a:
-     "prat_of_pnat (Abs_pnat (Suc 0)) < x ==>
-      \<exists>r \<in> Rep_preal(A). r*x \<notin> Rep_preal(A)"
-apply (rule lemma_gleason9_36)
-apply (simp (no_asm_simp) add: pnat_one_def)
+lemma inverse_mult_subset_lemma:
+     assumes rpos: "0 < r" 
+         and rless: "r < y"
+         and notin: "inverse y \<notin> Rep_preal R"
+         and q: "q \<in> Rep_preal R"
+     shows "r*q < 1"
+proof -
+  have "q < inverse y" using rpos rless
+    by (simp add: not_in_preal_ub [OF Rep_preal notin] q)
+  hence "r * q < r/y" using rpos
+    by (simp add: divide_inverse_zero mult_less_cancel_left)
+  also have "... \<le> 1" using rpos rless
+    by (simp add: pos_divide_le_eq)
+  finally show ?thesis .
+qed
+
+lemma inverse_mult_subset:
+     "Rep_preal(inverse R * R) \<subseteq> Rep_preal(preal_of_rat 1)"
+apply (auto simp add: Bex_def Rep_preal_of_rat mem_Rep_preal_inverse_iff 
+                      mem_Rep_preal_mult_iff)
+apply (simp add: zero_less_mult_iff preal_imp_pos [OF Rep_preal]) 
+apply (blast intro: inverse_mult_subset_lemma) 
+done
+
+lemma preal_mult_inverse:
+     "inverse R * R = (preal_of_rat 1)"
+apply (rule inj_Rep_preal [THEN injD])
+apply (rule equalityI [OF inverse_mult_subset subset_inverse_mult]) 
+done
+
+lemma preal_mult_inverse_right:
+     "R * inverse R = (preal_of_rat 1)"
+apply (rule preal_mult_commute [THEN subst])
+apply (rule preal_mult_inverse)
 done
 
 
-subsection{*Existence of Inverse: Part 2*}
-lemma preal_mem_mult_invI:
-     "x \<in> Rep_preal(preal_of_prat (prat_of_pnat (Abs_pnat (Suc 0))))
-      ==> x \<in> Rep_preal(pinv(A)*A)"
-apply (auto intro!: mem_Rep_preal_multI simp add: pinv_def preal_of_prat_def)
-apply (rule preal_mem_inv_set [THEN Abs_preal_inverse, THEN ssubst])
-apply (drule prat_qinv_gt_1)
-apply (drule_tac A = A in lemma_gleason9_36a, auto)
-apply (drule Rep_preal [THEN prealE_lemma4a])
-apply (auto, drule qinv_prat_less)
-apply (rule_tac x = "qinv (u) *x" in exI)
-apply (rule conjI)
-apply (rule_tac x = "qinv (r) *x" in exI)
-apply (auto intro: prat_mult_less2_mono1 simp add: qinv_mult_eq qinv_qinv)
-apply (rule_tac x = u in bexI)
-apply (auto simp add: prat_mult_assoc prat_mult_left_commute)
-done
-
-lemma preal_mult_inv:
-     "pinv(A)*A = (preal_of_prat (prat_of_pnat (Abs_pnat (Suc 0))))"
-apply (rule inj_Rep_preal [THEN injD])
-apply (fast dest: preal_mem_mult_invD preal_mem_mult_invI)
-done
+text{*Theorems needing @{text Gleason9_34}*}
 
-lemma preal_mult_inv_right:
-     "A*pinv(A) = (preal_of_prat (prat_of_pnat (Abs_pnat (Suc 0))))"
-apply (rule preal_mult_commute [THEN subst])
-apply (rule preal_mult_inv)
-done
-
-
-text{*Theorems needing @{text lemma_gleason9_34}*}
+lemma Rep_preal_self_subset: "Rep_preal (R) \<subseteq> Rep_preal(R + S)"
+proof 
+  fix r
+  assume r: "r \<in> Rep_preal R"
+  have rpos: "0<r" by (rule preal_imp_pos [OF Rep_preal r])
+  from mem_Rep_preal_Ex 
+  obtain y where y: "y \<in> Rep_preal S" ..
+  have ypos: "0<y" by (rule preal_imp_pos [OF Rep_preal y])
+  have ry: "r+y \<in> Rep_preal(R + S)" using r y
+    by (auto simp add: mem_Rep_preal_add_iff)
+  show "r \<in> Rep_preal(R + S)" using r ypos rpos 
+    by (simp add:  preal_downwards_closed [OF Rep_preal ry]) 
+qed
 
-lemma Rep_preal_self_subset: "Rep_preal (R1) \<subseteq> Rep_preal(R1 + R2)"
-apply (cut_tac X = R2 in mem_Rep_preal_Ex)
-apply (auto intro!: bexI 
-            intro: Rep_preal [THEN prealE_lemma3b] prat_self_less_add_left 
-                   mem_Rep_preal_addI)
-done
+lemma Rep_preal_sum_not_subset: "~ Rep_preal (R + S) \<subseteq> Rep_preal(R)"
+proof -
+  from mem_Rep_preal_Ex 
+  obtain y where y: "y \<in> Rep_preal S" ..
+  have ypos: "0<y" by (rule preal_imp_pos [OF Rep_preal y])
+  from  Gleason9_34 [OF Rep_preal ypos]
+  obtain r where r: "r \<in> Rep_preal R" and notin: "r + y \<notin> Rep_preal R" ..
+  have "r + y \<in> Rep_preal (R + S)" using r y
+    by (auto simp add: mem_Rep_preal_add_iff)
+  thus ?thesis using notin by blast
+qed
 
-lemma Rep_preal_sum_not_subset: "~ Rep_preal (R1 + R2) \<subseteq> Rep_preal(R1)"
-apply (cut_tac X = R2 in mem_Rep_preal_Ex)
-apply (erule exE)
-apply (cut_tac R = R1 in lemma_gleason9_34)
-apply (auto intro: mem_Rep_preal_addI)
-done
-
-lemma Rep_preal_sum_not_eq: "Rep_preal (R1 + R2) \<noteq> Rep_preal(R1)"
-apply (rule notI)
-apply (erule equalityE)
-apply (simp add: Rep_preal_sum_not_subset)
-done
+lemma Rep_preal_sum_not_eq: "Rep_preal (R + S) \<noteq> Rep_preal(R)"
+by (insert Rep_preal_sum_not_subset, blast)
 
 text{*at last, Gleason prop. 9-3.5(iii) page 123*}
-lemma preal_self_less_add_left: "(R1::preal) < R1 + R2"
+lemma preal_self_less_add_left: "(R::preal) < R + S"
 apply (unfold preal_less_def psubset_def)
 apply (simp add: Rep_preal_self_subset Rep_preal_sum_not_eq [THEN not_sym])
 done
 
-lemma preal_self_less_add_right: "(R1::preal) < R2 + R1"
-apply (simp add: preal_add_commute preal_self_less_add_left)
-done
+lemma preal_self_less_add_right: "(R::preal) < S + R"
+by (simp add: preal_add_commute preal_self_less_add_left)
+
+lemma preal_not_eq_self: "x \<noteq> x + (y::preal)"
+by (insert preal_self_less_add_left [of x y], auto)
 
 
-subsection{*The @{text "\<le>"} Ordering*}
-
-lemma preal_less_le_iff: "(~(w < z)) = (z \<le> (w::preal))"
-apply (unfold preal_le_def psubset_def preal_less_def)
-apply (insert preal_linear [of w z])
-apply (auto simp add: preal_less_def psubset_def)
-done
-
-lemma preal_le_iff_less_or_eq:
-     "((x::preal) \<le> y) = (x < y | x = y)"
-apply (unfold preal_le_def preal_less_def psubset_def)
-apply (auto intro: inj_Rep_preal [THEN injD])
-done
-
-lemma preal_le_refl: "w \<le> (w::preal)"
-apply (simp add: preal_le_def)
-done
-
-lemma preal_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::preal)"
-apply (simp add: preal_le_iff_less_or_eq) 
-apply (blast intro: preal_less_trans)
-done
-
-lemma preal_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::preal)"
-apply (simp add: preal_le_iff_less_or_eq) 
-apply (blast intro: preal_less_asym)
-done
+subsection{*Subtraction for Positive Reals*}
 
-lemma preal_neq_iff: "(w \<noteq> z) = (w<z | z < (w::preal))"
-apply (insert preal_linear [of w z])
-apply (auto elim: preal_less_irrefl)
-done
-
-(* Axiom 'order_less_le' of class 'order': *)
-lemma preal_less_le: "((w::preal) < z) = (w \<le> z & w \<noteq> z)"
-apply (simp (no_asm) add: preal_less_le_iff [symmetric] preal_neq_iff)
-apply (blast elim!: preal_less_asym)
-done
-
-instance preal :: order
-proof qed
- (assumption |
-  rule preal_le_refl preal_le_trans preal_le_anti_sym preal_less_le)+
-
-lemma preal_le_linear: "x <= y | y <= (x::preal)"
-apply (insert preal_linear [of x y]) 
-apply (auto simp add: order_less_le) 
-done
-
-instance preal :: linorder
-  by (intro_classes, rule preal_le_linear)
-
-
-subsection{*Gleason prop. 9-3.5(iv), page 123*}
-
-text{*Proving @{term "A < B ==> \<exists>D. A + D = B"}*}
-
-text{*Define the claimed D and show that it is a positive real*}
+text{*Gleason prop. 9-3.5(iv), page 123: proving @{term "A < B ==> \<exists>D. A + D =
+B"}. We define the claimed @{term D} and show that it is a positive real*}
 
 text{*Part 1 of Dedekind sections definition*}
-lemma lemma_ex_mem_less_left_add1:
-     "A < B ==>
-      \<exists>q. q \<in> {d. \<exists>n. n \<notin> Rep_preal(A) & n + d \<in> Rep_preal(B)}"
-apply (unfold preal_less_def psubset_def)
-apply (clarify) 
-apply (drule_tac x1 = B in Rep_preal [THEN prealE_lemma4a])
-apply (auto simp add: prat_less_def)
-done
-
-lemma preal_less_set_not_empty:
-     "A < B ==> {} < {d. \<exists>n. n \<notin> Rep_preal(A) & n + d \<in> Rep_preal(B)}"
-apply (drule lemma_ex_mem_less_left_add1)
-apply (auto intro!: psubsetI)
+lemma diff_set_not_empty:
+     "R < S ==> {} \<subset> diff_set (Rep_preal S) (Rep_preal R)"
+apply (auto simp add: preal_less_def diff_set_def elim!: equalityE) 
+apply (frule_tac x1 = S in Rep_preal [THEN preal_exists_greater])
+apply (drule preal_imp_pos [OF Rep_preal], clarify)
+apply (cut_tac a=x and b=u in add_eq_exists, force) 
 done
 
 text{*Part 2 of Dedekind sections definition*}
-lemma lemma_ex_not_mem_less_left_add1:
-     "\<exists>q. q \<notin> {d. \<exists>n. n \<notin> Rep_preal(A) & n + d \<in> Rep_preal(B)}"
-apply (cut_tac X = B in not_mem_Rep_preal_Ex)
+lemma diff_set_nonempty:
+     "\<exists>q. 0 < q & q \<notin> diff_set (Rep_preal S) (Rep_preal R)"
+apply (cut_tac X = S in Rep_preal_exists_bound)
 apply (erule exE)
 apply (rule_tac x = x in exI, auto)
-apply (cut_tac x = x and y = n in prat_self_less_add_right)
-apply (auto dest: Rep_preal [THEN prealE_lemma3b])
+apply (simp add: diff_set_def) 
+apply (auto dest: Rep_preal [THEN preal_downwards_closed])
 done
 
-lemma preal_less_set_not_prat_set:
-     "{d. \<exists>n. n \<notin> Rep_preal(A) & n + d \<in> Rep_preal(B)} < UNIV"
-apply (auto intro!: psubsetI)
-apply (cut_tac A = A and B = B in lemma_ex_not_mem_less_left_add1, auto)
-done
+lemma diff_set_not_rat_set:
+     "diff_set (Rep_preal S) (Rep_preal R) < {r. 0 < r}" (is "?lhs < ?rhs")
+proof
+  show "?lhs \<subseteq> ?rhs" by (auto simp add: diff_set_def) 
+  show "?lhs \<noteq> ?rhs" using diff_set_nonempty by blast
+qed
 
 text{*Part 3 of Dedekind sections definition*}
-lemma preal_less_set_lemma3:
-     "A < B ==> \<forall>y \<in> {d. \<exists>n. n \<notin> Rep_preal(A) & n + d \<in> Rep_preal(B)}.
-     \<forall>z. z < y --> z \<in> {d. \<exists>n. n \<notin> Rep_preal(A) & n + d \<in> Rep_preal(B)}"
-apply auto
-apply (drule_tac x = n in prat_add_less2_mono2)
-apply (drule Rep_preal [THEN prealE_lemma3b], auto)
+lemma diff_set_lemma3:
+     "[|R < S; u \<in> diff_set (Rep_preal S) (Rep_preal R); 0 < z; z < u|] 
+      ==> z \<in> diff_set (Rep_preal S) (Rep_preal R)"
+apply (auto simp add: diff_set_def) 
+apply (rule_tac x=x in exI) 
+apply (drule Rep_preal [THEN preal_downwards_closed], auto)
 done
 
-lemma preal_less_set_lemma4:
-     "A < B ==> \<forall>y \<in> {d. \<exists>n. n \<notin> Rep_preal(A) & n + d \<in> Rep_preal(B)}.
-        Bex {d. \<exists>n. n \<notin> Rep_preal(A) & n + d \<in> Rep_preal(B)} (op < y)"
-apply auto
-apply (drule Rep_preal [THEN prealE_lemma4a])
-apply (auto simp add: prat_less_def prat_add_assoc)
+text{*Part 4 of Dedekind sections definition*}
+lemma diff_set_lemma4:
+     "[|R < S; y \<in> diff_set (Rep_preal S) (Rep_preal R)|] 
+      ==> \<exists>u \<in> diff_set (Rep_preal S) (Rep_preal R). y < u"
+apply (auto simp add: diff_set_def) 
+apply (drule Rep_preal [THEN preal_exists_greater], clarify) 
+apply (cut_tac a="x+y" and b=u in add_eq_exists, clarify)  
+apply (rule_tac x="y+xa" in exI) 
+apply (auto simp add: add_ac)
 done
 
-lemma preal_mem_less_set:
-     "!! (A::preal). A < B ==>
-      {d. \<exists>n. n \<notin> Rep_preal(A) & n + d \<in> Rep_preal(B)}: preal"
-apply (rule prealI2)
-apply (rule preal_less_set_not_empty)
-apply (rule_tac [2] preal_less_set_not_prat_set)
-apply (rule_tac [2] preal_less_set_lemma3)
-apply (rule_tac [3] preal_less_set_lemma4, auto)
+lemma mem_diff_set:
+     "R < S ==> diff_set (Rep_preal S) (Rep_preal R) \<in> preal"
+apply (unfold preal_def cut_def)
+apply (blast intro!: diff_set_not_empty diff_set_not_rat_set
+                     diff_set_lemma3 diff_set_lemma4)
+done
+
+lemma mem_Rep_preal_diff_iff:
+      "R < S ==>
+       (z \<in> Rep_preal(S-R)) = 
+       (\<exists>x. 0 < x & 0 < z & x \<notin> Rep_preal R & x + z \<in> Rep_preal S)"
+apply (simp add: preal_diff_def mem_diff_set Rep_preal)
+apply (force simp add: diff_set_def) 
 done
 
-text{*proving that @{term "A + D \<le> B"}*}
-lemma preal_less_add_left_subsetI:
-       "!! (A::preal). A < B ==>
-          A + Abs_preal({d. \<exists>n. n \<notin> Rep_preal(A) & n + d \<in> Rep_preal(B)}) \<le> B"
-apply (unfold preal_le_def)
-apply (rule subsetI)
-apply (drule mem_Rep_preal_addD)
-apply (auto simp add: preal_mem_less_set [THEN Abs_preal_inverse])
-apply (drule not_in_preal_ub)
-apply (drule bspec, assumption)
-apply (drule_tac x = y in prat_add_less2_mono1)
-apply (drule_tac x1 = B in Rep_preal [THEN prealE_lemma3b], auto)
+
+text{*proving that @{term "R + D \<le> S"}*}
+
+lemma less_add_left_lemma:
+  assumes Rless: "R < S"
+      and a: "a \<in> Rep_preal R"
+      and cb: "c + b \<in> Rep_preal S"
+      and "c \<notin> Rep_preal R"
+      and "0 < b"
+      and "0 < c"
+  shows "a + b \<in> Rep_preal S"
+proof -
+  have "0<a" by (rule preal_imp_pos [OF Rep_preal a])
+  moreover
+  have "a < c" using prems
+    by (blast intro: not_in_Rep_preal_ub ) 
+  ultimately show ?thesis using prems
+    by (simp add: preal_downwards_closed [OF Rep_preal cb]) 
+qed
+
+lemma less_add_left_le1:
+       "R < (S::preal) ==> R + (S-R) \<le> S"
+apply (auto simp add: Bex_def preal_le_def mem_Rep_preal_add_iff 
+                      mem_Rep_preal_diff_iff)
+apply (blast intro: less_add_left_lemma) 
 done
 
-subsection{*proving that @{term "B \<le> A + D"} --- trickier*}
+subsection{*proving that @{term "S \<le> R + D"} --- trickier*}
 
 lemma lemma_sum_mem_Rep_preal_ex:
-     "x \<in> Rep_preal(B) ==> \<exists>e. x + e \<in> Rep_preal(B)"
-apply (drule Rep_preal [THEN prealE_lemma4a])
-apply (auto simp add: prat_less_def)
+     "x \<in> Rep_preal S ==> \<exists>e. 0 < e & x + e \<in> Rep_preal S"
+apply (drule Rep_preal [THEN preal_exists_greater], clarify) 
+apply (cut_tac a=x and b=u in add_eq_exists, auto) 
 done
 
-lemma preal_less_add_left_subsetI2:
-       "!! (A::preal). A < B ==>
-          B \<le> A + Abs_preal({d. \<exists>n. n \<notin> Rep_preal(A) & n + d \<in> Rep_preal(B)})"
-apply (unfold preal_le_def)
-apply (rule subsetI)
-apply (rule_tac Q = "x \<in> Rep_preal (A) " in excluded_middle [THEN disjE])
-apply (rule mem_Rep_preal_addI)
-apply (drule lemma_sum_mem_Rep_preal_ex)
-apply (erule exE)
-apply (cut_tac R = A and x = e in lemma_gleason9_34, erule bexE)
-apply (drule not_in_preal_ub, drule bspec, assumption)
-apply (erule prat_lessE)
-apply (rule_tac x = r in bexI)
-apply (rule_tac x = Q3 in bexI)
-apply (cut_tac [4] Rep_preal_self_subset)
-apply (auto simp add: preal_mem_less_set [THEN Abs_preal_inverse])
-apply (rule_tac x = "r+e" in exI)
-apply (simp add: prat_add_ac)
+lemma less_add_left_lemma2:
+  assumes Rless: "R < S"
+      and x:     "x \<in> Rep_preal S"
+      and xnot: "x \<notin>  Rep_preal R"
+  shows "\<exists>u v z. 0 < v & 0 < z & u \<in> Rep_preal R & z \<notin> Rep_preal R & 
+                     z + v \<in> Rep_preal S & x = u + v"
+proof -
+  have xpos: "0<x" by (rule preal_imp_pos [OF Rep_preal x])
+  from lemma_sum_mem_Rep_preal_ex [OF x]
+  obtain e where epos: "0 < e" and xe: "x + e \<in> Rep_preal S" by blast
+  from  Gleason9_34 [OF Rep_preal epos]
+  obtain r where r: "r \<in> Rep_preal R" and notin: "r + e \<notin> Rep_preal R" ..
+  with x xnot xpos have rless: "r < x" by (blast intro: not_in_Rep_preal_ub)
+  from add_eq_exists [of r x]
+  obtain y where eq: "x = r+y" by auto
+  show ?thesis 
+  proof (intro exI conjI)
+    show "r \<in> Rep_preal R" by (rule r)
+    show "r + e \<notin> Rep_preal R" by (rule notin)
+    show "r + e + y \<in> Rep_preal S" using xe eq by (simp add: add_ac)
+    show "x = r + y" by (simp add: eq)
+    show "0 < r + e" using epos preal_imp_pos [OF Rep_preal r]
+      by simp
+    show "0 < y" using rless eq by arith
+  qed
+qed
+
+lemma less_add_left_le2: "R < (S::preal) ==> S \<le> R + (S-R)"
+apply (auto simp add: preal_le_def)
+apply (case_tac "x \<in> Rep_preal R")
+apply (cut_tac Rep_preal_self_subset [of R], force)
+apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_diff_iff)
+apply (blast dest: less_add_left_lemma2)
 done
 
-(*** required proof ***)
-lemma preal_less_add_left:
-     "!! (A::preal). A < B ==>
-          A + Abs_preal({d. \<exists>n. n \<notin> Rep_preal(A) & n + d \<in> Rep_preal(B)}) = B"
-apply (blast intro: preal_le_anti_sym preal_less_add_left_subsetI preal_less_add_left_subsetI2)
-done
+lemma less_add_left: "R < (S::preal) ==> R + (S-R) = S"
+by (blast intro: preal_le_anti_sym [OF less_add_left_le1 less_add_left_le2])
 
-lemma preal_less_add_left_Ex: "!! (A::preal). A < B ==> \<exists>D. A + D = B"
-by (fast dest: preal_less_add_left)
+lemma less_add_left_Ex: "R < (S::preal) ==> \<exists>D. R + D = S"
+by (fast dest: less_add_left)
 
-lemma preal_add_less2_mono1: "!!(A::preal). A < B ==> A + C < B + C"
-apply (auto dest!: preal_less_add_left_Ex simp add: preal_add_assoc)
+lemma preal_add_less2_mono1: "R < (S::preal) ==> R + T < S + T"
+apply (auto dest!: less_add_left_Ex simp add: preal_add_assoc)
 apply (rule_tac y1 = D in preal_add_commute [THEN subst])
 apply (auto intro: preal_self_less_add_left simp add: preal_add_assoc [symmetric])
 done
 
-lemma preal_add_less2_mono2: "!!(A::preal). A < B ==> C + A < C + B"
-by (auto intro: preal_add_less2_mono1 simp add: preal_add_commute)
+lemma preal_add_less2_mono2: "R < (S::preal) ==> T + R < T + S"
+by (auto intro: preal_add_less2_mono1 simp add: preal_add_commute [of T])
 
-lemma preal_mult_less_mono1:
-      "!!(q1::preal). q1 < q2 ==> q1 * x < q2 * x"
-apply (drule preal_less_add_left_Ex)
-apply (auto simp add: preal_add_mult_distrib preal_self_less_add_left)
-done
-
-lemma preal_mult_left_less_mono1: "!!(q1::preal). q1 < q2  ==> x * q1 < x * q2"
-by (auto dest: preal_mult_less_mono1 simp add: preal_mult_commute)
-
-lemma preal_mult_left_le_mono1: "!!(q1::preal). q1 \<le> q2  ==> x * q1 \<le> x * q2"
-apply (simp add: preal_le_iff_less_or_eq) 
-apply (blast intro!: preal_mult_left_less_mono1)
+lemma preal_add_right_less_cancel: "R + T < S + T ==> R < (S::preal)"
+apply (insert linorder_less_linear [of R S], auto)
+apply (drule_tac R = S and T = T in preal_add_less2_mono1)
+apply (blast dest: order_less_trans) 
 done
 
-lemma preal_mult_le_mono1: "!!(q1::preal). q1 \<le> q2  ==> q1 * x \<le> q2 * x"
-by (auto dest: preal_mult_left_le_mono1 simp add: preal_mult_commute)
-
-lemma preal_add_left_le_mono1: "!!(q1::preal). q1 \<le> q2  ==> x + q1 \<le> x + q2"
-apply (simp add: preal_le_iff_less_or_eq) 
-apply (auto intro!: preal_add_less2_mono1 simp add: preal_add_commute)
-done
-
-lemma preal_add_le_mono1: "!!(q1::preal). q1 \<le> q2  ==> q1 + x \<le> q2 + x"
-by (auto dest: preal_add_left_le_mono1 simp add: preal_add_commute)
+lemma preal_add_left_less_cancel: "T + R < T + S ==> R <  (S::preal)"
+by (auto elim: preal_add_right_less_cancel simp add: preal_add_commute [of T])
 
-lemma preal_add_right_less_cancel: "!!(A::preal). A + C < B + C ==> A < B"
-apply (cut_tac preal_linear)
-apply (auto elim: preal_less_irrefl)
-apply (drule_tac A = B and C = C in preal_add_less2_mono1)
-apply (fast dest: preal_less_trans elim: preal_less_irrefl)
-done
-
-lemma preal_add_left_less_cancel: "!!(A::preal). C + A < C + B ==> A < B"
-by (auto elim: preal_add_right_less_cancel simp add: preal_add_commute)
-
-lemma preal_add_less_iff1 [simp]: "((A::preal) + C < B + C) = (A < B)"
+lemma preal_add_less_cancel_right: "((R::preal) + T < S + T) = (R < S)"
 by (blast intro: preal_add_less2_mono1 preal_add_right_less_cancel)
 
-lemma preal_add_less_iff2 [simp]: "(C + (A::preal) < C + B) = (A < B)"
+lemma preal_add_less_cancel_left: "(T + (R::preal) < T + S) = (R < S)"
 by (blast intro: preal_add_less2_mono2 preal_add_left_less_cancel)
 
+lemma preal_add_le_cancel_right: "((R::preal) + T \<le> S + T) = (R \<le> S)"
+by (simp add: linorder_not_less [symmetric] preal_add_less_cancel_right) 
+
+lemma preal_add_le_cancel_left: "(T + (R::preal) \<le> T + S) = (R \<le> S)"
+by (simp add: linorder_not_less [symmetric] preal_add_less_cancel_left) 
+
 lemma preal_add_less_mono:
      "[| x1 < y1; x2 < y2 |] ==> x1 + x2 < y1 + (y2::preal)"
-apply (auto dest!: preal_less_add_left_Ex simp add: preal_add_ac)
+apply (auto dest!: less_add_left_Ex simp add: preal_add_ac)
 apply (rule preal_add_assoc [THEN subst])
 apply (rule preal_self_less_add_right)
 done
 
-lemma preal_mult_less_mono:
-     "[| x1 < y1; x2 < y2 |] ==> x1 * x2 < y1 * (y2::preal)"
-apply (auto dest!: preal_less_add_left_Ex simp add: preal_add_mult_distrib preal_add_mult_distrib2 preal_self_less_add_left preal_add_assoc preal_mult_ac)
+lemma preal_add_right_cancel: "(R::preal) + T = S + T ==> R = S"
+apply (insert linorder_less_linear [of R S], safe)
+apply (drule_tac [!] T = T in preal_add_less2_mono1, auto)
 done
 
-lemma preal_add_right_cancel: "(A::preal) + C = B + C ==> A = B"
-apply (cut_tac preal_linear [of A B], safe)
-apply (drule_tac [!] C = C in preal_add_less2_mono1)
-apply (auto elim: preal_less_irrefl)
-done
-
-lemma preal_add_left_cancel: "!!(A::preal). C + A = C + B ==> A = B"
+lemma preal_add_left_cancel: "C + A = C + B ==> A = (B::preal)"
 by (auto intro: preal_add_right_cancel simp add: preal_add_commute)
 
-lemma preal_add_left_cancel_iff [simp]: "(C + A = C + B) = ((A::preal) = B)"
+lemma preal_add_left_cancel_iff: "(C + A = C + B) = ((A::preal) = B)"
 by (fast intro: preal_add_left_cancel)
 
-lemma preal_add_right_cancel_iff [simp]: "(A + C = B + C) = ((A::preal) = B)"
+lemma preal_add_right_cancel_iff: "(A + C = B + C) = ((A::preal) = B)"
 by (fast intro: preal_add_right_cancel)
 
+lemmas preal_cancels =
+    preal_add_less_cancel_right preal_add_less_cancel_left
+    preal_add_le_cancel_right preal_add_le_cancel_left
+    preal_add_left_cancel_iff preal_add_right_cancel_iff
 
 
 subsection{*Completeness of type @{typ preal}*}
 
 text{*Prove that supremum is a cut*}
 
-lemma preal_sup_mem_Ex:
-     "\<exists>X. X \<in> P ==> \<exists>q.  q \<in> {w. \<exists>X. X \<in> P & w \<in> Rep_preal X}"
-apply safe
-apply (cut_tac X = X in mem_Rep_preal_Ex, auto)
+text{*Part 1 of Dedekind sections definition*}
+
+lemma preal_sup_set_not_empty:
+     "P \<noteq> {} ==> {} \<subset> (\<Union>X \<in> P. Rep_preal(X))"
+apply auto
+apply (cut_tac X = x in mem_Rep_preal_Ex, auto)
 done
 
-text{*Part 1 of Dedekind sections definition*}
-lemma preal_sup_set_not_empty:
-     "\<exists>(X::preal). X \<in> P ==>
-          {} < {w. \<exists>X \<in> P. w \<in> Rep_preal X}"
-apply (drule preal_sup_mem_Ex)
-apply (auto intro!: psubsetI)
-done
 
 text{*Part 2 of Dedekind sections definition*}
-lemma preal_sup_not_mem_Ex:
-     "\<exists>Y. (\<forall>X \<in> P. X < Y)
-          ==> \<exists>q. q \<notin> {w. \<exists>X. X \<in> P & w \<in> Rep_preal(X)}"
-apply (unfold preal_less_def)
-apply (auto simp add: psubset_def)
-apply (cut_tac X = Y in not_mem_Rep_preal_Ex)
-apply (erule exE)
-apply (rule_tac x = x in exI)
-apply (auto dest!: bspec)
+
+lemma preal_sup_not_exists:
+     "\<forall>X \<in> P. X \<le> Y ==> \<exists>q. 0 < q & q \<notin> (\<Union>X \<in> P. Rep_preal(X))"
+apply (cut_tac X = Y in Rep_preal_exists_bound)
+apply (auto simp add: preal_le_def)
 done
 
-lemma preal_sup_not_mem_Ex1:
-     "\<exists>Y. (\<forall>X \<in> P. X \<le> Y)
-          ==> \<exists>q. q \<notin> {w. \<exists>X. X \<in> P & w \<in> Rep_preal(X)}"
-apply (unfold preal_le_def, safe)
-apply (cut_tac X = Y in not_mem_Rep_preal_Ex)
-apply (erule exE)
-apply (rule_tac x = x in exI)
-apply (auto dest!: bspec)
-done
-
-lemma preal_sup_set_not_prat_set:
-     "\<exists>Y. (\<forall>X \<in> P. X < Y) ==> {w. \<exists>X \<in> P. w \<in> Rep_preal(X)} < UNIV"
-apply (drule preal_sup_not_mem_Ex)
-apply (auto intro!: psubsetI)
-done
-
-lemma preal_sup_set_not_prat_set1:
-     "\<exists>Y. (\<forall>X \<in> P. X \<le> Y) ==> {w. \<exists>X \<in> P. w \<in> Rep_preal(X)} < UNIV"
-apply (drule preal_sup_not_mem_Ex1)
-apply (auto intro!: psubsetI)
+lemma preal_sup_set_not_rat_set:
+     "\<forall>X \<in> P. X \<le> Y ==> (\<Union>X \<in> P. Rep_preal(X)) < {r. 0 < r}"
+apply (drule preal_sup_not_exists)
+apply (blast intro: preal_imp_pos [OF Rep_preal])  
 done
 
 text{*Part 3 of Dedekind sections definition*}
 lemma preal_sup_set_lemma3:
-     "[|\<exists>(X::preal). X \<in> P; \<exists>Y. (\<forall>X \<in> P. X < Y) |]
-          ==> \<forall>y \<in> {w. \<exists>X \<in> P. w \<in> Rep_preal X}.
-              \<forall>z. z < y --> z \<in> {w. \<exists>X \<in> P. w \<in> Rep_preal X}"
-apply (auto elim: Rep_preal [THEN prealE_lemma3b])
-done
-
-lemma preal_sup_set_lemma3_1:
-     "[|\<exists>(X::preal). X \<in> P; \<exists>Y. (\<forall>X \<in> P. X \<le> Y) |]
-          ==> \<forall>y \<in> {w. \<exists>X \<in> P. w \<in> Rep_preal X}.
-              \<forall>z. z < y --> z \<in> {w. \<exists>X \<in> P. w \<in> Rep_preal X}"
-apply (auto elim: Rep_preal [THEN prealE_lemma3b])
-done
+     "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; u \<in> (\<Union>X \<in> P. Rep_preal(X)); 0 < z; z < u|]
+      ==> z \<in> (\<Union>X \<in> P. Rep_preal(X))"
+by (auto elim: Rep_preal [THEN preal_downwards_closed])
 
+text{*Part 4 of Dedekind sections definition*}
 lemma preal_sup_set_lemma4:
-     "[|\<exists>(X::preal). X \<in> P; \<exists>Y. (\<forall>X \<in> P. X < Y) |]
-          ==>  \<forall>y \<in> {w. \<exists>X \<in> P. w \<in> Rep_preal X}.
-              Bex {w. \<exists>X \<in> P. w \<in> Rep_preal X} (op < y)"
-apply (blast dest: Rep_preal [THEN prealE_lemma4a])
-done
-
-lemma preal_sup_set_lemma4_1:
-     "[|\<exists>(X::preal). X \<in> P; \<exists>Y. (\<forall>X \<in> P. X \<le> Y) |]
-          ==>  \<forall>y \<in> {w. \<exists>X \<in> P. w \<in> Rep_preal X}.
-              Bex {w. \<exists>X \<in> P. w \<in> Rep_preal X} (op < y)"
-apply (blast dest: Rep_preal [THEN prealE_lemma4a])
-done
+     "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; y \<in> (\<Union>X \<in> P. Rep_preal(X)) |]
+          ==> \<exists>u \<in> (\<Union>X \<in> P. Rep_preal(X)). y < u"
+by (blast dest: Rep_preal [THEN preal_exists_greater])
 
 lemma preal_sup:
-     "[|\<exists>(X::preal). X \<in> P; \<exists>Y. (\<forall>X \<in> P. X < Y) |]
-          ==> {w. \<exists>X \<in> P. w \<in> Rep_preal(X)}: preal"
-apply (rule prealI2)
-apply (rule preal_sup_set_not_empty)
-apply (rule_tac [2] preal_sup_set_not_prat_set)
-apply (rule_tac [3] preal_sup_set_lemma3)
-apply (rule_tac [5] preal_sup_set_lemma4, auto)
+     "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y|] ==> (\<Union>X \<in> P. Rep_preal(X)) \<in> preal"
+apply (unfold preal_def cut_def)
+apply (blast intro!: preal_sup_set_not_empty preal_sup_set_not_rat_set
+                     preal_sup_set_lemma3 preal_sup_set_lemma4)
 done
 
-lemma preal_sup1:
-     "[|\<exists>(X::preal). X \<in> P; \<exists>Y. (\<forall>X \<in> P. X \<le> Y) |]
-          ==> {w. \<exists>X \<in> P. w \<in> Rep_preal(X)}: preal"
-apply (rule prealI2)
-apply (rule preal_sup_set_not_empty)
-apply (rule_tac [2] preal_sup_set_not_prat_set1)
-apply (rule_tac [3] preal_sup_set_lemma3_1)
-apply (rule_tac [5] preal_sup_set_lemma4_1, auto)
-done
-
-lemma preal_psup_leI: "\<exists>Y. (\<forall>X \<in> P. X < Y) ==> \<forall>x \<in> P. x \<le> psup P"
-apply (unfold psup_def)
-apply (auto simp add: preal_le_def)
-apply (rule preal_sup [THEN Abs_preal_inverse, THEN ssubst], auto)
-done
-
-lemma preal_psup_leI2: "\<exists>Y. (\<forall>X \<in> P. X \<le> Y) ==> \<forall>x \<in> P. x \<le> psup P"
-apply (unfold psup_def)
-apply (auto simp add: preal_le_def)
-apply (rule preal_sup1 [THEN Abs_preal_inverse, THEN ssubst])
-apply (auto simp add: preal_le_def)
+lemma preal_psup_le:
+     "[| \<forall>X \<in> P. X \<le> Y;  x \<in> P |] ==> x \<le> psup P"
+apply (simp (no_asm_simp) add: preal_le_def) 
+apply (subgoal_tac "P \<noteq> {}") 
+apply (auto simp add: psup_def preal_sup) 
 done
 
-lemma preal_psup_leI2b:
-     "[| \<exists>Y. (\<forall>X \<in> P. X < Y); x \<in> P |] ==> x \<le> psup P"
-apply (blast dest!: preal_psup_leI)
-done
-
-lemma preal_psup_leI2a:
-     "[| \<exists>Y. (\<forall>X \<in> P. X \<le> Y); x \<in> P |] ==> x \<le> psup P"
-apply (blast dest!: preal_psup_leI2)
-done
-
-lemma psup_le_ub: "[| \<exists>X. X \<in> P; \<forall>X \<in> P. X < Y |] ==> psup P \<le> Y"
-apply (unfold psup_def)
+lemma psup_le_ub: "[| P \<noteq> {}; \<forall>X \<in> P. X \<le> Y |] ==> psup P \<le> Y"
+apply (simp (no_asm_simp) add: preal_le_def)
+apply (simp add: psup_def preal_sup) 
 apply (auto simp add: preal_le_def)
-apply (drule preal_sup [OF exI exI, THEN Abs_preal_inverse, THEN subst])
-apply (rotate_tac [2] 1)
-prefer 2 apply assumption
-apply (auto dest!: bspec simp add: preal_less_def psubset_def)
-done
-
-lemma psup_le_ub1: "[| \<exists>X. X \<in> P; \<forall>X \<in> P. X \<le> Y |] ==> psup P \<le> Y"
-apply (unfold psup_def)
-apply (auto simp add: preal_le_def)
-apply (drule preal_sup1 [OF exI exI, THEN Abs_preal_inverse, THEN subst])
-apply (rotate_tac [2] 1)
-prefer 2 apply assumption
-apply (auto dest!: bspec simp add: preal_less_def psubset_def preal_le_def)
 done
 
 text{*Supremum property*}
 lemma preal_complete:
-     "[|\<exists>(X::preal). X \<in> P; \<exists>Y. (\<forall>X \<in> P. X < Y) |]
-          ==> (\<forall>Y. (\<exists>X \<in> P. Y < X) = (Y < psup P))"
-apply (frule preal_sup [THEN Abs_preal_inverse], fast)
-apply (auto simp add: psup_def preal_less_def)
-apply (cut_tac x = Xa and y = Ya in preal_linear)
-apply (auto dest: psubsetD simp add: preal_less_def)
+     "[| P \<noteq> {}; \<forall>X \<in> P. X \<le> Y |] ==> (\<exists>X \<in> P. Z < X) = (Z < psup P)"
+apply (simp add: preal_less_def psup_def preal_sup)
+apply (auto simp add: preal_le_def)
+apply (rename_tac U) 
+apply (cut_tac x = U and y = Z in linorder_less_linear)
+apply (auto simp add: preal_less_def)
 done
 
 
-subsection{*The Embadding from @{typ prat} into @{typ preal}*}
+subsection{*The Embadding from @{typ rat} into @{typ preal}*}
 
-lemma lemma_preal_rat_less: "x < z1 + z2 ==> x * z1 * qinv (z1 + z2) < z1"
-apply (drule_tac x = "z1 * qinv (z1 + z2) " in prat_mult_less2_mono1)
-apply (simp add: prat_mult_ac)
-done
-
-lemma lemma_preal_rat_less2: "x < z1 + z2 ==> x * z2 * qinv (z1 + z2) < z2"
-apply (subst prat_add_commute)
-apply (drule prat_add_commute [THEN subst])
-apply (erule lemma_preal_rat_less)
+lemma preal_of_rat_add_lemma1:
+     "[|x < y + z; 0 < x; 0 < y|] ==> x * y * inverse (y + z) < (y::rat)"
+apply (frule_tac c = "y * inverse (y + z) " in mult_strict_right_mono)
+apply (simp add: zero_less_mult_iff) 
+apply (simp add: mult_ac)
 done
 
-lemma preal_of_prat_add:
-      "preal_of_prat ((z1::prat) + z2) =
-       preal_of_prat z1 + preal_of_prat z2"
-apply (unfold preal_of_prat_def preal_add_def)
+lemma preal_of_rat_add_lemma2:
+  assumes "u < x + y"
+      and "0 < x"
+      and "0 < y"
+      and "0 < u"
+  shows "\<exists>v w::rat. w < y & 0 < v & v < x & 0 < w & u = v + w"
+proof (intro exI conjI)
+  show "u * x * inverse(x+y) < x" using prems 
+    by (simp add: preal_of_rat_add_lemma1) 
+  show "u * y * inverse(x+y) < y" using prems 
+    by (simp add: preal_of_rat_add_lemma1 add_commute [of x]) 
+  show "0 < u * x * inverse (x + y)" using prems
+    by (simp add: zero_less_mult_iff) 
+  show "0 < u * y * inverse (x + y)" using prems
+    by (simp add: zero_less_mult_iff) 
+  show "u = u * x * inverse (x + y) + u * y * inverse (x + y)" using prems
+    by (simp add: left_distrib [symmetric] right_distrib [symmetric] mult_ac)
+qed
+
+lemma preal_of_rat_add:
+     "[| 0 < x; 0 < y|] 
+      ==> preal_of_rat ((x::rat) + y) = preal_of_rat x + preal_of_rat y"
+apply (unfold preal_of_rat_def preal_add_def)
+apply (simp add: rat_mem_preal) 
 apply (rule_tac f = Abs_preal in arg_cong)
-apply (auto intro: prat_add_less_mono 
-            simp add: lemma_prat_less_set_mem_preal [THEN Abs_preal_inverse])
-apply (rule_tac x = "x*z1*qinv (z1+z2) " in exI, rule conjI)
-apply (erule lemma_preal_rat_less)
-apply (rule_tac x = "x*z2*qinv (z1+z2) " in exI, rule conjI)
-apply (erule lemma_preal_rat_less2)
-apply (simp add: prat_add_mult_distrib [symmetric] 
-                 prat_add_mult_distrib2 [symmetric] prat_mult_ac)
+apply (auto simp add: add_set_def) 
+apply (blast dest: preal_of_rat_add_lemma2) 
+done
+
+lemma preal_of_rat_mult_lemma1:
+     "[|x < y; 0 < x; 0 < z|] ==> x * z * inverse y < (z::rat)"
+apply (frule_tac c = "z * inverse y" in mult_strict_right_mono)
+apply (simp add: zero_less_mult_iff)
+apply (subgoal_tac "y * (z * inverse y) = z * (y * inverse y)")
+apply (simp_all add: mult_ac)
 done
 
-lemma lemma_preal_rat_less3: "x < xa ==> x*z1*qinv(xa) < z1"
-apply (drule_tac x = "z1 * qinv xa" in prat_mult_less2_mono1)
-apply (drule prat_mult_left_commute [THEN subst])
-apply (simp add: prat_mult_ac)
-done
+lemma preal_of_rat_mult_lemma2: 
+  assumes xless: "x < y * z"
+      and xpos: "0 < x"
+      and ypos: "0 < y"
+  shows "x * z * inverse y * inverse z < (z::rat)"
+proof -
+  have "0 < y * z" using prems by simp
+  hence zpos:  "0 < z" using prems by (simp add: zero_less_mult_iff)
+  have "x * z * inverse y * inverse z = x * inverse y * (z * inverse z)"
+    by (simp add: mult_ac)
+  also have "... = x/y" using zpos
+    by (simp add: divide_inverse_zero)
+  also have "... < z"
+    by (simp add: pos_divide_less_eq [OF ypos] mult_commute) 
+  finally show ?thesis .
+qed
 
-lemma lemma_preal_rat_less4: "xa < z1 * z2 ==> xa*z2*qinv(z1*z2) < z2"
-apply (drule_tac x = "z2 * qinv (z1*z2) " in prat_mult_less2_mono1)
-apply (drule prat_mult_left_commute [THEN subst])
-apply (simp add: prat_mult_ac)
+lemma preal_of_rat_mult_lemma3:
+  assumes uless: "u < x * y"
+      and "0 < x"
+      and "0 < y"
+      and "0 < u"
+  shows "\<exists>v w::rat. v < x & w < y & 0 < v & 0 < w & u = v * w"
+proof -
+  from dense [OF uless] 
+  obtain r where "u < r" "r < x * y" by blast
+  thus ?thesis
+  proof (intro exI conjI)
+  show "u * x * inverse r < x" using prems 
+    by (simp add: preal_of_rat_mult_lemma1) 
+  show "r * y * inverse x * inverse y < y" using prems
+    by (simp add: preal_of_rat_mult_lemma2)
+  show "0 < u * x * inverse r" using prems
+    by (simp add: zero_less_mult_iff) 
+  show "0 < r * y * inverse x * inverse y" using prems
+    by (simp add: zero_less_mult_iff) 
+  have "u * x * inverse r * (r * y * inverse x * inverse y) =
+        u * (r * inverse r) * (x * inverse x) * (y * inverse y)"
+    by (simp only: mult_ac)
+  thus "u = u * x * inverse r * (r * y * inverse x * inverse y)" using prems
+    by simp
+  qed
+qed
+
+lemma preal_of_rat_mult:
+     "[| 0 < x; 0 < y|] 
+      ==> preal_of_rat ((x::rat) * y) = preal_of_rat x * preal_of_rat y"
+apply (unfold preal_of_rat_def preal_mult_def)
+apply (simp add: rat_mem_preal) 
+apply (rule_tac f = Abs_preal in arg_cong)
+apply (auto simp add: zero_less_mult_iff mult_strict_mono mult_set_def) 
+apply (blast dest: preal_of_rat_mult_lemma3) 
 done
 
-lemma preal_of_prat_mult:
-      "preal_of_prat ((z1::prat) * z2) =
-       preal_of_prat z1 * preal_of_prat z2"
-apply (unfold preal_of_prat_def preal_mult_def)
-apply (rule_tac f = Abs_preal in arg_cong)
-apply (auto intro: prat_mult_less_mono 
-            simp add: lemma_prat_less_set_mem_preal [THEN Abs_preal_inverse])
-apply (drule prat_dense, safe)
-apply (rule_tac x = "x*z1*qinv (xa) " in exI, rule conjI)
-apply (erule lemma_preal_rat_less3)
-apply (rule_tac x = " xa*z2*qinv (z1*z2) " in exI, rule conjI)
-apply (erule lemma_preal_rat_less4)
-apply (simp add: qinv_mult_eq [symmetric] prat_mult_ac)
-apply (simp add: prat_mult_assoc [symmetric])
-done
+lemma preal_of_rat_less_iff:
+      "[| 0 < x; 0 < y|] ==> (preal_of_rat x < preal_of_rat y) = (x < y)"
+by (force simp add: preal_of_rat_def preal_less_def rat_mem_preal) 
 
-lemma preal_of_prat_less_iff [simp]:
-      "(preal_of_prat p < preal_of_prat q) = (p < q)"
-apply (unfold preal_of_prat_def preal_less_def)
-apply (auto dest!: lemma_prat_set_eq elim: prat_less_trans 
-        simp add: lemma_prat_less_set_mem_preal psubset_def prat_less_not_refl)
-apply (rule_tac x = p and y = q in prat_linear_less2)
-apply (auto intro: prat_less_irrefl)
-done
+lemma preal_of_rat_le_iff:
+      "[| 0 < x; 0 < y|] ==> (preal_of_rat x \<le> preal_of_rat y) = (x \<le> y)"
+by (simp add: preal_of_rat_less_iff linorder_not_less [symmetric]) 
+
+lemma preal_of_rat_eq_iff:
+      "[| 0 < x; 0 < y|] ==> (preal_of_rat x = preal_of_rat y) = (x = y)"
+by (simp add: preal_of_rat_le_iff order_eq_iff) 
 
 
 ML
 {*
 val inj_on_Abs_preal = thm"inj_on_Abs_preal";
 val inj_Rep_preal = thm"inj_Rep_preal";
-val empty_not_mem_preal = thm"empty_not_mem_preal";
-val one_set_mem_preal = thm"one_set_mem_preal";
-val preal_psubset_empty = thm"preal_psubset_empty";
 val mem_Rep_preal_Ex = thm"mem_Rep_preal_Ex";
-val inj_preal_of_prat = thm"inj_preal_of_prat";
-val not_in_preal_ub = thm"not_in_preal_ub";
-val preal_less_not_refl = thm"preal_less_not_refl";
-val preal_less_trans = thm"preal_less_trans";
-val preal_less_not_sym = thm"preal_less_not_sym";
-val preal_linear = thm"preal_linear";
 val preal_add_commute = thm"preal_add_commute";
-val preal_add_set_not_empty = thm"preal_add_set_not_empty";
-val preal_not_mem_add_set_Ex = thm"preal_not_mem_add_set_Ex";
-val preal_add_set_not_prat_set = thm"preal_add_set_not_prat_set";
-val preal_mem_add_set = thm"preal_mem_add_set";
 val preal_add_assoc = thm"preal_add_assoc";
 val preal_add_left_commute = thm"preal_add_left_commute";
 val preal_mult_commute = thm"preal_mult_commute";
-val preal_mult_set_not_empty = thm"preal_mult_set_not_empty";
-val preal_not_mem_mult_set_Ex = thm"preal_not_mem_mult_set_Ex";
-val preal_mult_set_not_prat_set = thm"preal_mult_set_not_prat_set";
-val preal_mem_mult_set = thm"preal_mem_mult_set";
 val preal_mult_assoc = thm"preal_mult_assoc";
 val preal_mult_left_commute = thm"preal_mult_left_commute";
-val preal_mult_1 = thm"preal_mult_1";
-val preal_mult_1_right = thm"preal_mult_1_right";
-val mem_Rep_preal_addD = thm"mem_Rep_preal_addD";
-val mem_Rep_preal_addI = thm"mem_Rep_preal_addI";
-val mem_Rep_preal_add_iff = thm"mem_Rep_preal_add_iff";
-val mem_Rep_preal_multD = thm"mem_Rep_preal_multD";
-val mem_Rep_preal_multI = thm"mem_Rep_preal_multI";
-val mem_Rep_preal_mult_iff = thm"mem_Rep_preal_mult_iff";
 val preal_add_mult_distrib2 = thm"preal_add_mult_distrib2";
 val preal_add_mult_distrib = thm"preal_add_mult_distrib";
-val qinv_not_mem_Rep_preal_Ex = thm"qinv_not_mem_Rep_preal_Ex";
-val preal_inv_set_not_empty = thm"preal_inv_set_not_empty";
-val qinv_mem_Rep_preal_Ex = thm"qinv_mem_Rep_preal_Ex";
-val preal_not_mem_inv_set_Ex = thm"preal_not_mem_inv_set_Ex";
-val preal_inv_set_not_prat_set = thm"preal_inv_set_not_prat_set";
-val preal_mem_inv_set = thm"preal_mem_inv_set";
-val preal_mem_mult_invD = thm"preal_mem_mult_invD";
-val preal_mem_mult_invI = thm"preal_mem_mult_invI";
-val preal_mult_inv = thm"preal_mult_inv";
-val preal_mult_inv_right = thm"preal_mult_inv_right";
-val Rep_preal_self_subset = thm"Rep_preal_self_subset";
-val Rep_preal_sum_not_subset = thm"Rep_preal_sum_not_subset";
-val Rep_preal_sum_not_eq = thm"Rep_preal_sum_not_eq";
 val preal_self_less_add_left = thm"preal_self_less_add_left";
 val preal_self_less_add_right = thm"preal_self_less_add_right";
-val preal_less_le_iff = thm"preal_less_le_iff";
-val preal_le_refl = thm"preal_le_refl";
-val preal_le_trans = thm"preal_le_trans";
-val preal_le_anti_sym = thm"preal_le_anti_sym";
-val preal_neq_iff = thm"preal_neq_iff";
-val preal_less_le = thm"preal_less_le";
-val psubset_trans = thm"psubset_trans";
-val preal_less_set_not_empty = thm"preal_less_set_not_empty";
-val preal_less_set_not_prat_set = thm"preal_less_set_not_prat_set";
-val preal_mem_less_set = thm"preal_mem_less_set";
-val preal_less_add_left_subsetI = thm"preal_less_add_left_subsetI";
-val preal_less_add_left_subsetI2 = thm"preal_less_add_left_subsetI2";
-val preal_less_add_left = thm"preal_less_add_left";
-val preal_less_add_left_Ex = thm"preal_less_add_left_Ex";
+val less_add_left = thm"less_add_left";
 val preal_add_less2_mono1 = thm"preal_add_less2_mono1";
 val preal_add_less2_mono2 = thm"preal_add_less2_mono2";
-val preal_mult_less_mono1 = thm"preal_mult_less_mono1";
-val preal_mult_left_less_mono1 = thm"preal_mult_left_less_mono1";
-val preal_mult_left_le_mono1 = thm"preal_mult_left_le_mono1";
-val preal_mult_le_mono1 = thm"preal_mult_le_mono1";
-val preal_add_left_le_mono1 = thm"preal_add_left_le_mono1";
-val preal_add_le_mono1 = thm"preal_add_le_mono1";
 val preal_add_right_less_cancel = thm"preal_add_right_less_cancel";
 val preal_add_left_less_cancel = thm"preal_add_left_less_cancel";
-val preal_add_less_iff1 = thm"preal_add_less_iff1";
-val preal_add_less_iff2 = thm"preal_add_less_iff2";
-val preal_add_less_mono = thm"preal_add_less_mono";
-val preal_mult_less_mono = thm"preal_mult_less_mono";
 val preal_add_right_cancel = thm"preal_add_right_cancel";
 val preal_add_left_cancel = thm"preal_add_left_cancel";
 val preal_add_left_cancel_iff = thm"preal_add_left_cancel_iff";
 val preal_add_right_cancel_iff = thm"preal_add_right_cancel_iff";
-val preal_sup_mem_Ex = thm"preal_sup_mem_Ex";
-val preal_sup_set_not_empty = thm"preal_sup_set_not_empty";
-val preal_sup_not_mem_Ex = thm"preal_sup_not_mem_Ex";
-val preal_sup_not_mem_Ex1 = thm"preal_sup_not_mem_Ex1";
-val preal_sup_set_not_prat_set = thm"preal_sup_set_not_prat_set";
-val preal_sup_set_not_prat_set1 = thm"preal_sup_set_not_prat_set1";
-val preal_sup = thm"preal_sup";
-val preal_sup1 = thm"preal_sup1";
-val preal_psup_leI = thm"preal_psup_leI";
-val preal_psup_leI2 = thm"preal_psup_leI2";
-val preal_psup_leI2b = thm"preal_psup_leI2b";
-val preal_psup_leI2a = thm"preal_psup_leI2a";
+val preal_psup_le = thm"preal_psup_le";
 val psup_le_ub = thm"psup_le_ub";
-val psup_le_ub1 = thm"psup_le_ub1";
 val preal_complete = thm"preal_complete";
-val preal_of_prat_add = thm"preal_of_prat_add";
-val preal_of_prat_mult = thm"preal_of_prat_mult";
+val preal_of_rat_add = thm"preal_of_rat_add";
+val preal_of_rat_mult = thm"preal_of_rat_mult";
 
 val preal_add_ac = thms"preal_add_ac";
 val preal_mult_ac = thms"preal_mult_ac";
--- a/src/HOL/Real/RComplete.ML	Tue Jan 27 09:44:14 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,250 +0,0 @@
-(*  Title       : HOL/Real/RComplete.ML
-    ID          : $Id$
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-
-Completeness theorems for positive reals and reals.
-*) 
-
-Goal "x/2 + x/2 = (x::real)";
-by (Simp_tac 1);
-qed "real_sum_of_halves";
-
-(*---------------------------------------------------------
-       Completeness of reals: use supremum property of 
-       preal and theorems about real_preal. Theorems 
-       previously in Real.ML. 
- ---------------------------------------------------------*)
- (*a few lemmas*)
-Goal "ALL x:P. 0 < x ==> \ 
-\       ((EX x:P. y < x) = (EX X. real_of_preal X : P & \
-\                          y < real_of_preal X))";
-by (blast_tac (claset() addSDs [bspec, 
-		    real_gt_zero_preal_Ex RS iffD1]) 1);
-qed "real_sup_lemma1";
-
-Goal "[| ALL x:P. 0 < x;  a: P;   ALL x: P. x < y |] \
-\         ==> (EX X. X: {w. real_of_preal w : P}) & \
-\             (EX Y. ALL X: {w. real_of_preal w : P}. X < Y)";
-by (rtac conjI 1);
-by (blast_tac (claset() addDs [bspec, 
-                real_gt_zero_preal_Ex RS iffD1]) 1);
-by Auto_tac;
-by (dtac bspec 1 THEN assume_tac 1);
-by (ftac bspec 1  THEN assume_tac 1);
-by (dtac order_less_trans 1 THEN assume_tac 1); 
-by (dtac (real_gt_zero_preal_Ex RS iffD1) 1);
-by (Force_tac 1);     
-qed "real_sup_lemma2";
-
-(*-------------------------------------------------------------
-            Completeness of Positive Reals
- -------------------------------------------------------------*)
-
-(**
- Supremum property for the set of positive reals
- FIXME: long proof - should be improved
-**)
-
-(*Let P be a non-empty set of positive reals, with an upper bound y.
-  Then P has a least upper bound (written S).*)
-Goal "[| ALL x:P. (0::real) < x;  EX x. x: P;  EX y. ALL x: P. x<y |] \
-\     ==> (EX S. ALL y. (EX x: P. y < x) = (y < S))";
-by (res_inst_tac 
-   [("x","real_of_preal (psup({w. real_of_preal w : P}))")] exI 1);
-by (Clarify_tac 1); 
-by (case_tac "0 < ya" 1);
-by Auto_tac; 
-by (ftac real_sup_lemma2 1 THEN REPEAT (assume_tac 1));
-by (dtac (real_gt_zero_preal_Ex RS iffD1) 1);
-by (dtac (real_less_all_real2) 3);
-by Auto_tac;
-by (rtac (preal_complete RS spec RS iffD1) 1);
-by Auto_tac;
-by (ftac real_gt_preal_preal_Ex 1);
-by (Force_tac 1);   
-(* second part *)
-by (rtac (real_sup_lemma1 RS iffD2) 1 THEN assume_tac 1);
-by (auto_tac (claset() addSDs [real_less_all_real2,
-                               real_gt_zero_preal_Ex RS iffD1],
-	      simpset()));
-by (ftac real_sup_lemma2 2 THEN REPEAT (assume_tac 1));
-by (ftac real_sup_lemma2 1 THEN REPEAT (assume_tac 1));
-by (rtac (preal_complete RS spec RS iffD2 RS bexE) 1);
-by (Blast_tac 3);
-by (ALLGOALS(Blast_tac));
-qed "posreal_complete";
-
-(*--------------------------------------------------------
-   Completeness properties using isUb, isLub etc.
- -------------------------------------------------------*)
-
-Goal "[| isLub R S x; isLub R S y |] ==> x = (y::real)";
-by (ftac isLub_isUb 1);
-by (forw_inst_tac [("x","y")] isLub_isUb 1);
-by (blast_tac (claset() addSIs [real_le_anti_sym]
-                        addSDs [isLub_le_isUb]) 1);
-qed "real_isLub_unique";
-
-Goalw [setle_def,setge_def] "[| (x::real) <=* S'; S <= S' |] ==> x <=* S";
-by (Blast_tac 1);
-qed "real_order_restrict";
-
-(*----------------------------------------------------------------
-           Completeness theorem for the positive reals(again)
- ----------------------------------------------------------------*)
-
-Goal "[| ALL x: S. 0 < x; \
-\        EX x. x: S; \
-\        EX u. isUb (UNIV::real set) S u \
-\     |] ==> EX t. isLub (UNIV::real set) S t";
-by (res_inst_tac 
-    [("x","real_of_preal(psup({w. real_of_preal w : S}))")] exI 1);
-by (auto_tac (claset(), simpset() addsimps 
-    [isLub_def,leastP_def,isUb_def]));
-by (auto_tac (claset() addSIs [setleI,setgeI] 
-	         addSDs [(real_gt_zero_preal_Ex) RS iffD1],
-	      simpset()));
-by (forw_inst_tac [("x","y")] bspec 1 THEN assume_tac 1);
-by (dtac ((real_gt_zero_preal_Ex) RS iffD1) 1);
-by (auto_tac (claset(), simpset() addsimps [real_of_preal_le_iff]));
-by (rtac preal_psup_leI2a 1);
-by (forw_inst_tac [("y","real_of_preal ya")] setleD 1 THEN assume_tac 1);
-by (ftac real_ge_preal_preal_Ex 1);
-by (Step_tac 1);
-by (res_inst_tac [("x","y")] exI 1);
-by (blast_tac (claset() addSDs [setleD] addIs [real_of_preal_le_iff RS iffD1]) 1);
-by (forw_inst_tac [("x","x")] bspec 1 THEN assume_tac 1);
-by (ftac isUbD2 1);
-by (dtac ((real_gt_zero_preal_Ex) RS iffD1) 1);
-by (auto_tac (claset() addSDs [isUbD, real_ge_preal_preal_Ex],
-	      simpset() addsimps [real_of_preal_le_iff]));
-by (blast_tac (claset() addSDs [setleD] addSIs [psup_le_ub1] 
-	                addIs [real_of_preal_le_iff RS iffD1]) 1);
-qed "posreals_complete";
-
-
-(*-------------------------------
-    Lemmas
- -------------------------------*)
-Goal "ALL y : {z. EX x: P. z = x + (-xa) + 1} Int {x. 0 < x}. 0 < y";
-by Auto_tac;
-qed "real_sup_lemma3";
- 
-Goal "(xa <= S + X + (-Z)) = (xa + (-X) + Z <= (S::real))";
-by Auto_tac;
-qed "lemma_le_swap2";
-
-Goal "[| (x::real) + (-X) + 1 <= S; xa <= x |] ==> xa <= S + X + (- 1)";
-by (arith_tac 1);
-by Auto_tac;
-qed "lemma_real_complete2b";
-
-(*----------------------------------------------------------
-      reals Completeness (again!)
- ----------------------------------------------------------*)
-Goal "[| EX X. X: S;  EX Y. isUb (UNIV::real set) S Y |]  \
-\     ==> EX t. isLub (UNIV :: real set) S t";
-by (Step_tac 1);
-by (subgoal_tac "EX u. u: {z. EX x: S. z = x + (-X) + 1} \
-\                Int {x. 0 < x}" 1);
-by (subgoal_tac "isUb (UNIV::real set) ({z. EX x: S. z = x + (-X) + 1} \
-\                Int {x. 0 < x})  (Y + (-X) + 1)" 1); 
-by (cut_inst_tac [("P","S"),("xa","X")] real_sup_lemma3 1);
-by (EVERY1[forward_tac [exI RSN (3,posreals_complete)], Blast_tac, Blast_tac, 
-	   Step_tac]);
-by (res_inst_tac [("x","t + X + (- 1)")] exI 1);
-by (rtac isLubI2 1);
-by (rtac setgeI 2 THEN Step_tac 2);
-by (subgoal_tac "isUb (UNIV:: real set) ({z. EX x: S. z = x + (-X) + 1} \
-\                Int {x. 0 < x})  (y + (-X) + 1)" 2); 
-by (dres_inst_tac [("y","(y + (- X) + 1)")] isLub_le_isUb 2 
-      THEN assume_tac 2);
-by (full_simp_tac
-    (simpset() addsimps [real_diff_def, diff_le_eq RS sym] @
-                        add_ac) 2);
-by (rtac (setleI RS isUbI) 1);
-by (Step_tac 1);
-by (res_inst_tac [("R1.0","x"),("R2.0","y")] real_linear_less2 1);
-by (stac lemma_le_swap2 1);
-by (ftac isLubD2 1 THEN assume_tac 2);
-by (Step_tac 1);
-by (Blast_tac 1);
-by (arith_tac 1);
-by (stac lemma_le_swap2 1);
-by (ftac isLubD2 1 THEN assume_tac 2);
-by (Blast_tac 1);
-by (rtac lemma_real_complete2b 1);
-by (etac order_less_imp_le 2);
-by (blast_tac (claset() addSIs [isLubD2]) 1 THEN Step_tac 1);
-by (full_simp_tac (simpset() addsimps [real_add_assoc]) 1);
-by (blast_tac (claset() addDs [isUbD] addSIs [setleI RS isUbI]
-                        addIs [add_right_mono]) 1);
-by (blast_tac (claset() addDs [isUbD] addSIs [setleI RS isUbI]
-                        addIs [add_right_mono]) 1);
-by (Auto_tac);
-qed "reals_complete";
-
-(*----------------------------------------------------------------
-        Related: Archimedean property of reals
- ----------------------------------------------------------------*)
-
-Goal "0 < real (Suc n)";
-by (res_inst_tac [("y","real n")] order_le_less_trans 1); 
-by (rtac (real_of_nat_ge_zero) 1);
-by (simp_tac (simpset() addsimps [real_of_nat_Suc]) 1); 
-qed "real_of_nat_Suc_gt_zero";
-
-Goal "0 < x ==> EX n. inverse (real(Suc n)) < x";
-by (rtac ccontr 1);
-by (subgoal_tac "ALL n. x * real (Suc n) <= 1" 1);
-by (asm_full_simp_tac
-    (simpset() addsimps [linorder_not_less, inverse_eq_divide]) 2); 
-by (Clarify_tac 2);
-by (dres_inst_tac [("x","n")] spec 2); 
-by (dres_inst_tac [("c","real (Suc n)")] (mult_right_mono) 2); 
-by (rtac real_of_nat_ge_zero 2);
-by (asm_full_simp_tac (simpset()  
-	 addsimps [real_of_nat_Suc_gt_zero RS real_not_refl2 RS not_sym, 
-                   real_mult_commute]) 2); 
-by (subgoal_tac "isUb (UNIV::real set) \
-\                     {z. EX n. z = x*(real (Suc n))} 1" 1);
-by (subgoal_tac "EX X. X : {z. EX n. z = x*(real (Suc n))}" 1);
-by (dtac reals_complete 1);
-by (auto_tac (claset() addIs [isUbI,setleI],simpset()));
-by (subgoal_tac "ALL m. x*(real(Suc m)) <= t" 1);
-by (asm_full_simp_tac (simpset() addsimps 
-                       [real_of_nat_Suc, right_distrib]) 1);
-by (blast_tac (claset() addIs [isLubD2]) 2);
-by (asm_full_simp_tac
-    (simpset() addsimps [le_diff_eq RS sym, real_diff_def]) 1);
-by (subgoal_tac "isUb (UNIV::real set) \
-\                  {z. EX n. z = x*(real (Suc n))} (t + (-x))" 1);
-by (blast_tac (claset() addSIs [isUbI,setleI]) 2);
-by (dres_inst_tac [("y","t+(-x)")] isLub_le_isUb 1);
-by (auto_tac (claset(), 
-	      simpset() addsimps [real_of_nat_Suc,right_distrib]));
-qed "reals_Archimedean";
-
-(*There must be other proofs, e.g. Suc of the largest integer in the
-  cut representing x*)
-Goal "EX n. (x::real) < real (n::nat)";
-by (res_inst_tac [("R1.0","x"),("R2.0","0")] real_linear_less2 1);
-by (res_inst_tac [("x","0")] exI 1);
-by (res_inst_tac [("x","1")] exI 2);
-by (auto_tac (claset() addEs [order_less_trans],
-	      simpset() addsimps [real_of_nat_one]));
-by (ftac (positive_imp_inverse_positive RS reals_Archimedean) 1);
-by (Step_tac 1 THEN res_inst_tac [("x","Suc n")] exI 1);
-by (forw_inst_tac [("b","inverse x")] mult_strict_right_mono 1);
-by Auto_tac;  
-qed "reals_Archimedean2";
-
-Goal "0 < x ==> ALL y. EX (n::nat). y < real n * x";
-by (Step_tac 1);
-by (cut_inst_tac [("x","y*inverse(x)")] reals_Archimedean2 1);
-by (Step_tac 1);
-by (forw_inst_tac [("a","y * inverse x")] (mult_strict_right_mono) 1);
-by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,real_of_nat_def]));
-qed "reals_Archimedean3";
-
--- a/src/HOL/Real/RComplete.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Real/RComplete.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -6,5 +6,221 @@
                   reals and reals 
 *) 
 
-RComplete = Lubs + RealArith
+header{*Completeness Theorems for Positive Reals and Reals.*}
+
+theory RComplete = Lubs + RealArith:
+
+lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
+apply (simp)
+done
+
+
+subsection{*Completeness of Reals by Supremum Property of type @{typ preal}*} 
+
+ (*a few lemmas*)
+lemma real_sup_lemma1:
+     "\<forall>x \<in> P. 0 < x ==>   
+      ((\<exists>x \<in> P. y < x) = (\<exists>X. real_of_preal X \<in> P & y < real_of_preal X))"
+by (blast dest!: bspec real_gt_zero_preal_Ex [THEN iffD1])
+
+lemma real_sup_lemma2:
+     "[| \<forall>x \<in> P. 0 < x;  a \<in> P;   \<forall>x \<in> P. x < y |]  
+      ==> (\<exists>X. X\<in> {w. real_of_preal w \<in> P}) &  
+          (\<exists>Y. \<forall>X\<in> {w. real_of_preal w \<in> P}. X < Y)"
+apply (rule conjI)
+apply (blast dest: bspec real_gt_zero_preal_Ex [THEN iffD1], auto)
+apply (drule bspec, assumption)
+apply (frule bspec, assumption)
+apply (drule order_less_trans, assumption)
+apply (drule real_gt_zero_preal_Ex [THEN iffD1])
+apply (force) 
+done
+
+(*-------------------------------------------------------------
+            Completeness of Positive Reals
+ -------------------------------------------------------------*)
+
+(**
+ Supremum property for the set of positive reals
+ FIXME: long proof - should be improved
+**)
+
+(*Let P be a non-empty set of positive reals, with an upper bound y.
+  Then P has a least upper bound (written S).  
+FIXME: Can the premise be weakened to \<forall>x \<in> P. x\<le> y ??*)
+lemma posreal_complete: "[| \<forall>x \<in> P. (0::real) < x;  \<exists>x. x \<in> P;  \<exists>y. \<forall>x \<in> P. x<y |]  
+      ==> (\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S))"
+apply (rule_tac x = "real_of_preal (psup ({w. real_of_preal w \<in> P}))" in exI)
+apply clarify
+apply (case_tac "0 < ya", auto)
+apply (frule real_sup_lemma2, assumption+)
+apply (drule real_gt_zero_preal_Ex [THEN iffD1])
+apply (drule_tac [3] real_less_all_real2)
+apply (auto)
+apply (rule preal_complete [THEN iffD1])
+apply (auto intro: order_less_imp_le)
+apply (frule real_gt_preal_preal_Ex)
+apply (force)
+(* second part *)
+apply (rule real_sup_lemma1 [THEN iffD2], assumption)
+apply (auto dest!: real_less_all_real2 real_gt_zero_preal_Ex [THEN iffD1])
+apply (frule_tac [2] real_sup_lemma2)
+apply (frule real_sup_lemma2, assumption+, clarify) 
+apply (rule preal_complete [THEN iffD2, THEN bexE])
+prefer 3 apply blast
+apply (blast intro!: order_less_imp_le)+
+done
+
+(*--------------------------------------------------------
+   Completeness properties using isUb, isLub etc.
+ -------------------------------------------------------*)
+
+lemma real_isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::real)"
+apply (frule isLub_isUb)
+apply (frule_tac x = y in isLub_isUb)
+apply (blast intro!: real_le_anti_sym dest!: isLub_le_isUb)
+done
+
+lemma real_order_restrict: "[| (x::real) <=* S'; S <= S' |] ==> x <=* S"
+by (unfold setle_def setge_def, blast)
+
+(*----------------------------------------------------------------
+           Completeness theorem for the positive reals(again)
+ ----------------------------------------------------------------*)
+
+lemma posreals_complete:
+     "[| \<forall>x \<in>S. 0 < x;  
+         \<exists>x. x \<in>S;  
+         \<exists>u. isUb (UNIV::real set) S u  
+      |] ==> \<exists>t. isLub (UNIV::real set) S t"
+apply (rule_tac x = "real_of_preal (psup ({w. real_of_preal w \<in> S}))" in exI)
+apply (auto simp add: isLub_def leastP_def isUb_def)
+apply (auto intro!: setleI setgeI dest!: real_gt_zero_preal_Ex [THEN iffD1])
+apply (frule_tac x = y in bspec, assumption)
+apply (drule real_gt_zero_preal_Ex [THEN iffD1])
+apply (auto simp add: real_of_preal_le_iff)
+apply (frule_tac y = "real_of_preal ya" in setleD, assumption)
+apply (frule real_ge_preal_preal_Ex, safe)
+apply (blast intro!: preal_psup_le dest!: setleD intro: real_of_preal_le_iff [THEN iffD1])
+apply (frule_tac x = x in bspec, assumption)
+apply (frule isUbD2)
+apply (drule real_gt_zero_preal_Ex [THEN iffD1])
+apply (auto dest!: isUbD real_ge_preal_preal_Ex simp add: real_of_preal_le_iff)
+apply (blast dest!: setleD intro!: psup_le_ub intro: real_of_preal_le_iff [THEN iffD1])
+done
+
 
+(*-------------------------------
+    Lemmas
+ -------------------------------*)
+lemma real_sup_lemma3: "\<forall>y \<in> {z. \<exists>x \<in> P. z = x + (-xa) + 1} Int {x. 0 < x}. 0 < y"
+by auto
+ 
+lemma lemma_le_swap2: "(xa <= S + X + (-Z)) = (xa + (-X) + Z <= (S::real))"
+by auto
+
+lemma lemma_real_complete2b: "[| (x::real) + (-X) + 1 <= S; xa <= x |] ==> xa <= S + X + (- 1)"
+by arith
+
+(*----------------------------------------------------------
+      reals Completeness (again!)
+ ----------------------------------------------------------*)
+lemma reals_complete: "[| \<exists>X. X \<in>S;  \<exists>Y. isUb (UNIV::real set) S Y |]   
+      ==> \<exists>t. isLub (UNIV :: real set) S t"
+apply safe
+apply (subgoal_tac "\<exists>u. u\<in> {z. \<exists>x \<in>S. z = x + (-X) + 1} Int {x. 0 < x}")
+apply (subgoal_tac "isUb (UNIV::real set) ({z. \<exists>x \<in>S. z = x + (-X) + 1} Int {x. 0 < x}) (Y + (-X) + 1) ")
+apply (cut_tac P = S and xa = X in real_sup_lemma3)
+apply (frule posreals_complete [OF _ _ exI], blast, blast) 
+apply safe
+apply (rule_tac x = "t + X + (- 1) " in exI)
+apply (rule isLubI2)
+apply (rule_tac [2] setgeI, safe)
+apply (subgoal_tac [2] "isUb (UNIV:: real set) ({z. \<exists>x \<in>S. z = x + (-X) + 1} Int {x. 0 < x}) (y + (-X) + 1) ")
+apply (drule_tac [2] y = " (y + (- X) + 1) " in isLub_le_isUb)
+ prefer 2 apply assumption
+ prefer 2
+apply arith
+apply (rule setleI [THEN isUbI], safe)
+apply (rule_tac x = x and y = y in linorder_cases)
+apply (subst lemma_le_swap2)
+apply (frule isLubD2)
+ prefer 2 apply assumption
+apply safe
+apply blast
+apply arith
+apply (subst lemma_le_swap2)
+apply (frule isLubD2)
+ prefer 2 apply assumption
+apply blast
+apply (rule lemma_real_complete2b)
+apply (erule_tac [2] order_less_imp_le)
+apply (blast intro!: isLubD2, blast) 
+apply (simp (no_asm_use) add: real_add_assoc)
+apply (blast dest: isUbD intro!: setleI [THEN isUbI] intro: add_right_mono)
+apply (blast dest: isUbD intro!: setleI [THEN isUbI] intro: add_right_mono, auto)
+done
+
+
+subsection{*Corollary: the Archimedean Property of the Reals*}
+
+lemma reals_Archimedean: "0 < x ==> \<exists>n. inverse (real(Suc n)) < x"
+apply (rule ccontr)
+apply (subgoal_tac "\<forall>n. x * real (Suc n) <= 1")
+ prefer 2
+apply (simp add: linorder_not_less inverse_eq_divide, clarify) 
+apply (drule_tac x = n in spec)
+apply (drule_tac c = "real (Suc n)"  in mult_right_mono)
+apply (rule real_of_nat_ge_zero)
+apply (simp add: real_of_nat_Suc_gt_zero [THEN real_not_refl2, THEN not_sym] real_mult_commute)
+apply (subgoal_tac "isUb (UNIV::real set) {z. \<exists>n. z = x* (real (Suc n))} 1")
+apply (subgoal_tac "\<exists>X. X \<in> {z. \<exists>n. z = x* (real (Suc n))}")
+apply (drule reals_complete)
+apply (auto intro: isUbI setleI)
+apply (subgoal_tac "\<forall>m. x* (real (Suc m)) <= t")
+apply (simp add: real_of_nat_Suc right_distrib)
+prefer 2 apply (blast intro: isLubD2)
+apply (simp add: le_diff_eq [symmetric] real_diff_def)
+apply (subgoal_tac "isUb (UNIV::real set) {z. \<exists>n. z = x* (real (Suc n))} (t + (-x))")
+prefer 2 apply (blast intro!: isUbI setleI)
+apply (drule_tac y = "t+ (-x) " in isLub_le_isUb)
+apply (auto simp add: real_of_nat_Suc right_distrib)
+done
+
+(*There must be other proofs, e.g. Suc of the largest integer in the
+  cut representing x*)
+lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"
+apply (rule_tac x = x and y = 0 in linorder_cases)
+apply (rule_tac x = 0 in exI)
+apply (rule_tac [2] x = 1 in exI)
+apply (auto elim: order_less_trans simp add: real_of_nat_one)
+apply (frule positive_imp_inverse_positive [THEN reals_Archimedean], safe)
+apply (rule_tac x = "Suc n" in exI)
+apply (frule_tac b = "inverse x" in mult_strict_right_mono, auto)
+done
+
+lemma reals_Archimedean3: "0 < x ==> \<forall>y. \<exists>(n::nat). y < real n * x"
+apply safe
+apply (cut_tac x = "y*inverse (x) " in reals_Archimedean2)
+apply safe
+apply (frule_tac a = "y * inverse x" in mult_strict_right_mono)
+apply (auto simp add: mult_assoc real_of_nat_def)
+done
+
+ML
+{*
+val real_sum_of_halves = thm "real_sum_of_halves";
+val posreal_complete = thm "posreal_complete";
+val real_isLub_unique = thm "real_isLub_unique";
+val real_order_restrict = thm "real_order_restrict";
+val posreals_complete = thm "posreals_complete";
+val reals_complete = thm "reals_complete";
+val reals_Archimedean = thm "reals_Archimedean";
+val reals_Archimedean2 = thm "reals_Archimedean2";
+val reals_Archimedean3 = thm "reals_Archimedean3";
+*}
+
+end
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/RatArith.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -0,0 +1,162 @@
+(*  Title:      HOL/RatArith.thy
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   2004  University of Cambridge
+
+Binary arithmetic and simplification for the rats
+
+This case is reduced to that for the integers
+*)
+
+theory RatArith = Rational
+files ("rat_arith.ML"):
+
+instance rat :: number ..
+
+defs
+  rat_number_of_def:
+    "number_of v == Fract (number_of v) 1"
+     (*::bin=>rat         ::bin=>int*)
+
+theorem number_of_rat: "number_of b = rat (number_of b)"
+  by (simp add: rat_number_of_def rat_def)
+
+declare number_of_rat [symmetric, simp]
+
+lemma rat_numeral_0_eq_0: "Numeral0 = (0::rat)"
+by (simp add: rat_number_of_def zero_rat [symmetric])
+
+lemma rat_numeral_1_eq_1: "Numeral1 = (1::rat)"
+by (simp add: rat_number_of_def one_rat [symmetric])
+
+
+subsection{*Arithmetic Operations On Numerals*}
+
+lemma add_rat_number_of [simp]:
+     "(number_of v :: rat) + number_of v' = number_of (bin_add v v')"
+by (simp add: rat_number_of_def add_rat)
+
+lemma minus_rat_number_of [simp]:
+     "- (number_of w :: rat) = number_of (bin_minus w)"
+by (simp add: rat_number_of_def minus_rat)
+
+lemma diff_rat_number_of [simp]: 
+   "(number_of v :: rat) - number_of w = number_of (bin_add v (bin_minus w))"
+by (simp add: rat_number_of_def diff_rat)
+
+lemma mult_rat_number_of [simp]:
+     "(number_of v :: rat) * number_of v' = number_of (bin_mult v v')"
+by (simp add: rat_number_of_def mult_rat)
+
+text{*Lemmas for specialist use, NOT as default simprules*}
+lemma rat_mult_2: "2 * z = (z+z::rat)"
+proof -
+  have eq: "(2::rat) = 1 + 1" by (simp add: rat_numeral_1_eq_1 [symmetric])
+  thus ?thesis by (simp add: eq left_distrib)
+qed
+
+lemma rat_mult_2_right: "z * 2 = (z+z::rat)"
+by (subst mult_commute, rule rat_mult_2)
+
+
+subsection{*Comparisons On Numerals*}
+
+lemma eq_rat_number_of [simp]:
+     "((number_of v :: rat) = number_of v') =  
+      iszero (number_of (bin_add v (bin_minus v')))"
+by (simp add: rat_number_of_def eq_rat)
+
+text{*@{term neg} is used in rewrite rules for binary comparisons*}
+lemma less_rat_number_of [simp]:
+     "((number_of v :: rat) < number_of v') =  
+      neg (number_of (bin_add v (bin_minus v')))"
+by (simp add: rat_number_of_def less_rat)
+
+
+text{*New versions of existing theorems involving 0, 1*}
+
+lemma rat_minus_1_eq_m1 [simp]: "- 1 = (-1::rat)"
+by (simp add: rat_numeral_1_eq_1 [symmetric])
+
+lemma rat_mult_minus1 [simp]: "-1 * z = -(z::rat)"
+proof -
+  have  "-1 * z = (- 1) * z" by (simp add: rat_minus_1_eq_m1)
+  also have "... = - (1 * z)" by (simp only: minus_mult_left) 
+  also have "... = -z" by simp
+  finally show ?thesis .
+qed
+
+lemma rat_mult_minus1_right [simp]: "z * -1 = -(z::rat)"
+by (subst mult_commute, rule rat_mult_minus1)
+
+
+subsection{*Simplification of Arithmetic when Nested to the Right*}
+
+lemma rat_add_number_of_left [simp]:
+     "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::rat)"
+by (simp add: add_assoc [symmetric])
+
+lemma rat_mult_number_of_left [simp]:
+     "number_of v * (number_of w * z) = (number_of(bin_mult v w) * z::rat)"
+apply (simp add: mult_assoc [symmetric])
+done
+
+lemma rat_add_number_of_diff1 [simp]: 
+     "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::rat)"
+apply (unfold diff_rat_def)
+apply (rule rat_add_number_of_left)
+done
+
+lemma rat_add_number_of_diff2 [simp]:
+     "number_of v + (c - number_of w) =  
+      number_of (bin_add v (bin_minus w)) + (c::rat)"
+apply (subst diff_rat_number_of [symmetric])
+apply (simp only: diff_rat_def add_ac)
+done
+
+
+declare rat_numeral_0_eq_0 [simp] rat_numeral_1_eq_1 [simp]
+
+lemmas rat_add_0_left = add_0 [where ?'a = rat]
+lemmas rat_add_0_right = add_0_right [where ?'a = rat]
+lemmas rat_mult_1_left = mult_1 [where ?'a = rat]
+lemmas rat_mult_1_right = mult_1_right [where ?'a = rat]
+
+
+declare diff_rat_def [symmetric]
+
+
+use "rat_arith.ML"
+
+setup rat_arith_setup
+
+
+subsubsection{*Division By @{term "-1"}*}
+
+lemma rat_divide_minus1 [simp]: "x/-1 = -(x::rat)" 
+by simp
+
+lemma rat_minus1_divide [simp]: "-1/(x::rat) = - (1/x)"
+by (simp add: divide_rat_def inverse_minus_eq)
+
+subsection{*Absolute Value Function for the Rats*}
+
+lemma abs_nat_number_of [simp]: 
+     "abs (number_of v :: rat) =  
+        (if neg (number_of v) then number_of (bin_minus v)  
+         else number_of v)"
+by (simp add: abs_if)
+
+lemma abs_minus_one [simp]: "abs (-1) = (1::rat)"
+by (simp add: abs_if)
+
+
+ML
+{*
+val rat_divide_minus1 = thm "rat_divide_minus1";
+val rat_minus1_divide = thm "rat_minus1_divide";
+val abs_nat_number_of = thm "abs_nat_number_of";
+val abs_minus_one = thm "abs_minus_one";
+*}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/Rational.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -0,0 +1,736 @@
+(*  Title: HOL/Library/Rational.thy
+    ID:    $Id$
+    Author: Markus Wenzel, TU Muenchen
+    License: GPL (GNU GENERAL PUBLIC LICENSE)
+*)
+
+header {*
+  \title{Rational numbers}
+  \author{Markus Wenzel}
+*}
+
+theory Rational = Quotient + Ring_and_Field:
+
+subsection {* Fractions *}
+
+subsubsection {* The type of fractions *}
+
+typedef fraction = "{(a, b) :: int \<times> int | a b. b \<noteq> 0}"
+proof
+  show "(0, 1) \<in> ?fraction" by simp
+qed
+
+constdefs
+  fract :: "int => int => fraction"
+  "fract a b == Abs_fraction (a, b)"
+  num :: "fraction => int"
+  "num Q == fst (Rep_fraction Q)"
+  den :: "fraction => int"
+  "den Q == snd (Rep_fraction Q)"
+
+lemma fract_num [simp]: "b \<noteq> 0 ==> num (fract a b) = a"
+  by (simp add: fract_def num_def fraction_def Abs_fraction_inverse)
+
+lemma fract_den [simp]: "b \<noteq> 0 ==> den (fract a b) = b"
+  by (simp add: fract_def den_def fraction_def Abs_fraction_inverse)
+
+lemma fraction_cases [case_names fract, cases type: fraction]:
+  "(!!a b. Q = fract a b ==> b \<noteq> 0 ==> C) ==> C"
+proof -
+  assume r: "!!a b. Q = fract a b ==> b \<noteq> 0 ==> C"
+  obtain a b where "Q = fract a b" and "b \<noteq> 0"
+    by (cases Q) (auto simp add: fract_def fraction_def)
+  thus C by (rule r)
+qed
+
+lemma fraction_induct [case_names fract, induct type: fraction]:
+    "(!!a b. b \<noteq> 0 ==> P (fract a b)) ==> P Q"
+  by (cases Q) simp
+
+
+subsubsection {* Equivalence of fractions *}
+
+instance fraction :: eqv ..
+
+defs (overloaded)
+  equiv_fraction_def: "Q \<sim> R == num Q * den R = num R * den Q"
+
+lemma equiv_fraction_iff [iff]:
+    "b \<noteq> 0 ==> b' \<noteq> 0 ==> (fract a b \<sim> fract a' b') = (a * b' = a' * b)"
+  by (simp add: equiv_fraction_def)
+
+instance fraction :: equiv
+proof
+  fix Q R S :: fraction
+  {
+    show "Q \<sim> Q"
+    proof (induct Q)
+      fix a b :: int
+      assume "b \<noteq> 0" and "b \<noteq> 0"
+      with refl show "fract a b \<sim> fract a b" ..
+    qed
+  next
+    assume "Q \<sim> R" and "R \<sim> S"
+    show "Q \<sim> S"
+    proof (insert prems, induct Q, induct R, induct S)
+      fix a b a' b' a'' b'' :: int
+      assume b: "b \<noteq> 0" and b': "b' \<noteq> 0" and b'': "b'' \<noteq> 0"
+      assume "fract a b \<sim> fract a' b'" hence eq1: "a * b' = a' * b" ..
+      assume "fract a' b' \<sim> fract a'' b''" hence eq2: "a' * b'' = a'' * b'" ..
+      have "a * b'' = a'' * b"
+      proof cases
+        assume "a' = 0"
+        with b' eq1 eq2 have "a = 0 \<and> a'' = 0" by auto
+        thus ?thesis by simp
+      next
+        assume a': "a' \<noteq> 0"
+        from eq1 eq2 have "(a * b') * (a' * b'') = (a' * b) * (a'' * b')" by simp
+        hence "(a * b'') * (a' * b') = (a'' * b) * (a' * b')" by (simp only: mult_ac)
+        with a' b' show ?thesis by simp
+      qed
+      thus "fract a b \<sim> fract a'' b''" ..
+    qed
+  next
+    show "Q \<sim> R ==> R \<sim> Q"
+    proof (induct Q, induct R)
+      fix a b a' b' :: int
+      assume b: "b \<noteq> 0" and b': "b' \<noteq> 0"
+      assume "fract a b \<sim> fract a' b'"
+      hence "a * b' = a' * b" ..
+      hence "a' * b = a * b'" ..
+      thus "fract a' b' \<sim> fract a b" ..
+    qed
+  }
+qed
+
+lemma eq_fraction_iff [iff]:
+    "b \<noteq> 0 ==> b' \<noteq> 0 ==> (\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>) = (a * b' = a' * b)"
+  by (simp add: equiv_fraction_iff quot_equality)
+
+
+subsubsection {* Operations on fractions *}
+
+text {*
+ We define the basic arithmetic operations on fractions and
+ demonstrate their ``well-definedness'', i.e.\ congruence with respect
+ to equivalence of fractions.
+*}
+
+instance fraction :: zero ..
+instance fraction :: one ..
+instance fraction :: plus ..
+instance fraction :: minus ..
+instance fraction :: times ..
+instance fraction :: inverse ..
+instance fraction :: ord ..
+
+defs (overloaded)
+  zero_fraction_def: "0 == fract 0 1"
+  one_fraction_def: "1 == fract 1 1"
+  add_fraction_def: "Q + R ==
+    fract (num Q * den R + num R * den Q) (den Q * den R)"
+  minus_fraction_def: "-Q == fract (-(num Q)) (den Q)"
+  mult_fraction_def: "Q * R == fract (num Q * num R) (den Q * den R)"
+  inverse_fraction_def: "inverse Q == fract (den Q) (num Q)"
+  le_fraction_def: "Q \<le> R ==
+    (num Q * den R) * (den Q * den R) \<le> (num R * den Q) * (den Q * den R)"
+
+lemma is_zero_fraction_iff: "b \<noteq> 0 ==> (\<lfloor>fract a b\<rfloor> = \<lfloor>0\<rfloor>) = (a = 0)"
+  by (simp add: zero_fraction_def eq_fraction_iff)
+
+theorem add_fraction_cong:
+  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>
+    ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0
+    ==> \<lfloor>fract a b + fract c d\<rfloor> = \<lfloor>fract a' b' + fract c' d'\<rfloor>"
+proof -
+  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
+  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" ..
+  assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" ..
+  have "\<lfloor>fract (a * d + c * b) (b * d)\<rfloor> = \<lfloor>fract (a' * d' + c' * b') (b' * d')\<rfloor>"
+  proof
+    show "(a * d + c * b) * (b' * d') = (a' * d' + c' * b') * (b * d)"
+      (is "?lhs = ?rhs")
+    proof -
+      have "?lhs = (a * b') * (d * d') + (c * d') * (b * b')"
+        by (simp add: int_distrib mult_ac)
+      also have "... = (a' * b) * (d * d') + (c' * d) * (b * b')"
+        by (simp only: eq1 eq2)
+      also have "... = ?rhs"
+        by (simp add: int_distrib mult_ac)
+      finally show "?lhs = ?rhs" .
+    qed
+    from neq show "b * d \<noteq> 0" by simp
+    from neq show "b' * d' \<noteq> 0" by simp
+  qed
+  with neq show ?thesis by (simp add: add_fraction_def)
+qed
+
+theorem minus_fraction_cong:
+  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> b \<noteq> 0 ==> b' \<noteq> 0
+    ==> \<lfloor>-(fract a b)\<rfloor> = \<lfloor>-(fract a' b')\<rfloor>"
+proof -
+  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"
+  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>"
+  hence "a * b' = a' * b" ..
+  hence "-a * b' = -a' * b" by simp
+  hence "\<lfloor>fract (-a) b\<rfloor> = \<lfloor>fract (-a') b'\<rfloor>" ..
+  with neq show ?thesis by (simp add: minus_fraction_def)
+qed
+
+theorem mult_fraction_cong:
+  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>
+    ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0
+    ==> \<lfloor>fract a b * fract c d\<rfloor> = \<lfloor>fract a' b' * fract c' d'\<rfloor>"
+proof -
+  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
+  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" ..
+  assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" ..
+  have "\<lfloor>fract (a * c) (b * d)\<rfloor> = \<lfloor>fract (a' * c') (b' * d')\<rfloor>"
+  proof
+    from eq1 eq2 have "(a * b') * (c * d') = (a' * b) * (c' * d)" by simp
+    thus "(a * c) * (b' * d') = (a' * c') * (b * d)" by (simp add: mult_ac)
+    from neq show "b * d \<noteq> 0" by simp
+    from neq show "b' * d' \<noteq> 0" by simp
+  qed
+  with neq show "\<lfloor>fract a b * fract c d\<rfloor> = \<lfloor>fract a' b' * fract c' d'\<rfloor>"
+    by (simp add: mult_fraction_def)
+qed
+
+theorem inverse_fraction_cong:
+  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor> ==> \<lfloor>fract a' b'\<rfloor> \<noteq> \<lfloor>0\<rfloor>
+    ==> b \<noteq> 0 ==> b' \<noteq> 0
+    ==> \<lfloor>inverse (fract a b)\<rfloor> = \<lfloor>inverse (fract a' b')\<rfloor>"
+proof -
+  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"
+  assume "\<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor>" and "\<lfloor>fract a' b'\<rfloor> \<noteq> \<lfloor>0\<rfloor>"
+  with neq obtain "a \<noteq> 0" and "a' \<noteq> 0" by (simp add: is_zero_fraction_iff)
+  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>"
+  hence "a * b' = a' * b" ..
+  hence "b * a' = b' * a" by (simp only: mult_ac)
+  hence "\<lfloor>fract b a\<rfloor> = \<lfloor>fract b' a'\<rfloor>" ..
+  with neq show ?thesis by (simp add: inverse_fraction_def)
+qed
+
+theorem le_fraction_cong:
+  "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>
+    ==> b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0
+    ==> (fract a b \<le> fract c d) = (fract a' b' \<le> fract c' d')"
+proof -
+  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
+  assume "\<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor>" hence eq1: "a * b' = a' * b" ..
+  assume "\<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor>" hence eq2: "c * d' = c' * d" ..
+
+  let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))"
+  {
+    fix a b c d x :: int assume x: "x \<noteq> 0"
+    have "?le a b c d = ?le (a * x) (b * x) c d"
+    proof -
+      from x have "0 < x * x" by (auto simp add: zero_less_mult_iff)
+      hence "?le a b c d =
+          ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))"
+        by (simp add: mult_le_cancel_right)
+      also have "... = ?le (a * x) (b * x) c d"
+        by (simp add: mult_ac)
+      finally show ?thesis .
+    qed
+  } note le_factor = this
+
+  let ?D = "b * d" and ?D' = "b' * d'"
+  from neq have D: "?D \<noteq> 0" by simp
+  from neq have "?D' \<noteq> 0" by simp
+  hence "?le a b c d = ?le (a * ?D') (b * ?D') c d"
+    by (rule le_factor)
+  also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')"
+    by (simp add: mult_ac)
+  also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')"
+    by (simp only: eq1 eq2)
+  also have "... = ?le (a' * ?D) (b' * ?D) c' d'"
+    by (simp add: mult_ac)
+  also from D have "... = ?le a' b' c' d'"
+    by (rule le_factor [symmetric])
+  finally have "?le a b c d = ?le a' b' c' d'" .
+  with neq show ?thesis by (simp add: le_fraction_def)
+qed
+
+
+subsection {* Rational numbers *}
+
+subsubsection {* The type of rational numbers *}
+
+typedef (Rat)
+  rat = "UNIV :: fraction quot set" ..
+
+lemma RatI [intro, simp]: "Q \<in> Rat"
+  by (simp add: Rat_def)
+
+constdefs
+  fraction_of :: "rat => fraction"
+  "fraction_of q == pick (Rep_Rat q)"
+  rat_of :: "fraction => rat"
+  "rat_of Q == Abs_Rat \<lfloor>Q\<rfloor>"
+
+theorem rat_of_equality [iff?]: "(rat_of Q = rat_of Q') = (\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor>)"
+  by (simp add: rat_of_def Abs_Rat_inject)
+
+lemma rat_of: "\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor> ==> rat_of Q = rat_of Q'" ..
+
+constdefs
+  Fract :: "int => int => rat"
+  "Fract a b == rat_of (fract a b)"
+
+theorem Fract_inverse: "\<lfloor>fraction_of (Fract a b)\<rfloor> = \<lfloor>fract a b\<rfloor>"
+  by (simp add: fraction_of_def rat_of_def Fract_def Abs_Rat_inverse pick_inverse)
+
+theorem Fract_equality [iff?]:
+    "(Fract a b = Fract c d) = (\<lfloor>fract a b\<rfloor> = \<lfloor>fract c d\<rfloor>)"
+  by (simp add: Fract_def rat_of_equality)
+
+theorem eq_rat:
+    "b \<noteq> 0 ==> d \<noteq> 0 ==> (Fract a b = Fract c d) = (a * d = c * b)"
+  by (simp add: Fract_equality eq_fraction_iff)
+
+theorem Rat_cases [case_names Fract, cases type: rat]:
+  "(!!a b. q = Fract a b ==> b \<noteq> 0 ==> C) ==> C"
+proof -
+  assume r: "!!a b. q = Fract a b ==> b \<noteq> 0 ==> C"
+  obtain x where "q = Abs_Rat x" by (cases q)
+  moreover obtain Q where "x = \<lfloor>Q\<rfloor>" by (cases x)
+  moreover obtain a b where "Q = fract a b" and "b \<noteq> 0" by (cases Q)
+  ultimately have "q = Fract a b" by (simp only: Fract_def rat_of_def)
+  thus ?thesis by (rule r)
+qed
+
+theorem Rat_induct [case_names Fract, induct type: rat]:
+    "(!!a b. b \<noteq> 0 ==> P (Fract a b)) ==> P q"
+  by (cases q) simp
+
+
+subsubsection {* Canonical function definitions *}
+
+text {*
+  Note that the unconditional version below is much easier to read.
+*}
+
+theorem rat_cond_function:
+  "(!!q r. P \<lfloor>fraction_of q\<rfloor> \<lfloor>fraction_of r\<rfloor> ==>
+      f q r == g (fraction_of q) (fraction_of r)) ==>
+    (!!a b a' b' c d c' d'.
+      \<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor> ==>
+      P \<lfloor>fract a b\<rfloor> \<lfloor>fract c d\<rfloor> ==> P \<lfloor>fract a' b'\<rfloor> \<lfloor>fract c' d'\<rfloor> ==>
+      b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0 ==>
+      g (fract a b) (fract c d) = g (fract a' b') (fract c' d')) ==>
+    P \<lfloor>fract a b\<rfloor> \<lfloor>fract c d\<rfloor> ==>
+      f (Fract a b) (Fract c d) = g (fract a b) (fract c d)"
+  (is "PROP ?eq ==> PROP ?cong ==> ?P ==> _")
+proof -
+  assume eq: "PROP ?eq" and cong: "PROP ?cong" and P: ?P
+  have "f (Abs_Rat \<lfloor>fract a b\<rfloor>) (Abs_Rat \<lfloor>fract c d\<rfloor>) = g (fract a b) (fract c d)"
+  proof (rule quot_cond_function)
+    fix X Y assume "P X Y"
+    with eq show "f (Abs_Rat X) (Abs_Rat Y) == g (pick X) (pick Y)"
+      by (simp add: fraction_of_def pick_inverse Abs_Rat_inverse)
+  next
+    fix Q Q' R R' :: fraction
+    show "\<lfloor>Q\<rfloor> = \<lfloor>Q'\<rfloor> ==> \<lfloor>R\<rfloor> = \<lfloor>R'\<rfloor> ==>
+        P \<lfloor>Q\<rfloor> \<lfloor>R\<rfloor> ==> P \<lfloor>Q'\<rfloor> \<lfloor>R'\<rfloor> ==> g Q R = g Q' R'"
+      by (induct Q, induct Q', induct R, induct R') (rule cong)
+  qed
+  thus ?thesis by (unfold Fract_def rat_of_def)
+qed
+
+theorem rat_function:
+  "(!!q r. f q r == g (fraction_of q) (fraction_of r)) ==>
+    (!!a b a' b' c d c' d'.
+      \<lfloor>fract a b\<rfloor> = \<lfloor>fract a' b'\<rfloor> ==> \<lfloor>fract c d\<rfloor> = \<lfloor>fract c' d'\<rfloor> ==>
+      b \<noteq> 0 ==> b' \<noteq> 0 ==> d \<noteq> 0 ==> d' \<noteq> 0 ==>
+      g (fract a b) (fract c d) = g (fract a' b') (fract c' d')) ==>
+    f (Fract a b) (Fract c d) = g (fract a b) (fract c d)"
+proof -
+  case rule_context from this TrueI
+  show ?thesis by (rule rat_cond_function)
+qed
+
+
+subsubsection {* Standard operations on rational numbers *}
+
+instance rat :: zero ..
+instance rat :: one ..
+instance rat :: plus ..
+instance rat :: minus ..
+instance rat :: times ..
+instance rat :: inverse ..
+instance rat :: ord ..
+
+defs (overloaded)
+  zero_rat_def: "0 == rat_of 0"
+  one_rat_def: "1 == rat_of 1"
+  add_rat_def: "q + r == rat_of (fraction_of q + fraction_of r)"
+  minus_rat_def: "-q == rat_of (-(fraction_of q))"
+  diff_rat_def: "q - r == q + (-(r::rat))"
+  mult_rat_def: "q * r == rat_of (fraction_of q * fraction_of r)"
+  inverse_rat_def: "inverse q == 
+                    if q=0 then 0 else rat_of (inverse (fraction_of q))"
+  divide_rat_def: "q / r == q * inverse (r::rat)"
+  le_rat_def: "q \<le> r == fraction_of q \<le> fraction_of r"
+  less_rat_def: "q < r == q \<le> r \<and> q \<noteq> (r::rat)"
+  abs_rat_def: "\<bar>q\<bar> == if q < 0 then -q else (q::rat)"
+
+theorem zero_rat: "0 = Fract 0 1"
+  by (simp add: zero_rat_def zero_fraction_def rat_of_def Fract_def)        
+
+theorem one_rat: "1 = Fract 1 1"
+  by (simp add: one_rat_def one_fraction_def rat_of_def Fract_def)
+
+theorem add_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
+  Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
+proof -
+  have "Fract a b + Fract c d = rat_of (fract a b + fract c d)"
+    by (rule rat_function, rule add_rat_def, rule rat_of, rule add_fraction_cong)
+  also
+  assume "b \<noteq> 0"  "d \<noteq> 0"
+  hence "fract a b + fract c d = fract (a * d + c * b) (b * d)"
+    by (simp add: add_fraction_def)
+  finally show ?thesis by (unfold Fract_def)
+qed
+
+theorem minus_rat: "b \<noteq> 0 ==> -(Fract a b) = Fract (-a) b"
+proof -
+  have "-(Fract a b) = rat_of (-(fract a b))"
+    by (rule rat_function, rule minus_rat_def, rule rat_of, rule minus_fraction_cong)
+  also assume "b \<noteq> 0" hence "-(fract a b) = fract (-a) b"
+    by (simp add: minus_fraction_def)
+  finally show ?thesis by (unfold Fract_def)
+qed
+
+theorem diff_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
+    Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
+  by (simp add: diff_rat_def add_rat minus_rat)
+
+theorem mult_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
+  Fract a b * Fract c d = Fract (a * c) (b * d)"
+proof -
+  have "Fract a b * Fract c d = rat_of (fract a b * fract c d)"
+    by (rule rat_function, rule mult_rat_def, rule rat_of, rule mult_fraction_cong)
+  also
+  assume "b \<noteq> 0"  "d \<noteq> 0"
+  hence "fract a b * fract c d = fract (a * c) (b * d)"
+    by (simp add: mult_fraction_def)
+  finally show ?thesis by (unfold Fract_def)
+qed
+
+theorem inverse_rat: "Fract a b \<noteq> 0 ==> b \<noteq> 0 ==>
+  inverse (Fract a b) = Fract b a"
+proof -
+  assume neq: "b \<noteq> 0" and nonzero: "Fract a b \<noteq> 0"
+  hence "\<lfloor>fract a b\<rfloor> \<noteq> \<lfloor>0\<rfloor>"
+    by (simp add: zero_rat eq_rat is_zero_fraction_iff)
+  with _ inverse_fraction_cong [THEN rat_of]
+  have "inverse (Fract a b) = rat_of (inverse (fract a b))"
+  proof (rule rat_cond_function)
+    fix q assume cond: "\<lfloor>fraction_of q\<rfloor> \<noteq> \<lfloor>0\<rfloor>"
+    have "q \<noteq> 0"
+    proof (cases q)
+      fix a b assume "b \<noteq> 0" and "q = Fract a b"
+      from this cond show ?thesis
+        by (simp add: Fract_inverse is_zero_fraction_iff zero_rat eq_rat)
+    qed
+    thus "inverse q == rat_of (inverse (fraction_of q))"
+      by (simp add: inverse_rat_def)
+  qed
+  also from neq nonzero have "inverse (fract a b) = fract b a"
+    by (simp add: inverse_fraction_def)
+  finally show ?thesis by (unfold Fract_def)
+qed
+
+theorem divide_rat: "Fract c d \<noteq> 0 ==> b \<noteq> 0 ==> d \<noteq> 0 ==>
+  Fract a b / Fract c d = Fract (a * d) (b * c)"
+proof -
+  assume neq: "b \<noteq> 0"  "d \<noteq> 0" and nonzero: "Fract c d \<noteq> 0"
+  hence "c \<noteq> 0" by (simp add: zero_rat eq_rat)
+  with neq nonzero show ?thesis
+    by (simp add: divide_rat_def inverse_rat mult_rat)
+qed
+
+theorem le_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
+  (Fract a b \<le> Fract c d) = ((a * d) * (b * d) \<le> (c * b) * (b * d))"
+proof -
+  have "(Fract a b \<le> Fract c d) = (fract a b \<le> fract c d)"
+    by (rule rat_function, rule le_rat_def, rule le_fraction_cong)
+  also
+  assume "b \<noteq> 0"  "d \<noteq> 0"
+  hence "(fract a b \<le> fract c d) = ((a * d) * (b * d) \<le> (c * b) * (b * d))"
+    by (simp add: le_fraction_def)
+  finally show ?thesis .
+qed
+
+theorem less_rat: "b \<noteq> 0 ==> d \<noteq> 0 ==>
+    (Fract a b < Fract c d) = ((a * d) * (b * d) < (c * b) * (b * d))"
+  by (simp add: less_rat_def le_rat eq_rat int_less_le)
+
+theorem abs_rat: "b \<noteq> 0 ==> \<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
+  by (simp add: abs_rat_def minus_rat zero_rat less_rat eq_rat)
+     (auto simp add: mult_less_0_iff zero_less_mult_iff int_le_less 
+                split: abs_split)
+
+
+subsubsection {* The ordered field of rational numbers *}
+
+lemma rat_add_assoc: "(q + r) + s = q + (r + (s::rat))"
+  by (induct q, induct r, induct s) 
+     (simp add: add_rat add_ac mult_ac int_distrib)
+
+lemma rat_add_0: "0 + q = (q::rat)"
+  by (induct q) (simp add: zero_rat add_rat)
+
+lemma rat_left_minus: "(-q) + q = (0::rat)"
+  by (induct q) (simp add: zero_rat minus_rat add_rat eq_rat)
+
+
+instance rat :: field
+proof
+  fix q r s :: rat
+  show "(q + r) + s = q + (r + s)"
+    by (rule rat_add_assoc)
+  show "q + r = r + q"
+    by (induct q, induct r) (simp add: add_rat add_ac mult_ac)
+  show "0 + q = q"
+    by (induct q) (simp add: zero_rat add_rat)
+  show "(-q) + q = 0"
+    by (rule rat_left_minus)
+  show "q - r = q + (-r)"
+    by (induct q, induct r) (simp add: add_rat minus_rat diff_rat)
+  show "(q * r) * s = q * (r * s)"
+    by (induct q, induct r, induct s) (simp add: mult_rat mult_ac)
+  show "q * r = r * q"
+    by (induct q, induct r) (simp add: mult_rat mult_ac)
+  show "1 * q = q"
+    by (induct q) (simp add: one_rat mult_rat)
+  show "(q + r) * s = q * s + r * s"
+    by (induct q, induct r, induct s) 
+       (simp add: add_rat mult_rat eq_rat int_distrib)
+  show "q \<noteq> 0 ==> inverse q * q = 1"
+    by (induct q) (simp add: inverse_rat mult_rat one_rat zero_rat eq_rat)
+  show "r \<noteq> 0 ==> q / r = q * inverse r"
+    by (induct q, induct r)
+       (simp add: mult_rat divide_rat inverse_rat zero_rat eq_rat)
+  show "0 \<noteq> (1::rat)"
+    by (simp add: zero_rat one_rat eq_rat) 
+  assume eq: "s+q = s+r" 
+    hence "(-s + s) + q = (-s + s) + r" by (simp only: eq rat_add_assoc)
+    thus "q = r" by (simp add: rat_left_minus rat_add_0)
+qed
+
+instance rat :: linorder
+proof
+  fix q r s :: rat
+  {
+    assume "q \<le> r" and "r \<le> s"
+    show "q \<le> s"
+    proof (insert prems, induct q, induct r, induct s)
+      fix a b c d e f :: int
+      assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
+      assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract e f"
+      show "Fract a b \<le> Fract e f"
+      proof -
+        from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f"
+          by (auto simp add: zero_less_mult_iff linorder_neq_iff)
+        have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)"
+        proof -
+          from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
+            by (simp add: le_rat)
+          with ff show ?thesis by (simp add: mult_le_cancel_right)
+        qed
+        also have "... = (c * f) * (d * f) * (b * b)"
+          by (simp only: mult_ac)
+        also have "... \<le> (e * d) * (d * f) * (b * b)"
+        proof -
+          from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)"
+            by (simp add: le_rat)
+          with bb show ?thesis by (simp add: mult_le_cancel_right)
+        qed
+        finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)"
+          by (simp only: mult_ac)
+        with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)"
+          by (simp add: mult_le_cancel_right)
+        with neq show ?thesis by (simp add: le_rat)
+      qed
+    qed
+  next
+    assume "q \<le> r" and "r \<le> q"
+    show "q = r"
+    proof (insert prems, induct q, induct r)
+      fix a b c d :: int
+      assume neq: "b \<noteq> 0"  "d \<noteq> 0"
+      assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract a b"
+      show "Fract a b = Fract c d"
+      proof -
+        from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
+          by (simp add: le_rat)
+        also have "... \<le> (a * d) * (b * d)"
+        proof -
+          from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)"
+            by (simp add: le_rat)
+          thus ?thesis by (simp only: mult_ac)
+        qed
+        finally have "(a * d) * (b * d) = (c * b) * (b * d)" .
+        moreover from neq have "b * d \<noteq> 0" by simp
+        ultimately have "a * d = c * b" by simp
+        with neq show ?thesis by (simp add: eq_rat)
+      qed
+    qed
+  next
+    show "q \<le> q"
+      by (induct q) (simp add: le_rat)
+    show "(q < r) = (q \<le> r \<and> q \<noteq> r)"
+      by (simp only: less_rat_def)
+    show "q \<le> r \<or> r \<le> q"
+      by (induct q, induct r) (simp add: le_rat mult_ac, arith)
+  }
+qed
+
+instance rat :: ordered_field
+proof
+  fix q r s :: rat
+  show "0 < (1::rat)" 
+    by (simp add: zero_rat one_rat less_rat) 
+  show "q \<le> r ==> s + q \<le> s + r"
+  proof (induct q, induct r, induct s)
+    fix a b c d e f :: int
+    assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
+    assume le: "Fract a b \<le> Fract c d"
+    show "Fract e f + Fract a b \<le> Fract e f + Fract c d"
+    proof -
+      let ?F = "f * f" from neq have F: "0 < ?F"
+        by (auto simp add: zero_less_mult_iff)
+      from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
+        by (simp add: le_rat)
+      with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F"
+        by (simp add: mult_le_cancel_right)
+      with neq show ?thesis by (simp add: add_rat le_rat mult_ac int_distrib)
+    qed
+  qed
+  show "q < r ==> 0 < s ==> s * q < s * r"
+  proof (induct q, induct r, induct s)
+    fix a b c d e f :: int
+    assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
+    assume le: "Fract a b < Fract c d"
+    assume gt: "0 < Fract e f"
+    show "Fract e f * Fract a b < Fract e f * Fract c d"
+    proof -
+      let ?E = "e * f" and ?F = "f * f"
+      from neq gt have "0 < ?E"
+        by (auto simp add: zero_rat less_rat le_rat int_less_le eq_rat)
+      moreover from neq have "0 < ?F"
+        by (auto simp add: zero_less_mult_iff)
+      moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)"
+        by (simp add: less_rat)
+      ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F"
+        by (simp add: mult_less_cancel_right)
+      with neq show ?thesis
+        by (simp add: less_rat mult_rat mult_ac)
+    qed
+  qed
+  show "\<bar>q\<bar> = (if q < 0 then -q else q)"
+    by (simp only: abs_rat_def)
+qed
+
+instance rat :: division_by_zero
+proof
+  fix x :: rat
+  show "inverse 0 = (0::rat)"  by (simp add: inverse_rat_def)
+  show "x/0 = 0"   by (simp add: divide_rat_def inverse_rat_def)
+qed
+
+
+subsection {* Embedding integers: The Injection @{term rat} *}
+
+constdefs
+  rat :: "int => rat"    (* FIXME generalize int to any numeric subtype (?) *)
+  "rat z == Fract z 1"
+  int_set :: "rat set"    ("\<int>")    (* FIXME generalize rat to any numeric supertype (?) *)
+  "\<int> == range rat"
+
+lemma int_set_cases [case_names rat, cases set: int_set]:
+  "q \<in> \<int> ==> (!!z. q = rat z ==> C) ==> C"
+proof (unfold int_set_def)
+  assume "!!z. q = rat z ==> C"
+  assume "q \<in> range rat" thus C ..
+qed
+
+lemma int_set_induct [case_names rat, induct set: int_set]:
+  "q \<in> \<int> ==> (!!z. P (rat z)) ==> P q"
+  by (rule int_set_cases) auto
+
+lemma rat_of_int_zero [simp]: "rat (0::int) = (0::rat)"
+by (simp add: rat_def zero_rat [symmetric])
+
+lemma rat_of_int_one [simp]: "rat (1::int) = (1::rat)"
+by (simp add: rat_def one_rat [symmetric])
+
+lemma rat_of_int_add_distrib [simp]: "rat (x + y) = rat (x::int) + rat y"
+by (simp add: rat_def add_rat)
+
+lemma rat_of_int_minus_distrib [simp]: "rat (-x) = -rat (x::int)"
+by (simp add: rat_def minus_rat)
+
+lemma rat_of_int_diff_distrib [simp]: "rat (x - y) = rat (x::int) - rat y"
+by (simp add: rat_def diff_rat)
+
+lemma rat_of_int_mult_distrib [simp]: "rat (x * y) = rat (x::int) * rat y"
+by (simp add: rat_def mult_rat)
+
+lemma rat_inject [simp]: "(rat z = rat w) = (z = w)"
+proof
+  assume "rat z = rat w"
+  hence "Fract z 1 = Fract w 1" by (unfold rat_def)
+  hence "\<lfloor>fract z 1\<rfloor> = \<lfloor>fract w 1\<rfloor>" ..
+  thus "z = w" by auto
+next
+  assume "z = w"
+  thus "rat z = rat w" by simp
+qed
+
+
+lemma rat_of_int_zero_cancel [simp]: "(rat x = 0) = (x = 0)"
+proof -
+  have "(rat x = 0) = (rat x = rat 0)" by simp
+  also have "... = (x = 0)" by (rule rat_inject)
+  finally show ?thesis .
+qed
+
+lemma rat_of_int_less_iff [simp]: "rat (x::int) < rat y = (x < y)"
+by (simp add: rat_def less_rat) 
+
+lemma rat_of_int_le_iff [simp]: "(rat (x::int) \<le> rat y) = (x \<le> y)"
+by (simp add: linorder_not_less [symmetric])
+
+lemma zero_less_rat_of_int_iff [simp]: "(0 < rat y) = (0 < y)"
+by (insert rat_of_int_less_iff [of 0 y], simp)
+
+
+subsection {* Various Other Results *}
+
+lemma minus_rat_cancel [simp]: "b \<noteq> 0 ==> Fract (-a) (-b) = Fract a b"
+by (simp add: Fract_equality eq_fraction_iff) 
+
+theorem Rat_induct_pos [case_names Fract, induct type: rat]:
+  assumes step: "!!a b. 0 < b ==> P (Fract a b)"
+    shows "P q"
+proof (cases q)
+  have step': "!!a b. b < 0 ==> P (Fract a b)"
+  proof -
+    fix a::int and b::int
+    assume b: "b < 0"
+    hence "0 < -b" by simp
+    hence "P (Fract (-a) (-b))" by (rule step)
+    thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b])
+  qed
+  case (Fract a b)
+  thus "P q" by (force simp add: linorder_neq_iff step step')
+qed
+
+lemma zero_less_Fract_iff:
+     "0 < b ==> (0 < Fract a b) = (0 < a)"
+by (simp add: zero_rat less_rat order_less_imp_not_eq2 zero_less_mult_iff) 
+
+end
--- a/src/HOL/Real/RealArith.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Real/RealArith.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -1,6 +1,159 @@
-theory RealArith = RealBin
+(*  Title:      HOL/RealArith.thy
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1999  University of Cambridge
+
+Binary arithmetic and simplification for the reals
+
+This case is reduced to that for the integers
+*)
+
+theory RealArith = RealDef
 files ("real_arith.ML"):
 
+instance real :: number ..
+
+defs
+  real_number_of_def:
+    "number_of v == real (number_of v :: int)"
+     (*::bin=>real           ::bin=>int*)
+
+text{*Collapse applications of @{term real} to @{term number_of}*}
+declare real_number_of_def [symmetric, simp]
+
+lemma real_numeral_0_eq_0: "Numeral0 = (0::real)"
+by (simp add: real_number_of_def)
+
+lemma real_numeral_1_eq_1: "Numeral1 = (1::real)"
+apply (unfold real_number_of_def)
+apply (subst real_of_one [symmetric], simp)
+done
+
+
+subsection{*Arithmetic Operations On Numerals*}
+
+lemma add_real_number_of [simp]:
+     "(number_of v :: real) + number_of v' = number_of (bin_add v v')"
+by (simp only: real_number_of_def real_of_int_add number_of_add)
+
+lemma minus_real_number_of [simp]:
+     "- (number_of w :: real) = number_of (bin_minus w)"
+by (simp only: real_number_of_def number_of_minus real_of_int_minus)
+
+lemma diff_real_number_of [simp]: 
+   "(number_of v :: real) - number_of w = number_of (bin_add v (bin_minus w))"
+by (simp only: real_number_of_def diff_number_of_eq real_of_int_diff)
+
+lemma mult_real_number_of [simp]:
+     "(number_of v :: real) * number_of v' = number_of (bin_mult v v')"
+by (simp only: real_number_of_def real_of_int_mult number_of_mult)
+
+
+text{*Lemmas for specialist use, NOT as default simprules*}
+lemma real_mult_2: "2 * z = (z+z::real)"
+proof -
+  have eq: "(2::real) = 1 + 1" by (simp add: real_numeral_1_eq_1 [symmetric])
+  thus ?thesis by (simp add: eq left_distrib)
+qed
+
+lemma real_mult_2_right: "z * 2 = (z+z::real)"
+by (subst mult_commute, rule real_mult_2)
+
+
+subsection{*Comparisons On Numerals*}
+
+lemma eq_real_number_of [simp]:
+     "((number_of v :: real) = number_of v') =  
+      iszero (number_of (bin_add v (bin_minus v')))"
+by (simp only: real_number_of_def real_of_int_inject eq_number_of_eq)
+
+text{*@{term neg} is used in rewrite rules for binary comparisons*}
+lemma less_real_number_of [simp]:
+     "((number_of v :: real) < number_of v') =  
+      neg (number_of (bin_add v (bin_minus v')))"
+by (simp only: real_number_of_def real_of_int_less_iff less_number_of_eq_neg)
+
+
+text{*New versions of existing theorems involving 0, 1*}
+
+lemma real_minus_1_eq_m1 [simp]: "- 1 = (-1::real)"
+by (simp add: real_numeral_1_eq_1 [symmetric])
+
+lemma real_mult_minus1 [simp]: "-1 * z = -(z::real)"
+proof -
+  have  "-1 * z = (- 1) * z" by (simp add: real_minus_1_eq_m1)
+  also have "... = - (1 * z)" by (simp only: minus_mult_left) 
+  also have "... = -z" by simp
+  finally show ?thesis .
+qed
+
+lemma real_mult_minus1_right [simp]: "z * -1 = -(z::real)"
+by (subst mult_commute, rule real_mult_minus1)
+
+
+
+(** real from type "nat" **)
+
+lemma zero_less_real_of_nat_iff [iff]: "(0 < real (n::nat)) = (0<n)"
+by (simp only: real_of_nat_less_iff real_of_nat_zero [symmetric])
+
+lemma zero_le_real_of_nat_iff [iff]: "(0 <= real (n::nat)) = (0<=n)"
+by (simp only: real_of_nat_le_iff real_of_nat_zero [symmetric])
+
+
+(*Like the ones above, for "equals"*)
+declare real_of_nat_zero_iff [iff]
+
+
+subsection{*Simplification of Arithmetic when Nested to the Right*}
+
+lemma real_add_number_of_left [simp]:
+     "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::real)"
+by (simp add: add_assoc [symmetric])
+
+lemma real_mult_number_of_left [simp]:
+     "number_of v * (number_of w * z) = (number_of(bin_mult v w) * z::real)"
+apply (simp (no_asm) add: mult_assoc [symmetric])
+done
+
+lemma real_add_number_of_diff1 [simp]: 
+     "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::real)"
+apply (unfold real_diff_def)
+apply (rule real_add_number_of_left)
+done
+
+lemma real_add_number_of_diff2 [simp]:
+     "number_of v + (c - number_of w) =  
+      number_of (bin_add v (bin_minus w)) + (c::real)"
+apply (subst diff_real_number_of [symmetric])
+apply (simp only: real_diff_def add_ac)
+done
+
+
+text{*The constant @{term neg} is used in rewrite rules for binary
+comparisons. A complication in this proof is that both @{term real} and @{term
+number_of} are polymorphic, so that it's difficult to know what types subterms
+have. *}
+lemma real_of_nat_number_of [simp]:
+     "real (number_of v :: nat) =  
+        (if neg (number_of v) then 0  
+         else (number_of v :: real))"
+proof cases
+  assume "neg (number_of v)" thus ?thesis by simp
+next
+  assume neg: "~ neg (number_of v)"
+  thus ?thesis
+    by (simp only: nat_number_of_def real_of_nat_real_of_int [OF neg], simp) 
+qed
+
+declare real_numeral_0_eq_0 [simp] real_numeral_1_eq_1 [simp]
+
+(*Simplification of  x-y < 0, etc.*)
+declare less_iff_diff_less_0 [symmetric, iff]
+declare eq_iff_diff_eq_0 [symmetric, iff]
+declare le_iff_diff_le_0 [symmetric, iff]
+
+
 use "real_arith.ML"
 
 setup real_arith_setup
@@ -79,13 +232,7 @@
         (if neg (number_of v) then number_of (bin_minus v)  
          else number_of v)"
 by (simp add: real_abs_def bin_arith_simps minus_real_number_of
-       le_real_number_of_eq_not_less less_real_number_of real_of_int_le_iff)
-
-
-(*----------------------------------------------------------------------------
-       Properties of the absolute value function over the reals
-       (adapted version of previously proved theorems about abs)
- ----------------------------------------------------------------------------*)
+       less_real_number_of real_of_int_le_iff)
 
 text{*FIXME: these should go!*}
 lemma abs_eqI1: "(0::real)\<le>x ==> abs x = x"
@@ -95,7 +242,7 @@
 by (unfold real_abs_def, simp)
 
 lemma abs_minus_eqI2: "x < (0::real) ==> abs x = -x"
-by (unfold real_abs_def real_le_def, simp)
+by (simp add: real_abs_def linorder_not_less [symmetric])
 
 lemma abs_minus_add_cancel: "abs(x + (-y)) = abs (y + (-(x::real)))"
 by (unfold real_abs_def, simp)
--- a/src/HOL/Real/RealBin.ML	Tue Jan 27 09:44:14 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,590 +0,0 @@
-(*  Title:      HOL/Real/RealBin.ML
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1999  University of Cambridge
-
-Binary arithmetic for the reals (integer literals only).
-*)
-
-(** real (coercion from int to real) **)
-
-Goal "real (number_of w :: int) = number_of w";
-by (simp_tac (simpset() addsimps [real_number_of_def]) 1);
-qed "real_number_of";
-Addsimps [real_number_of];
-
-Goalw [real_number_of_def] "Numeral0 = (0::real)";
-by (simp_tac (simpset() addsimps [real_of_int_zero RS sym]) 1);
-qed "real_numeral_0_eq_0";
-
-Goalw [real_number_of_def] "Numeral1 = (1::real)";
-by (stac (real_of_one RS sym) 1);
-by (Simp_tac 1);
-qed "real_numeral_1_eq_1";
-
-(** Addition **)
-
-Goal "(number_of v :: real) + number_of v' = number_of (bin_add v v')";
-by (simp_tac
-    (HOL_ss addsimps [real_number_of_def,
-                                  real_of_int_add, number_of_add]) 1);
-qed "add_real_number_of";
-
-Addsimps [add_real_number_of];
-
-
-(** Subtraction **)
-
-Goalw [real_number_of_def] "- (number_of w :: real) = number_of (bin_minus w)";
-by (simp_tac
-    (HOL_ss addsimps [number_of_minus, real_of_int_minus]) 1);
-qed "minus_real_number_of";
-
-Goalw [real_number_of_def]
-   "(number_of v :: real) - number_of w = number_of (bin_add v (bin_minus w))";
-by (simp_tac
-    (HOL_ss addsimps [diff_number_of_eq, real_of_int_diff]) 1);
-qed "diff_real_number_of";
-
-Addsimps [minus_real_number_of, diff_real_number_of];
-
-
-(** Multiplication **)
-
-Goal "(number_of v :: real) * number_of v' = number_of (bin_mult v v')";
-by (simp_tac
-    (HOL_ss addsimps [real_number_of_def, real_of_int_mult,
-                      number_of_mult]) 1);
-qed "mult_real_number_of";
-
-Addsimps [mult_real_number_of];
-
-Goal "(2::real) = 1 + 1";
-by (simp_tac (simpset() addsimps [real_numeral_1_eq_1 RS sym]) 1);
-val lemma = result();
-
-(*For specialist use: NOT as default simprules*)
-Goal "2 * z = (z+z::real)";
-by (simp_tac (simpset () addsimps [lemma, left_distrib]) 1);
-qed "real_mult_2";
-
-Goal "z * 2 = (z+z::real)";
-by (stac real_mult_commute 1 THEN rtac real_mult_2 1);
-qed "real_mult_2_right";
-
-
-(*** Comparisons ***)
-
-(** Equals (=) **)
-
-Goal "((number_of v :: real) = number_of v') = \
-\     iszero (number_of (bin_add v (bin_minus v')))";
-by (simp_tac
-    (HOL_ss addsimps [real_number_of_def,
-                      real_of_int_inject, eq_number_of_eq]) 1);
-qed "eq_real_number_of";
-
-Addsimps [eq_real_number_of];
-
-(** Less-than (<) **)
-
-(*"neg" is used in rewrite rules for binary comparisons*)
-Goal "((number_of v :: real) < number_of v') = \
-\     neg (number_of (bin_add v (bin_minus v')))";
-by (simp_tac
-    (HOL_ss addsimps [real_number_of_def, real_of_int_less_iff,
-                                  less_number_of_eq_neg]) 1);
-qed "less_real_number_of";
-
-Addsimps [less_real_number_of];
-
-
-(** Less-than-or-equals (<=) **)
-
-Goal "(number_of x <= (number_of y::real)) = \
-\     (~ number_of y < (number_of x::real))";
-by (rtac (linorder_not_less RS sym) 1);
-qed "le_real_number_of_eq_not_less";
-
-Addsimps [le_real_number_of_eq_not_less];
-
-(*** New versions of existing theorems involving 0, 1 ***)
-
-Goal "- 1 = (-1::real)";
-by (simp_tac (simpset() addsimps [real_numeral_1_eq_1 RS sym]) 1);
-qed "real_minus_1_eq_m1";
-
-Goal "-1 * z = -(z::real)";
-by (simp_tac (simpset() addsimps [real_minus_1_eq_m1 RS sym]) 1);
-qed "real_mult_minus1";
-
-Goal "z * -1 = -(z::real)";
-by (stac real_mult_commute 1 THEN rtac real_mult_minus1 1);
-qed "real_mult_minus1_right";
-
-Addsimps [real_mult_minus1, real_mult_minus1_right];
-
-
-(*Maps 0 to Numeral0 and 1 to Numeral1 and -(Numeral1) to -1*)
-val real_numeral_ss =
-    HOL_ss addsimps [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym,
-                     real_minus_1_eq_m1];
-
-fun rename_numerals th =
-    asm_full_simplify real_numeral_ss (Thm.transfer (the_context ()) th);
-
-
-(** real from type "nat" **)
-
-Goal "(0 < real (n::nat)) = (0<n)";
-by (simp_tac (HOL_ss addsimps [real_of_nat_less_iff,
-                               real_of_nat_zero RS sym]) 1);
-qed "zero_less_real_of_nat_iff";
-AddIffs [zero_less_real_of_nat_iff];
-
-Goal "(0 <= real (n::nat)) = (0<=n)";
-by (simp_tac (HOL_ss addsimps [real_of_nat_le_iff,
-                               real_of_nat_zero RS sym]) 1);
-qed "zero_le_real_of_nat_iff";
-AddIffs [zero_le_real_of_nat_iff];
-
-
-(*Like the ones above, for "equals"*)
-AddIffs [real_of_nat_zero_iff];
-
-(** Simplification of arithmetic when nested to the right **)
-
-Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::real)";
-by (asm_full_simp_tac (simpset() addsimps [add_assoc RS sym]) 1); 
-qed "real_add_number_of_left";
-
-Goal "number_of v * (number_of w * z) = (number_of(bin_mult v w) * z::real)";
-by (simp_tac (simpset() addsimps [mult_assoc RS sym]) 1);
-qed "real_mult_number_of_left";
-
-Goalw [real_diff_def]
-     "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::real)";
-by (rtac real_add_number_of_left 1);
-qed "real_add_number_of_diff1";
-
-Goal "number_of v + (c - number_of w) = \
-\     number_of (bin_add v (bin_minus w)) + (c::real)";
-by (stac (diff_real_number_of RS sym) 1);
-by (asm_full_simp_tac (HOL_ss addsimps [real_diff_def]@add_ac) 1); 
-qed "real_add_number_of_diff2";
-
-Addsimps [real_add_number_of_left, real_mult_number_of_left,
-          real_add_number_of_diff1, real_add_number_of_diff2];
-
-
-(*"neg" is used in rewrite rules for binary comparisons*)
-Goal "real (number_of v :: nat) = \
-\       (if neg (number_of v) then 0 \
-\        else (number_of v :: real))";
-by (simp_tac
-    (HOL_ss addsimps [nat_number_of_def, real_of_nat_real_of_int,
-                      real_of_nat_neg_int, real_number_of,
-                      real_numeral_0_eq_0 RS sym]) 1);
-qed "real_of_nat_number_of";
-Addsimps [real_of_nat_number_of];
-
-
-(**** Simprocs for numeric literals ****)
-
-(** For combine_numerals **)
-
-Goal "i*u + (j*u + k) = (i+j)*u + (k::real)";
-by (asm_simp_tac (simpset() addsimps [left_distrib] @ add_ac) 1);
-qed "left_real_add_mult_distrib";
-
-
-(** For cancel_numerals **)
-
-val rel_iff_rel_0_rls = map (inst "b" "?u+?v")
-                   [less_iff_diff_less_0, eq_iff_diff_eq_0, le_iff_diff_le_0] @
-                 map (inst "b" "n")
-                   [less_iff_diff_less_0, eq_iff_diff_eq_0, le_iff_diff_le_0];
-
-Goal "!!i::real. (i*u + m = j*u + n) = ((i-j)*u + m = n)";
-by (asm_simp_tac (simpset() addsimps [real_diff_def, left_distrib]@
-                                     add_ac@rel_iff_rel_0_rls) 1);
-qed "real_eq_add_iff1";
-
-Goal "!!i::real. (i*u + m = j*u + n) = (m = (j-i)*u + n)";
-by (asm_simp_tac (simpset() addsimps [real_diff_def, left_distrib]@
-                                     add_ac@rel_iff_rel_0_rls) 1);
-qed "real_eq_add_iff2";
-
-Goal "!!i::real. (i*u + m < j*u + n) = ((i-j)*u + m < n)";
-by (asm_simp_tac (simpset() addsimps [real_diff_def, left_distrib]@
-                                     add_ac@rel_iff_rel_0_rls) 1);
-qed "real_less_add_iff1";
-
-Goal "!!i::real. (i*u + m < j*u + n) = (m < (j-i)*u + n)";
-by (asm_simp_tac (simpset() addsimps [real_diff_def, left_distrib]@
-                                     add_ac@rel_iff_rel_0_rls) 1);
-qed "real_less_add_iff2";
-
-Goal "!!i::real. (i*u + m <= j*u + n) = ((i-j)*u + m <= n)";
-by (asm_simp_tac (simpset() addsimps [real_diff_def, left_distrib]@
-                                     add_ac@rel_iff_rel_0_rls) 1);
-qed "real_le_add_iff1";
-
-Goal "!!i::real. (i*u + m <= j*u + n) = (m <= (j-i)*u + n)";
-by (asm_simp_tac (simpset() addsimps [real_diff_def, left_distrib]
-                                     @add_ac@rel_iff_rel_0_rls) 1);
-qed "real_le_add_iff2";
-
-
-Addsimps [real_numeral_0_eq_0, real_numeral_1_eq_1];
-
-
-structure Real_Numeral_Simprocs =
-struct
-
-(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in simprocs
-  isn't complicated by the abstract 0 and 1.*)
-val numeral_syms = [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym];
-
-(*Utilities*)
-
-fun mk_numeral n = HOLogic.number_of_const HOLogic.realT $ HOLogic.mk_bin n;
-
-(*Decodes a binary real constant, or 0, 1*)
-val dest_numeral = Int_Numeral_Simprocs.dest_numeral;
-val find_first_numeral = Int_Numeral_Simprocs.find_first_numeral;
-
-val zero = mk_numeral 0;
-val mk_plus = HOLogic.mk_binop "op +";
-
-val uminus_const = Const ("uminus", HOLogic.realT --> HOLogic.realT);
-
-(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
-fun mk_sum []        = zero
-  | mk_sum [t,u]     = mk_plus (t, u)
-  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
-
-(*this version ALWAYS includes a trailing zero*)
-fun long_mk_sum []        = zero
-  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
-
-val dest_plus = HOLogic.dest_bin "op +" HOLogic.realT;
-
-(*decompose additions AND subtractions as a sum*)
-fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
-        dest_summing (pos, t, dest_summing (pos, u, ts))
-  | dest_summing (pos, Const ("op -", _) $ t $ u, ts) =
-        dest_summing (pos, t, dest_summing (not pos, u, ts))
-  | dest_summing (pos, t, ts) =
-        if pos then t::ts else uminus_const$t :: ts;
-
-fun dest_sum t = dest_summing (true, t, []);
-
-val mk_diff = HOLogic.mk_binop "op -";
-val dest_diff = HOLogic.dest_bin "op -" HOLogic.realT;
-
-val one = mk_numeral 1;
-val mk_times = HOLogic.mk_binop "op *";
-
-fun mk_prod [] = one
-  | mk_prod [t] = t
-  | mk_prod (t :: ts) = if t = one then mk_prod ts
-                        else mk_times (t, mk_prod ts);
-
-val dest_times = HOLogic.dest_bin "op *" HOLogic.realT;
-
-fun dest_prod t =
-      let val (t,u) = dest_times t
-      in  dest_prod t @ dest_prod u  end
-      handle TERM _ => [t];
-
-(*DON'T do the obvious simplifications; that would create special cases*)
-fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts);
-
-(*Express t as a product of (possibly) a numeral with other sorted terms*)
-fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t
-  | dest_coeff sign t =
-    let val ts = sort Term.term_ord (dest_prod t)
-        val (n, ts') = find_first_numeral [] ts
-                          handle TERM _ => (1, ts)
-    in (sign*n, mk_prod ts') end;
-
-(*Find first coefficient-term THAT MATCHES u*)
-fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
-  | find_first_coeff past u (t::terms) =
-        let val (n,u') = dest_coeff 1 t
-        in  if u aconv u' then (n, rev past @ terms)
-                          else find_first_coeff (t::past) u terms
-        end
-        handle TERM _ => find_first_coeff (t::past) u terms;
-
-
-(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
-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_minus1, real_mult_minus1_right];
-
-(*To perform binary arithmetic*)
-val bin_simps =
-    [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym,
-     add_real_number_of, real_add_number_of_left, minus_real_number_of,
-     diff_real_number_of, mult_real_number_of, real_mult_number_of_left] @
-    bin_arith_simps @ bin_rel_simps;
-
-(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
-  during re-arrangement*)
-val non_add_bin_simps = 
-    bin_simps \\ [real_add_number_of_left, add_real_number_of];
-
-(*To evaluate binary negations of coefficients*)
-val real_minus_simps = NCons_simps @
-                   [real_minus_1_eq_m1, minus_real_number_of,
-                    bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min,
-                    bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min];
-
-(*To let us treat subtraction as addition*)
-val diff_simps = [real_diff_def, minus_add_distrib, minus_minus];
-
-(*to extract again any uncancelled minuses*)
-val real_minus_from_mult_simps =
-    [minus_minus, minus_mult_left RS sym, minus_mult_right RS sym];
-
-(*combine unary minus with numeric literals, however nested within a product*)
-val real_mult_minus_simps =
-    [real_mult_assoc, minus_mult_left, real_minus_mult_commute];
-
-(*Apply the given rewrite (if present) just once*)
-fun trans_tac None      = all_tac
-  | trans_tac (Some th) = ALLGOALS (rtac (th RS trans));
-
-(*Final simplification: cancel + and *  *)
-val simplify_meta_eq =
-    Int_Numeral_Simprocs.simplify_meta_eq
-         [add_0, add_0_right,
-          mult_zero_left, mult_zero_right, mult_1, mult_1_right];
-
-fun prep_simproc (name, pats, proc) =
-  Simplifier.simproc (Theory.sign_of (the_context ())) name pats proc;
-
-structure CancelNumeralsCommon =
-  struct
-  val mk_sum            = mk_sum
-  val dest_sum          = dest_sum
-  val mk_coeff          = mk_coeff
-  val dest_coeff        = dest_coeff 1
-  val find_first_coeff  = find_first_coeff []
-  val trans_tac         = trans_tac
-  val norm_tac =
-     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
-                                         real_minus_simps@add_ac))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@real_mult_minus_simps))
-     THEN ALLGOALS
-              (simp_tac (HOL_ss addsimps real_minus_from_mult_simps@
-                                         add_ac@mult_ac))
-  val numeral_simp_tac  = ALLGOALS (simp_tac (HOL_ss addsimps add_0s@bin_simps))
-  val simplify_meta_eq  = simplify_meta_eq
-  end;
-
-
-structure EqCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Bin_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.realT
-  val bal_add1 = real_eq_add_iff1 RS trans
-  val bal_add2 = real_eq_add_iff2 RS trans
-);
-
-structure LessCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Bin_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binrel "op <"
-  val dest_bal = HOLogic.dest_bin "op <" HOLogic.realT
-  val bal_add1 = real_less_add_iff1 RS trans
-  val bal_add2 = real_less_add_iff2 RS trans
-);
-
-structure LeCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Bin_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binrel "op <="
-  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.realT
-  val bal_add1 = real_le_add_iff1 RS trans
-  val bal_add2 = real_le_add_iff2 RS trans
-);
-
-val cancel_numerals =
-  map prep_simproc
-   [("realeq_cancel_numerals",
-     ["(l::real) + m = n", "(l::real) = m + n",
-      "(l::real) - m = n", "(l::real) = m - n",
-      "(l::real) * m = n", "(l::real) = m * n"],
-     EqCancelNumerals.proc),
-    ("realless_cancel_numerals",
-     ["(l::real) + m < n", "(l::real) < m + n",
-      "(l::real) - m < n", "(l::real) < m - n",
-      "(l::real) * m < n", "(l::real) < m * n"],
-     LessCancelNumerals.proc),
-    ("realle_cancel_numerals",
-     ["(l::real) + m <= n", "(l::real) <= m + n",
-      "(l::real) - m <= n", "(l::real) <= m - n",
-      "(l::real) * m <= n", "(l::real) <= m * n"],
-     LeCancelNumerals.proc)];
-
-
-structure CombineNumeralsData =
-  struct
-  val add               = op + : int*int -> int
-  val mk_sum            = long_mk_sum    (*to work for e.g. 2*x + 3*x *)
-  val dest_sum          = dest_sum
-  val mk_coeff          = mk_coeff
-  val dest_coeff        = dest_coeff 1
-  val left_distrib      = left_real_add_mult_distrib RS trans
-  val prove_conv        = Bin_Simprocs.prove_conv_nohyps
-  val trans_tac         = trans_tac
-  val norm_tac =
-     ALLGOALS (simp_tac (HOL_ss addsimps numeral_syms@add_0s@mult_1s@
-                                   diff_simps@real_minus_simps@add_ac))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@real_mult_minus_simps))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps real_minus_from_mult_simps@
-                                              add_ac@mult_ac))
-  val numeral_simp_tac  = ALLGOALS
-                    (simp_tac (HOL_ss addsimps add_0s@bin_simps))
-  val simplify_meta_eq  =
-        Int_Numeral_Simprocs.simplify_meta_eq (add_0s@mult_1s)
-  end;
-
-structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
-
-val combine_numerals =
-  prep_simproc ("real_combine_numerals", ["(i::real) + j", "(i::real) - j"], CombineNumerals.proc);
-
-
-(** Declarations for ExtractCommonTerm **)
-
-(*this version ALWAYS includes a trailing one*)
-fun long_mk_prod []        = one
-  | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts);
-
-(*Find first term that matches u*)
-fun find_first past u []         = raise TERM("find_first", [])
-  | find_first past u (t::terms) =
-        if u aconv t then (rev past @ terms)
-        else find_first (t::past) u terms
-        handle TERM _ => find_first (t::past) u terms;
-
-(*Final simplification: cancel + and *  *)
-fun cancel_simplify_meta_eq cancel_th th =
-    Int_Numeral_Simprocs.simplify_meta_eq
-        [real_mult_1, real_mult_1_right]
-        (([th, cancel_th]) MRS trans);
-
-(*** Making constant folding work for 0 and 1 too ***)
-
-structure RealAbstractNumeralsData =
-  struct
-  val dest_eq         = HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of
-  val is_numeral      = Bin_Simprocs.is_numeral
-  val numeral_0_eq_0  = real_numeral_0_eq_0
-  val numeral_1_eq_1  = real_numeral_1_eq_1
-  val prove_conv      = Bin_Simprocs.prove_conv_nohyps_novars
-  fun norm_tac simps  = ALLGOALS (simp_tac (HOL_ss addsimps simps))
-  val simplify_meta_eq = Bin_Simprocs.simplify_meta_eq
-  end
-
-structure RealAbstractNumerals = AbstractNumeralsFun (RealAbstractNumeralsData)
-
-(*For addition, we already have rules for the operand 0.
-  Multiplication is omitted because there are already special rules for
-  both 0 and 1 as operands.  Unary minus is trivial, just have - 1 = -1.
-  For the others, having three patterns is a compromise between just having
-  one (many spurious calls) and having nine (just too many!) *)
-val eval_numerals =
-  map prep_simproc
-   [("real_add_eval_numerals",
-     ["(m::real) + 1", "(m::real) + number_of v"],
-     RealAbstractNumerals.proc add_real_number_of),
-    ("real_diff_eval_numerals",
-     ["(m::real) - 1", "(m::real) - number_of v"],
-     RealAbstractNumerals.proc diff_real_number_of),
-    ("real_eq_eval_numerals",
-     ["(m::real) = 0", "(m::real) = 1", "(m::real) = number_of v"],
-     RealAbstractNumerals.proc eq_real_number_of),
-    ("real_less_eval_numerals",
-     ["(m::real) < 0", "(m::real) < 1", "(m::real) < number_of v"],
-     RealAbstractNumerals.proc less_real_number_of),
-    ("real_le_eval_numerals",
-     ["(m::real) <= 0", "(m::real) <= 1", "(m::real) <= number_of v"],
-     RealAbstractNumerals.proc le_real_number_of_eq_not_less)]
-
-end;
-
-
-Addsimprocs Real_Numeral_Simprocs.eval_numerals;
-Addsimprocs Real_Numeral_Simprocs.cancel_numerals;
-Addsimprocs [Real_Numeral_Simprocs.combine_numerals];
-
-(*examples:
-print_depth 22;
-set timing;
-set trace_simp;
-fun test s = (Goal s; by (Simp_tac 1));
-
-test "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::real)";
-
-test "2*u = (u::real)";
-test "(i + j + 12 + (k::real)) - 15 = y";
-test "(i + j + 12 + (k::real)) - 5 = y";
-
-test "y - b < (b::real)";
-test "y - (3*b + c) < (b::real) - 2*c";
-
-test "(2*x - (u*v) + y) - v*3*u = (w::real)";
-test "(2*x*u*v + (u*v)*4 + y) - v*u*4 = (w::real)";
-test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::real)";
-test "u*v - (x*u*v + (u*v)*4 + y) = (w::real)";
-
-test "(i + j + 12 + (k::real)) = u + 15 + y";
-test "(i + j*2 + 12 + (k::real)) = j + 5 + y";
-
-test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::real)";
-
-test "a + -(b+c) + b = (d::real)";
-test "a + -(b+c) - b = (d::real)";
-
-(*negative numerals*)
-test "(i + j + -2 + (k::real)) - (u + 5 + y) = zz";
-test "(i + j + -3 + (k::real)) < u + 5 + y";
-test "(i + j + 3 + (k::real)) < u + -6 + y";
-test "(i + j + -12 + (k::real)) - 15 = y";
-test "(i + j + 12 + (k::real)) - -15 = y";
-test "(i + j + -12 + (k::real)) - -15 = y";
-*)
-
-
-(** Constant folding for real plus and times **)
-
-(*We do not need
-    structure Real_Plus_Assoc = Assoc_Fold (Real_Plus_Assoc_Data);
-  because combine_numerals does the same thing*)
-
-structure Real_Times_Assoc_Data : ASSOC_FOLD_DATA =
-struct
-  val ss                = HOL_ss
-  val eq_reflection     = eq_reflection
-  val sg_ref    = Sign.self_ref (Theory.sign_of (the_context ()))
-  val T      = HOLogic.realT
-  val plus   = Const ("op *", [HOLogic.realT,HOLogic.realT] ---> HOLogic.realT)
-  val add_ac = mult_ac
-end;
-
-structure Real_Times_Assoc = Assoc_Fold (Real_Times_Assoc_Data);
-
-Addsimprocs [Real_Times_Assoc.conv];
-
-(*Simplification of  x-y < 0, etc.*)
-AddIffs [less_iff_diff_less_0 RS sym];
-AddIffs [eq_iff_diff_eq_0 RS sym];
-AddIffs [le_iff_diff_le_0 RS sym];
-
-Addsimps [real_minus_1_eq_m1];
--- a/src/HOL/Real/RealBin.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Real/RealBin.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -1,21 +0,0 @@
-(*  Title:      HOL/RealBin.thy
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1999  University of Cambridge
-
-Binary arithmetic for the reals
-
-This case is reduced to that for the integers
-*)
-
-RealBin = RealInt +
-
-instance
-  real :: number 
-
-defs
-  real_number_of_def
-    "number_of v == real (number_of v :: int)"
-     (*::bin=>real           ::bin=>int*)
-
-end
--- a/src/HOL/Real/RealDef.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Real/RealDef.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -35,13 +35,12 @@
 defs (overloaded)
 
   real_zero_def:
-  "0 == Abs_REAL(realrel``{(preal_of_prat(prat_of_pnat 1),
-			    preal_of_prat(prat_of_pnat 1))})"
+  "0 == Abs_REAL(realrel``{(preal_of_rat 1, preal_of_rat 1)})"
 
   real_one_def:
   "1 == Abs_REAL(realrel``
-               {(preal_of_prat(prat_of_pnat 1) + preal_of_prat(prat_of_pnat 1),
-		 preal_of_prat(prat_of_pnat 1))})"
+               {(preal_of_rat 1 + preal_of_rat 1,
+		 preal_of_rat 1)})"
 
   real_minus_def:
   "- R ==  Abs_REAL(UN (x,y):Rep_REAL(R). realrel``{(y,x)})"
@@ -61,17 +60,10 @@
 
   real_of_preal :: "preal => real"
   "real_of_preal m     ==
-           Abs_REAL(realrel``{(m + preal_of_prat(prat_of_pnat 1),
-                               preal_of_prat(prat_of_pnat 1))})"
-
-  real_of_posnat :: "nat => real"
-  "real_of_posnat n == real_of_preal(preal_of_prat(prat_of_pnat(pnat_of_nat n)))"
-
+           Abs_REAL(realrel``{(m + preal_of_rat 1, preal_of_rat 1)})"
 
 defs (overloaded)
 
-  real_of_nat_def:   "real n == real_of_posnat n + (- 1)"
-
   real_add_def:
   "P+Q == Abs_REAL(\<Union>p1\<in>Rep_REAL(P). \<Union>p2\<in>Rep_REAL(Q).
                    (%(x1,y1). (%(x2,y2). realrel``{(x1+x2, y1+y2)}) p2) p1)"
@@ -81,11 +73,12 @@
                    (%(x1,y1). (%(x2,y2). realrel``{(x1*x2+y1*y2,x1*y2+x2*y1)})
 		   p2) p1)"
 
-  real_less_def:
-  "P<Q == \<exists>x1 y1 x2 y2. x1 + y2 < x2 + y1 &
-                            (x1,y1)\<in>Rep_REAL(P) & (x2,y2)\<in>Rep_REAL(Q)"
+  real_less_def: "(x < (y::real)) == (x \<le> y & x \<noteq> y)"
+
+
   real_le_def:
-  "P \<le> (Q::real) == ~(Q < P)"
+  "P \<le> (Q::real) == \<exists>x1 y1 x2 y2. x1 + y2 \<le> x2 + y1 &
+                            (x1,y1) \<in> Rep_REAL(P) & (x2,y2) \<in> Rep_REAL(Q)"
 
   real_abs_def:  "abs (r::real) == (if 0 \<le> r then r else -r)"
 
@@ -95,18 +88,31 @@
   Nats      :: "'a set"                   ("\<nat>")
 
 
+defs (overloaded)
+  real_of_int_def:
+   "real z == Abs_REAL(\<Union>(i,j) \<in> Rep_Integ z. realrel ``
+		       {(preal_of_rat(rat(int(Suc i))),
+			 preal_of_rat(rat(int(Suc j))))})"
+
+  real_of_nat_def:   "real n == real (int n)"
+
+
 subsection{*Proving that realrel is an equivalence relation*}
 
 lemma preal_trans_lemma:
-     "[| (x1::preal) + y2 = x2 + y1; x2 + y3 = x3 + y2 |]
-      ==> x1 + y3 = x3 + y1"
-apply (rule_tac C = y2 in preal_add_right_cancel)
-apply (rotate_tac 1, drule sym)
-apply (simp add: preal_add_ac)
-apply (rule preal_add_left_commute [THEN subst])
-apply (rule_tac x1 = x1 in preal_add_assoc [THEN subst])
-apply (simp add: preal_add_ac)
-done
+  assumes "x + y1 = x1 + y"
+      and "x + y2 = x2 + y"
+  shows "x1 + y2 = x2 + (y1::preal)"
+proof -
+  have "(x1 + y2) + x = (x + y2) + x1" by (simp add: preal_add_ac) 
+  also have "... = (x2 + y) + x1"  by (simp add: prems)
+  also have "... = x2 + (x1 + y)"  by (simp add: preal_add_ac)
+  also have "... = x2 + (x + y1)"  by (simp add: prems)
+  also have "... = (x2 + y1) + x"  by (simp add: preal_add_ac)
+  finally have "(x1 + y2) + x = (x2 + y1) + x" .
+  thus ?thesis by (simp add: preal_add_right_cancel_iff) 
+qed
+
 
 lemma realrel_iff [simp]: "(((x1,y1),(x2,y2)): realrel) = (x1 + y2 = x2 + y1)"
 by (unfold realrel_def, blast)
@@ -117,8 +123,8 @@
 done
 
 lemma equiv_realrel: "equiv UNIV realrel"
-apply (unfold equiv_def refl_def sym_def trans_def realrel_def)
-apply (fast elim!: sym preal_trans_lemma)
+apply (auto simp add: equiv_def refl_def sym_def trans_def realrel_def)
+apply (blast dest: preal_trans_lemma) 
 done
 
 (* (realrel `` {x} = realrel `` {y}) = ((x,y) : realrel) *)
@@ -130,6 +136,7 @@
 lemma realrel_in_real [simp]: "realrel``{(x,y)}: REAL"
 by (unfold REAL_def realrel_def quotient_def, blast)
 
+
 lemma inj_on_Abs_REAL: "inj_on Abs_REAL REAL"
 apply (rule inj_on_inverseI)
 apply (erule Abs_REAL_inverse)
@@ -154,7 +161,7 @@
 apply (rule realrel_in_real)+
 apply (drule eq_equiv_class)
 apply (rule equiv_realrel, blast)
-apply (simp add: realrel_def)
+apply (simp add: realrel_def preal_add_right_cancel_iff)
 done
 
 lemma eq_Abs_REAL: 
@@ -165,6 +172,30 @@
 apply (simp add: Rep_REAL_inverse)
 done
 
+lemma real_eq_iff:
+     "[|(x1,y1) \<in> Rep_REAL w; (x2,y2) \<in> Rep_REAL z|] 
+      ==> (z = w) = (x1+y2 = x2+y1)"
+apply (insert quotient_eq_iff
+                [OF equiv_realrel, 
+                 of "Rep_REAL w" "Rep_REAL z" "(x1,y1)" "(x2,y2)"])
+apply (simp add: Rep_REAL [unfolded REAL_def] Rep_REAL_inject eq_commute) 
+done 
+
+lemma mem_REAL_imp_eq:
+     "[|R \<in> REAL; (x1,y1) \<in> R; (x2,y2) \<in> R|] ==> x1+y2 = x2+y1" 
+apply (auto simp add: REAL_def realrel_def quotient_def)
+apply (blast dest: preal_trans_lemma) 
+done
+
+lemma Rep_REAL_cancel_right:
+     "((x + z, y + z) \<in> Rep_REAL R) = ((x, y) \<in> Rep_REAL R)"
+apply (rule_tac z = R in eq_Abs_REAL, simp) 
+apply (rename_tac u v) 
+apply (subgoal_tac "(u + (y + z) = x + z + v) = ((u + y) + z = (x + v) + z)")
+ prefer 2 apply (simp add: preal_add_ac) 
+apply (simp add: preal_add_right_cancel_iff) 
+done
+
 
 subsection{*Congruence property for addition*}
 
@@ -280,7 +311,7 @@
 done
 
 lemma real_mult_1: "(1::real) * z = z"
-apply (unfold real_one_def pnat_one_def)
+apply (unfold real_one_def)
 apply (rule_tac z = z in eq_Abs_REAL)
 apply (simp add: real_mult preal_add_mult_distrib2 preal_mult_1_right
                  preal_mult_ac preal_add_ac)
@@ -294,39 +325,44 @@
 done
 
 text{*one and zero are distinct*}
-lemma real_zero_not_eq_one: "0 ~= (1::real)"
+lemma real_zero_not_eq_one: "0 \<noteq> (1::real)"
+apply (subgoal_tac "preal_of_rat 1 < preal_of_rat 1 + preal_of_rat 1")
+ prefer 2 apply (simp add: preal_self_less_add_left) 
 apply (unfold real_zero_def real_one_def)
-apply (auto simp add: preal_self_less_add_left [THEN preal_not_refl2])
+apply (auto simp add: preal_add_right_cancel_iff)
 done
 
 subsection{*existence of inverse*}
-(** lemma -- alternative definition of 0 **)
-lemma real_zero_iff: "0 = Abs_REAL (realrel `` {(x, x)})"
+
+lemma real_zero_iff: "Abs_REAL (realrel `` {(x, x)}) = 0"
 apply (unfold real_zero_def)
 apply (auto simp add: preal_add_commute)
 done
 
-lemma real_mult_inv_left_ex: "x ~= 0 ==> \<exists>y. y*x = (1::real)"
+text{*Instead of using an existential quantifier and constructing the inverse
+within the proof, we could define the inverse explicitly.*}
+
+lemma real_mult_inverse_left_ex: "x \<noteq> 0 ==> \<exists>y. y*x = (1::real)"
 apply (unfold real_zero_def real_one_def)
 apply (rule_tac z = x in eq_Abs_REAL)
 apply (cut_tac x = xa and y = y in linorder_less_linear)
-apply (auto dest!: preal_less_add_left_Ex simp add: real_zero_iff [symmetric])
+apply (auto dest!: less_add_left_Ex simp add: real_zero_iff)
 apply (rule_tac
-        x = "Abs_REAL (realrel `` { (preal_of_prat (prat_of_pnat 1), 
-                            pinv (D) + preal_of_prat (prat_of_pnat 1))}) " 
+        x = "Abs_REAL (realrel `` { (preal_of_rat 1, 
+                            inverse (D) + preal_of_rat 1)}) " 
        in exI)
 apply (rule_tac [2]
-        x = "Abs_REAL (realrel `` { (pinv (D) + preal_of_prat (prat_of_pnat 1),
-                   preal_of_prat (prat_of_pnat 1))})" 
+        x = "Abs_REAL (realrel `` { (inverse (D) + preal_of_rat 1,
+                   preal_of_rat 1)})" 
        in exI)
-apply (auto simp add: real_mult pnat_one_def preal_mult_1_right
+apply (auto simp add: real_mult preal_mult_1_right
               preal_add_mult_distrib2 preal_add_mult_distrib preal_mult_1
-              preal_mult_inv_right preal_add_ac preal_mult_ac)
+              preal_mult_inverse_right preal_add_ac preal_mult_ac)
 done
 
-lemma real_mult_inv_left: "x ~= 0 ==> inverse(x)*x = (1::real)"
+lemma real_mult_inverse_left: "x \<noteq> 0 ==> inverse(x)*x = (1::real)"
 apply (unfold real_inverse_def)
-apply (frule real_mult_inv_left_ex, safe)
+apply (frule real_mult_inverse_left_ex, safe)
 apply (rule someI2, auto)
 done
 
@@ -346,7 +382,7 @@
   show "1 * x = x" by (rule real_mult_1)
   show "(x + y) * z = x * z + y * z" by (simp add: real_add_mult_distrib)
   show "0 \<noteq> (1::real)" by (rule real_zero_not_eq_one)
-  show "x \<noteq> 0 ==> inverse x * x = 1" by (rule real_mult_inv_left)
+  show "x \<noteq> 0 ==> inverse x * x = 1" by (rule real_mult_inverse_left)
   show "y \<noteq> 0 ==> x / y = x * inverse y" by (simp add: real_divide_def)
   assume eq: "z+x = z+y" 
     hence "(-z + z) + x = (-z + z) + y" by (simp only: eq real_add_assoc)
@@ -377,414 +413,199 @@
 declare minus_mult_right [symmetric, simp] 
         minus_mult_left [symmetric, simp]
 
-text{*Used in RealBin*}
-lemma real_minus_mult_commute: "(-x) * y = x * (- y :: real)"
-by simp
-
 lemma real_mult_1_right: "z * (1::real) = z"
   by (rule Ring_and_Field.mult_1_right)
 
 
-subsection{*Theorems for Ordering*}
-
-(* real_less is a strict order: irreflexive *)
-
-text{*lemmas*}
-lemma preal_lemma_eq_rev_sum:
-     "!!(x::preal). [| x = y; x1 = y1 |] ==> x + y1 = x1 + y"
-by (simp add: preal_add_commute)
-
-lemma preal_add_left_commute_cancel:
-     "!!(b::preal). x + (b + y) = x1 + (b + y1) ==> x + y = x1 + y1"
-by (simp add: preal_add_ac)
-
-lemma preal_lemma_for_not_refl:
-     "!!(x::preal). [| x + y2a = x2a + y;
-                       x + y2b = x2b + y |]
-                    ==> x2a + y2b = x2b + y2a"
-apply (drule preal_lemma_eq_rev_sum, assumption)
-apply (erule_tac V = "x + y2b = x2b + y" in thin_rl)
-apply (simp add: preal_add_ac)
-apply (drule preal_add_left_commute_cancel)
-apply (simp add: preal_add_ac)
-done
-
-lemma real_less_not_refl: "~ (R::real) < R"
-apply (rule_tac z = R in eq_Abs_REAL)
-apply (auto simp add: real_less_def)
-apply (drule preal_lemma_for_not_refl, assumption, auto)
-done
-
-(*** y < y ==> P ***)
-lemmas real_less_irrefl = real_less_not_refl [THEN notE, standard]
-declare real_less_irrefl [elim!]
-
-lemma real_not_refl2: "!!(x::real). x < y ==> x ~= y"
-by (auto simp add: real_less_not_refl)
-
-(* lemma re-arranging and eliminating terms *)
-lemma preal_lemma_trans: "!! (a::preal). [| a + b = c + d;
-             x2b + d + (c + y2e) < a + y2b + (x2e + b) |]
-          ==> x2b + y2e < x2e + y2b"
-apply (simp add: preal_add_ac)
-apply (rule_tac C = "c+d" in preal_add_left_less_cancel)
-apply (simp add: preal_add_assoc [symmetric])
-done
-
-(** A MESS!  heavy re-writing involved*)
-lemma real_less_trans: "!!(R1::real). [| R1 < R2; R2 < R3 |] ==> R1 < R3"
-apply (rule_tac z = R1 in eq_Abs_REAL)
-apply (rule_tac z = R2 in eq_Abs_REAL)
-apply (rule_tac z = R3 in eq_Abs_REAL)
-apply (auto simp add: real_less_def)
-apply (rule exI)+
-apply (rule conjI, rule_tac [2] conjI)
- prefer 2 apply blast 
- prefer 2 apply blast 
-apply (drule preal_lemma_for_not_refl, assumption)
-apply (blast dest: preal_add_less_mono intro: preal_lemma_trans)
-done
-
-lemma real_less_not_sym: "!! (R1::real). R1 < R2 ==> ~ (R2 < R1)"
-apply (rule notI)
-apply (drule real_less_trans, assumption)
-apply (simp add: real_less_not_refl)
-done
-
-(* [| x < y;  ~P ==> y < x |] ==> P *)
-lemmas real_less_asym = real_less_not_sym [THEN contrapos_np, standard]
+subsection{*The @{text "\<le>"} Ordering*}
 
-lemma real_of_preal_add:
-     "real_of_preal ((z1::preal) + z2) =
-      real_of_preal z1 + real_of_preal z2"
-apply (unfold real_of_preal_def)
-apply (simp add: real_add preal_add_mult_distrib preal_mult_1 add: preal_add_ac)
-done
-
-lemma real_of_preal_mult:
-     "real_of_preal ((z1::preal) * z2) =
-      real_of_preal z1* real_of_preal z2"
-apply (unfold real_of_preal_def)
-apply (simp (no_asm_use) add: real_mult preal_add_mult_distrib2 preal_mult_1 preal_mult_1_right pnat_one_def preal_add_ac preal_mult_ac)
-done
-
-lemma real_of_preal_ExI:
-      "!!(x::preal). y < x ==>
-       \<exists>m. Abs_REAL (realrel `` {(x,y)}) = real_of_preal m"
-apply (unfold real_of_preal_def)
-apply (auto dest!: preal_less_add_left_Ex simp add: preal_add_ac)
-done
-
-lemma real_of_preal_ExD:
-      "!!(x::preal). \<exists>m. Abs_REAL (realrel `` {(x,y)}) =
-                     real_of_preal m ==> y < x"
-apply (unfold real_of_preal_def)
-apply (auto simp add: preal_add_commute preal_add_assoc)
-apply (simp add: preal_add_assoc [symmetric] preal_self_less_add_left)
-done
-
-lemma real_of_preal_iff:
-     "(\<exists>m. Abs_REAL (realrel `` {(x,y)}) = real_of_preal m) = (y < x)"
-by (blast intro!: real_of_preal_ExI real_of_preal_ExD)
-
-text{*Gleason prop 9-4.4 p 127*}
-lemma real_of_preal_trichotomy:
-      "\<exists>m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)"
-apply (unfold real_of_preal_def real_zero_def)
-apply (rule_tac z = x in eq_Abs_REAL)
-apply (auto simp add: real_minus preal_add_ac)
-apply (cut_tac x = x and y = y in linorder_less_linear)
-apply (auto dest!: preal_less_add_left_Ex simp add: preal_add_assoc [symmetric])
-apply (auto simp add: preal_add_commute)
-done
-
-lemma real_of_preal_trichotomyE:
-     "!!P. [| !!m. x = real_of_preal m ==> P;
-              x = 0 ==> P;
-              !!m. x = -(real_of_preal m) ==> P |] ==> P"
-apply (cut_tac x = x in real_of_preal_trichotomy, auto)
-done
-
-lemma real_of_preal_lessD:
-      "real_of_preal m1 < real_of_preal m2 ==> m1 < m2"
-apply (unfold real_of_preal_def)
-apply (auto simp add: real_less_def preal_add_ac)
-apply (auto simp add: preal_add_assoc [symmetric])
-apply (auto simp add: preal_add_ac)
-done
-
-lemma real_of_preal_lessI: "m1 < m2 ==> real_of_preal m1 < real_of_preal m2"
-apply (drule preal_less_add_left_Ex)
-apply (auto simp add: real_of_preal_add real_of_preal_def real_less_def)
-apply (rule exI)+
-apply (rule conjI, rule_tac [2] conjI)
- apply (rule_tac [2] refl)+
-apply (simp add: preal_self_less_add_left del: preal_add_less_iff2)
-done
-
-lemma real_of_preal_less_iff1:
-     "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)"
-by (blast intro: real_of_preal_lessI real_of_preal_lessD)
-
-declare real_of_preal_less_iff1 [simp]
-
-lemma real_of_preal_minus_less_self: "- real_of_preal m < real_of_preal m"
-apply (auto simp add: real_of_preal_def real_less_def real_minus)
-apply (rule exI)+
-apply (rule conjI, rule_tac [2] conjI)
- apply (rule_tac [2] refl)+
-apply (simp (no_asm_use) add: preal_add_ac)
-apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_assoc [symmetric])
+lemma real_le_refl: "w \<le> (w::real)"
+apply (rule_tac z = w in eq_Abs_REAL)
+apply (force simp add: real_le_def)
 done
 
-lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0"
-apply (unfold real_zero_def)
-apply (auto simp add: real_of_preal_def real_less_def real_minus)
-apply (rule exI)+
-apply (rule conjI, rule_tac [2] conjI)
- apply (rule_tac [2] refl)+
-apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_ac)
-done
-
-lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m"
-apply (cut_tac real_of_preal_minus_less_zero)
-apply (fast dest: real_less_trans elim: real_less_irrefl)
-done
-
-lemma real_of_preal_zero_less: "0 < real_of_preal m"
-apply (unfold real_zero_def)
-apply (auto simp add: real_of_preal_def real_less_def real_minus)
-apply (rule exI)+
-apply (rule conjI, rule_tac [2] conjI)
- apply (rule_tac [2] refl)+
-apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_ac)
-done
-
-lemma real_of_preal_not_less_zero: "~ real_of_preal m < 0"
-apply (cut_tac real_of_preal_zero_less)
-apply (blast dest: real_less_trans elim: real_less_irrefl)
-done
-
-lemma real_minus_minus_zero_less: "0 < - (- real_of_preal m)"
-by (simp add: real_of_preal_zero_less)
+lemma real_le_trans_lemma:
+  assumes le1: "x1 + y2 \<le> x2 + y1"
+      and le2: "u1 + v2 \<le> u2 + v1"
+      and eq: "x2 + v1 = u1 + y2"
+  shows "x1 + v2 + u1 + y2 \<le> u2 + u1 + y2 + (y1::preal)"
+proof -
+  have "x1 + v2 + u1 + y2 = (x1 + y2) + (u1 + v2)" by (simp add: preal_add_ac)
+  also have "... \<le> (x2 + y1) + (u1 + v2)"
+         by (simp add: prems preal_add_le_cancel_right)
+  also have "... \<le> (x2 + y1) + (u2 + v1)"
+         by (simp add: prems preal_add_le_cancel_left)
+  also have "... = (x2 + v1) + (u2 + y1)" by (simp add: preal_add_ac)
+  also have "... = (u1 + y2) + (u2 + y1)" by (simp add: prems)
+  also have "... = u2 + u1 + y2 + y1" by (simp add: preal_add_ac)
+  finally show ?thesis .
+qed						 
 
-(* another lemma *)
-lemma real_of_preal_sum_zero_less:
-      "0 < real_of_preal m + real_of_preal m1"
-apply (unfold real_zero_def)
-apply (auto simp add: real_of_preal_def real_less_def real_add)
-apply (rule exI)+
-apply (rule conjI, rule_tac [2] conjI)
- apply (rule_tac [2] refl)+
-apply (simp (no_asm_use) add: preal_add_ac)
-apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_assoc [symmetric])
-done
-
-lemma real_of_preal_minus_less_all: "- real_of_preal m < real_of_preal m1"
-apply (auto simp add: real_of_preal_def real_less_def real_minus)
-apply (rule exI)+
-apply (rule conjI, rule_tac [2] conjI)
- apply (rule_tac [2] refl)+
-apply (simp (no_asm_use) add: preal_add_ac)
-apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_assoc [symmetric])
-done
-
-lemma real_of_preal_not_minus_gt_all: "~ real_of_preal m < - real_of_preal m1"
-apply (cut_tac real_of_preal_minus_less_all)
-apply (blast dest: real_less_trans elim: real_less_irrefl)
-done
-
-lemma real_of_preal_minus_less_rev1:
-     "- real_of_preal m1 < - real_of_preal m2
-      ==> real_of_preal m2 < real_of_preal m1"
-apply (auto simp add: real_of_preal_def real_less_def real_minus)
-apply (rule exI)+
-apply (rule conjI, rule_tac [2] conjI)
- apply (rule_tac [2] refl)+
-apply (auto simp add: preal_add_ac)
-apply (simp add: preal_add_assoc [symmetric])
-apply (auto simp add: preal_add_ac)
+lemma real_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::real)"
+apply (simp add: real_le_def, clarify)
+apply (rename_tac x1 u1 y1 v1 x2 u2 y2 v2) 
+apply (drule mem_REAL_imp_eq [OF Rep_REAL], assumption)  
+apply (rule_tac x=x1 in exI) 
+apply (rule_tac x=y1 in exI) 
+apply (rule_tac x="u2 + (x2 + v1)" in exI) 
+apply (rule_tac x="v2 + (u1 + y2)" in exI) 
+apply (simp add: Rep_REAL_cancel_right preal_add_le_cancel_right 
+                 preal_add_assoc [symmetric] real_le_trans_lemma)
 done
 
-lemma real_of_preal_minus_less_rev2:
-     "real_of_preal m1 < real_of_preal m2
-      ==> - real_of_preal m2 < - real_of_preal m1"
-apply (auto simp add: real_of_preal_def real_less_def real_minus)
-apply (rule exI)+
-apply (rule conjI, rule_tac [2] conjI)
- apply (rule_tac [2] refl)+
-apply (auto simp add: preal_add_ac)
-apply (simp add: preal_add_assoc [symmetric])
-apply (auto simp add: preal_add_ac)
-done
-
-lemma real_of_preal_minus_less_rev_iff:
-     "(- real_of_preal m1 < - real_of_preal m2) =
-      (real_of_preal m2 < real_of_preal m1)"
-apply (blast intro!: real_of_preal_minus_less_rev1 real_of_preal_minus_less_rev2)
-done
-
-
-subsection{*Linearity of the Ordering*}
-
-lemma real_linear: "(x::real) < y | x = y | y < x"
-apply (rule_tac x = x in real_of_preal_trichotomyE)
-apply (rule_tac [!] x = y in real_of_preal_trichotomyE)
-apply (auto dest!: preal_le_anti_sym 
-            simp add: preal_less_le_iff real_of_preal_minus_less_zero 
-                      real_of_preal_zero_less real_of_preal_minus_less_all
-                      real_of_preal_minus_less_rev_iff)
-done
-
-lemma real_neq_iff: "!!w::real. (w ~= z) = (w<z | z<w)"
-by (cut_tac real_linear, blast)
-
-
-lemma real_linear_less2:
-     "!!(R1::real). [| R1 < R2 ==> P;  R1 = R2 ==> P;
-                       R2 < R1 ==> P |] ==> P"
-apply (cut_tac x = R1 and y = R2 in real_linear, auto)
-done
-
-lemma real_minus_zero_less_iff: "(0 < -R) = (R < (0::real))"
-apply (rule_tac x = R in real_of_preal_trichotomyE)
-apply (auto simp add: real_of_preal_not_minus_gt_zero real_of_preal_not_less_zero real_of_preal_zero_less real_of_preal_minus_less_zero)
-done
-declare real_minus_zero_less_iff [simp]
-
-lemma real_minus_zero_less_iff2: "(-R < 0) = ((0::real) < R)"
-apply (rule_tac x = R in real_of_preal_trichotomyE)
-apply (auto simp add: real_of_preal_not_minus_gt_zero real_of_preal_not_less_zero real_of_preal_zero_less real_of_preal_minus_less_zero)
-done
-declare real_minus_zero_less_iff2 [simp]
-
-
-subsection{*Properties of Less-Than Or Equals*}
-
-lemma real_le_imp_less_or_eq: "!!(x::real). x \<le> y ==> x < y | x = y"
-apply (unfold real_le_def)
-apply (cut_tac real_linear)
-apply (blast elim: real_less_irrefl real_less_asym)
-done
-
-lemma real_less_or_eq_imp_le: "z<w | z=w ==> z \<le>(w::real)"
-apply (unfold real_le_def)
-apply (cut_tac real_linear)
-apply (fast elim: real_less_irrefl real_less_asym)
-done
-
-lemma real_le_less: "(x \<le> (y::real)) = (x < y | x=y)"
-by (blast intro!: real_less_or_eq_imp_le dest!: real_le_imp_less_or_eq)
-
-lemma real_le_refl: "w \<le> (w::real)"
-by (simp add: real_le_less)
-
-lemma real_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::real)"
-apply (drule real_le_imp_less_or_eq) 
-apply (drule real_le_imp_less_or_eq) 
-apply (rule real_less_or_eq_imp_le) 
-apply (blast intro: real_less_trans) 
-done
+lemma real_le_anti_sym_lemma: 
+  assumes le1: "x1 + y2 \<le> x2 + y1"
+      and le2: "u1 + v2 \<le> u2 + v1"
+      and eq1: "x1 + v2 = u2 + y1"
+      and eq2: "x2 + v1 = u1 + y2"
+  shows "x2 + y1 = x1 + (y2::preal)"
+proof (rule order_antisym)
+  show "x1 + y2 \<le> x2 + y1" .
+  have "(x2 + y1) + (u1+u2) = x2 + u1 + (u2 + y1)" by (simp add: preal_add_ac)
+  also have "... = x2 + u1 + (x1 + v2)" by (simp add: prems)
+  also have "... = (x2 + x1) + (u1 + v2)" by (simp add: preal_add_ac)
+  also have "... \<le> (x2 + x1) + (u2 + v1)" 
+                                  by (simp add: preal_add_le_cancel_left)
+  also have "... = (x1 + u2) + (x2 + v1)" by (simp add: preal_add_ac)
+  also have "... = (x1 + u2) + (u1 + y2)" by (simp add: prems)
+  also have "... = (x1 + y2) + (u1 + u2)" by (simp add: preal_add_ac)
+  finally show "x2 + y1 \<le> x1 + y2" by (simp add: preal_add_le_cancel_right)
+qed  
 
 lemma real_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::real)"
-apply (drule real_le_imp_less_or_eq) 
-apply (drule real_le_imp_less_or_eq) 
-apply (fast elim: real_less_irrefl real_less_asym)
+apply (simp add: real_le_def, clarify) 
+apply (rule real_eq_iff [THEN iffD2], assumption+)
+apply (drule mem_REAL_imp_eq [OF Rep_REAL], assumption)+
+apply (blast intro: real_le_anti_sym_lemma) 
 done
 
 (* Axiom 'order_less_le' of class 'order': *)
 lemma real_less_le: "((w::real) < z) = (w \<le> z & w \<noteq> z)"
-apply (simp add: real_le_def real_neq_iff)
-apply (blast elim!: real_less_asym)
+by (simp add: real_less_def)
+
+instance real :: order
+proof qed
+ (assumption |
+  rule real_le_refl real_le_trans real_le_anti_sym real_less_le)+
+
+text{*Simplifies a strange formula that occurs quantified.*}
+lemma preal_strange_le_eq: "(x1 + x2 \<le> x2 + y1) = (x1 \<le> (y1::preal))"
+by (simp add: preal_add_commute [of x1] preal_add_le_cancel_left) 
+
+text{*This is the nicest way to prove linearity*}
+lemma real_le_linear_0: "(z::real) \<le> 0 | 0 \<le> z"
+apply (rule_tac z = z in eq_Abs_REAL)
+apply (auto simp add: real_le_def real_zero_def preal_add_ac 
+                      preal_cancels preal_strange_le_eq)
+apply (cut_tac x=x and y=y in linorder_linear, auto) 
+done
+
+lemma real_minus_zero_le_iff: "(0 \<le> -R) = (R \<le> (0::real))"
+apply (rule_tac z = R in eq_Abs_REAL)
+apply (force simp add: real_le_def real_zero_def real_minus preal_add_ac 
+                       preal_cancels preal_strange_le_eq)
 done
 
-instance real :: order
-  by (intro_classes,
-      (assumption | 
-       rule real_le_refl real_le_trans real_le_anti_sym real_less_le)+)
+lemma real_le_imp_diff_le_0: "x \<le> y ==> x-y \<le> (0::real)"
+apply (rule_tac z = x in eq_Abs_REAL)
+apply (rule_tac z = y in eq_Abs_REAL)
+apply (auto simp add: real_le_def real_zero_def real_diff_def real_minus 
+    real_add preal_add_ac preal_cancels preal_strange_le_eq)
+apply (rule exI)+
+apply (rule conjI, assumption)
+apply (subgoal_tac " x + (x2 + y1 + ya) = (x + y1) + (x2 + ya)")
+ prefer 2 apply (simp (no_asm) only: preal_add_ac) 
+apply (subgoal_tac "x1 + y2 + (xa + y) = (x1 + y) + (xa + y2)")
+ prefer 2 apply (simp (no_asm) only: preal_add_ac) 
+apply simp 
+done
+
+lemma real_diff_le_0_imp_le: "x-y \<le> (0::real) ==> x \<le> y"
+apply (rule_tac z = x in eq_Abs_REAL)
+apply (rule_tac z = y in eq_Abs_REAL)
+apply (auto simp add: real_le_def real_zero_def real_diff_def real_minus 
+    real_add preal_add_ac preal_cancels preal_strange_le_eq)
+apply (rule exI)+
+apply (rule conjI, rule_tac [2] conjI)
+ apply (rule_tac [2] refl)+
+apply (subgoal_tac "(x + ya) + (x1 + y1) \<le> (xa + y) + (x1 + y1)") 
+apply (simp add: preal_cancels)
+apply (subgoal_tac "x1 + (x + (y1 + ya)) \<le> y1 + (x1 + (xa + y))")
+ apply (simp add: preal_add_ac) 
+apply (simp add: preal_cancels)
+done
+
+lemma real_le_eq_diff: "(x \<le> y) = (x-y \<le> (0::real))"
+by (blast intro!: real_diff_le_0_imp_le real_le_imp_diff_le_0)
+
 
 (* Axiom 'linorder_linear' of class 'linorder': *)
 lemma real_le_linear: "(z::real) \<le> w | w \<le> z"
-apply (simp add: real_le_less)
-apply (cut_tac real_linear, blast)
+apply (insert real_le_linear_0 [of "z-w"])
+apply (auto simp add: real_le_eq_diff [of w] real_le_eq_diff [of z] 
+                      real_minus_zero_le_iff [symmetric])
 done
 
 instance real :: linorder
   by (intro_classes, rule real_le_linear)
 
 
-subsection{*Theorems About the Ordering*}
+lemma real_add_left_mono: "x \<le> y ==> z + x \<le> z + (y::real)"
+apply (auto simp add: real_le_eq_diff [of x] real_le_eq_diff [of "z+x"])
+apply (subgoal_tac "z + x - (z + y) = (z + -z) + (x - y)")
+ prefer 2 apply (simp add: diff_minus add_ac, simp) 
+done
 
-lemma real_gt_zero_preal_Ex: "(0 < x) = (\<exists>y. x = real_of_preal y)"
-apply (auto simp add: real_of_preal_zero_less)
-apply (cut_tac x = x in real_of_preal_trichotomy)
-apply (blast elim!: real_less_irrefl real_of_preal_not_minus_gt_zero [THEN notE])
+
+lemma real_minus_zero_le_iff2: "(-R \<le> 0) = (0 \<le> (R::real))"
+apply (rule_tac z = R in eq_Abs_REAL)
+apply (force simp add: real_le_def real_zero_def real_minus preal_add_ac 
+                       preal_cancels preal_strange_le_eq)
 done
 
-lemma real_gt_preal_preal_Ex:
-     "real_of_preal z < x ==> \<exists>y. x = real_of_preal y"
-by (blast dest!: real_of_preal_zero_less [THEN real_less_trans]
-             intro: real_gt_zero_preal_Ex [THEN iffD1])
+lemma real_minus_zero_less_iff: "(0 < -R) = (R < (0::real))"
+by (simp add: linorder_not_le [symmetric] real_minus_zero_le_iff2) 
+
+lemma real_minus_zero_less_iff2: "(-R < 0) = ((0::real) < R)"
+by (simp add: linorder_not_le [symmetric] real_minus_zero_le_iff) 
 
-lemma real_ge_preal_preal_Ex:
-     "real_of_preal z \<le> x ==> \<exists>y. x = real_of_preal y"
-by (blast dest: order_le_imp_less_or_eq real_gt_preal_preal_Ex)
+lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)"
+by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
+
+text{*Used a few times in Lim and Transcendental*}
+lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))"
+by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
 
-lemma real_less_all_preal: "y \<le> 0 ==> \<forall>x. y < real_of_preal x"
-by (auto elim: order_le_imp_less_or_eq [THEN disjE] 
-            intro: real_of_preal_zero_less [THEN [2] real_less_trans] 
-            simp add: real_of_preal_zero_less)
-
-lemma real_less_all_real2: "~ 0 < y ==> \<forall>x. y < real_of_preal x"
-by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1])
+text{*Handles other strange cases that arise in these proofs.*}
+lemma forall_imp_less: "\<forall>u v. u \<le> v \<longrightarrow> x + v \<noteq> u + (y::preal) ==> y < x";
+apply (drule_tac x=x in spec) 
+apply (drule_tac x=y in spec) 
+apply (simp add: preal_add_commute linorder_not_le) 
+done
 
-lemma real_of_preal_le_iff:
-     "(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)"
-by (auto intro!: preal_le_iff_less_or_eq [THEN iffD1]  
-          simp add: linorder_not_less [symmetric])
-
-
-subsection{*Monotonicity of Addition*}
+text{*The arithmetic decision procedure is not set up for type preal.*}
+lemma preal_eq_le_imp_le:
+  assumes eq: "a+b = c+d" and le: "c \<le> a"
+  shows "b \<le> (d::preal)"
+proof -
+  have "c+d \<le> a+d" by (simp add: prems preal_cancels)
+  hence "a+b \<le> a+d" by (simp add: prems)
+  thus "b \<le> d" by (simp add: preal_cancels)
+qed
 
 lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y"
-apply (auto simp add: real_gt_zero_preal_Ex)
-apply (rule_tac x = "y*ya" in exI)
-apply (simp (no_asm_use) add: real_of_preal_mult)
-done
-
-(*Alternative definition for real_less*)
-lemma real_less_add_positive_left_Ex: "R < S ==> \<exists>T::real. 0 < T & R + T = S"
-apply (rule_tac x = R in real_of_preal_trichotomyE)
-apply (rule_tac [!] x = S in real_of_preal_trichotomyE)
-apply (auto dest!: preal_less_add_left_Ex 
-        simp add: real_of_preal_not_minus_gt_all real_of_preal_add
-                real_of_preal_not_less_zero real_less_not_refl 
-             real_of_preal_not_minus_gt_zero real_of_preal_minus_less_rev_iff)
-apply (rule real_of_preal_zero_less) 
-apply (rule_tac [1] x = "real_of_preal m+real_of_preal ma" in exI)
-apply (rule_tac [2] x = "real_of_preal D" in exI)
-apply (auto simp add: real_of_preal_minus_less_rev_iff real_of_preal_zero_less
-                 real_of_preal_sum_zero_less real_add_assoc)
-apply (simp add: real_add_assoc [symmetric])
-done
-
-lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))"
-apply (drule real_less_add_positive_left_Ex)
-apply (auto simp add: add_ac)
-done
-
-lemma real_lemma_change_eq_subj: "!!S::real. T = S + W ==> S = T + (-W)"
-by (simp add: add_ac)
-
-(* FIXME: long! *)
-lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)"
-apply (rule ccontr)
-apply (drule linorder_not_less [THEN iffD1, THEN real_le_imp_less_or_eq])
-apply (auto simp add: real_less_not_refl)
-apply (drule real_less_add_positive_left_Ex, clarify, simp)
-apply (drule real_lemma_change_eq_subj, auto)
-apply (drule real_less_sum_gt_zero)
-apply (auto elim: real_less_asym simp add: add_left_commute [of W] add_ac)
+apply (simp add: linorder_not_le [symmetric])
+  --{*Reduce to the (simpler) @{text "\<le>"} relation *}
+apply (rule_tac z = x in eq_Abs_REAL)
+apply (rule_tac z = y in eq_Abs_REAL)
+apply (auto simp add: real_zero_def real_le_def real_mult preal_add_ac 
+                      preal_cancels preal_strange_le_eq)
+apply (drule preal_eq_le_imp_le, assumption)
+apply (auto  dest!: forall_imp_less less_add_left_Ex 
+     simp add: preal_add_ac preal_mult_ac 
+         preal_add_mult_distrib preal_add_mult_distrib2)
+apply (insert preal_self_less_add_right)
+apply (simp add: linorder_not_le [symmetric])
 done
 
 lemma real_mult_less_mono2: "[| (0::real) < z; x < y |] ==> z * x < z * y"
@@ -794,57 +615,139 @@
 apply (simp add: right_distrib)
 done
 
-lemma real_less_sum_gt_0_iff: "(0 < S + (-W::real)) = (W < S)"
-by (blast intro: real_less_sum_gt_zero real_sum_gt_zero_less)
-
-lemma real_less_eq_diff: "(x<y) = (x-y < (0::real))"
-apply (unfold real_diff_def)
-apply (subst real_minus_zero_less_iff [symmetric])
-apply (simp add: real_add_commute real_less_sum_gt_0_iff)
-done
-
-lemma real_less_eqI: "(x::real) - y = x' - y' ==> (x<y) = (x'<y')"
-apply (subst real_less_eq_diff)
-apply (rule_tac y1 = y in real_less_eq_diff [THEN ssubst], simp)
+text{*lemma for proving @{term "0<(1::real)"}*}
+lemma real_zero_le_one: "0 \<le> (1::real)"
+apply (auto simp add: real_zero_def real_one_def real_le_def preal_add_ac 
+                      preal_cancels)
+apply (rule_tac x="preal_of_rat 1 + preal_of_rat 1" in exI) 
+apply (rule_tac x="preal_of_rat 1" in exI) 
+apply (auto simp add: preal_add_ac preal_self_less_add_left order_less_imp_le)
 done
 
-lemma real_le_eqI: "(x::real) - y = x' - y' ==> (y\<le>x) = (y'\<le>x')"
-apply (drule real_less_eqI)
-apply (simp add: real_le_def)
-done
-
-lemma real_add_left_mono: "x \<le> y ==> z + x \<le> z + (y::real)"
-apply (rule real_le_eqI [THEN iffD1]) 
- prefer 2 apply assumption
-apply (simp add: real_diff_def add_ac)
-done
-
-
 subsection{*The Reals Form an Ordered Field*}
 
 instance real :: ordered_field
 proof
   fix x y z :: real
-  show "0 < (1::real)" 
-    by (force intro: real_gt_zero_preal_Ex [THEN iffD2]
-              simp add: real_one_def real_of_preal_def)
+  show "0 < (1::real)"
+    by (simp add: real_less_def real_zero_le_one real_zero_not_eq_one)  
   show "x \<le> y ==> z + x \<le> z + y" by (rule real_add_left_mono)
   show "x < y ==> 0 < z ==> z * x < z * y" by (simp add: real_mult_less_mono2)
   show "\<bar>x\<bar> = (if x < 0 then -x else x)"
     by (auto dest: order_le_less_trans simp add: real_abs_def linorder_not_le)
 qed
 
-text{*These two need to be proved in @{text Ring_and_Field}*}
+
+
+text{*The function @{term real_of_preal} requires many proofs, but it seems
+to be essential for proving completeness of the reals from that of the
+positive reals.*}
+
+lemma real_of_preal_add:
+     "real_of_preal ((x::preal) + y) = real_of_preal x + real_of_preal y"
+by (simp add: real_of_preal_def real_add preal_add_mult_distrib preal_mult_1 
+              preal_add_ac)
+
+lemma real_of_preal_mult:
+     "real_of_preal ((x::preal) * y) = real_of_preal x* real_of_preal y"
+by (simp add: real_of_preal_def real_mult preal_add_mult_distrib2
+              preal_mult_1 preal_mult_1_right preal_add_ac preal_mult_ac)
+
+
+text{*Gleason prop 9-4.4 p 127*}
+lemma real_of_preal_trichotomy:
+      "\<exists>m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)"
+apply (unfold real_of_preal_def real_zero_def)
+apply (rule_tac z = x in eq_Abs_REAL)
+apply (auto simp add: real_minus preal_add_ac)
+apply (cut_tac x = x and y = y in linorder_less_linear)
+apply (auto dest!: less_add_left_Ex simp add: preal_add_assoc [symmetric])
+apply (auto simp add: preal_add_commute)
+done
+
+lemma real_of_preal_leD:
+      "real_of_preal m1 \<le> real_of_preal m2 ==> m1 \<le> m2"
+apply (unfold real_of_preal_def)
+apply (auto simp add: real_le_def preal_add_ac)
+apply (auto simp add: preal_add_assoc [symmetric] preal_add_right_cancel_iff)
+apply (auto simp add: preal_add_ac preal_add_le_cancel_left)
+done
+
+lemma real_of_preal_lessI: "m1 < m2 ==> real_of_preal m1 < real_of_preal m2"
+by (auto simp add: real_of_preal_leD linorder_not_le [symmetric])
+
+lemma real_of_preal_lessD:
+      "real_of_preal m1 < real_of_preal m2 ==> m1 < m2"
+apply (auto simp add: real_less_def)
+apply (drule real_of_preal_leD) 
+apply (auto simp add: order_le_less) 
+done
+
+lemma real_of_preal_less_iff [simp]:
+     "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)"
+by (blast intro: real_of_preal_lessI real_of_preal_lessD)
+
+lemma real_of_preal_le_iff:
+     "(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)"
+by (simp add: linorder_not_less [symmetric]) 
+
+lemma real_of_preal_zero_less: "0 < real_of_preal m"
+apply (auto simp add: real_zero_def real_of_preal_def real_less_def real_le_def
+            preal_add_ac preal_cancels)
+apply (simp_all add: preal_add_assoc [symmetric] preal_cancels)
+apply (blast intro: preal_self_less_add_left order_less_imp_le)
+apply (insert preal_not_eq_self [of "preal_of_rat 1" m]) 
+apply (simp add: preal_add_ac) 
+done
+
+lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0"
+by (simp add: real_of_preal_zero_less)
+
+lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m"
+apply (cut_tac real_of_preal_minus_less_zero)
+apply (fast dest: order_less_trans)
+done
+
+
+subsection{*Theorems About the Ordering*}
+
+text{*obsolete but used a lot*}
+
+lemma real_not_refl2: "x < y ==> x \<noteq> (y::real)"
+by blast 
+
+lemma real_le_imp_less_or_eq: "!!(x::real). x \<le> y ==> x < y | x = y"
+by (simp add: order_le_less)
+
+lemma real_gt_zero_preal_Ex: "(0 < x) = (\<exists>y. x = real_of_preal y)"
+apply (auto simp add: real_of_preal_zero_less)
+apply (cut_tac x = x in real_of_preal_trichotomy)
+apply (blast elim!: real_of_preal_not_minus_gt_zero [THEN notE])
+done
+
+lemma real_gt_preal_preal_Ex:
+     "real_of_preal z < x ==> \<exists>y. x = real_of_preal y"
+by (blast dest!: real_of_preal_zero_less [THEN order_less_trans]
+             intro: real_gt_zero_preal_Ex [THEN iffD1])
+
+lemma real_ge_preal_preal_Ex:
+     "real_of_preal z \<le> x ==> \<exists>y. x = real_of_preal y"
+by (blast dest: order_le_imp_less_or_eq real_gt_preal_preal_Ex)
+
+lemma real_less_all_preal: "y \<le> 0 ==> \<forall>x. y < real_of_preal x"
+by (auto elim: order_le_imp_less_or_eq [THEN disjE] 
+            intro: real_of_preal_zero_less [THEN [2] order_less_trans] 
+            simp add: real_of_preal_zero_less)
+
+lemma real_less_all_real2: "~ 0 < y ==> \<forall>x. y < real_of_preal x"
+by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1])
+
 lemma real_add_less_le_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::real)"
-apply (erule add_strict_right_mono [THEN order_less_le_trans])
-apply (erule add_left_mono) 
-done
+  by (rule Ring_and_Field.add_less_le_mono)
 
 lemma real_add_le_less_mono:
      "!!z z'::real. [| w'\<le>w; z'<z |] ==> w' + z' < w + z"
-apply (erule add_right_mono [THEN order_le_less_trans])
-apply (erule add_strict_left_mono) 
-done
+  by (rule Ring_and_Field.add_le_less_mono)
 
 lemma real_zero_less_one: "0 < (1::real)"
   by (rule Ring_and_Field.zero_less_one)
@@ -871,7 +774,9 @@
             simp add: Ring_and_Field.mult_less_cancel_right)
 
 lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \<le> y*z) = (x\<le>y)"
-by (auto simp add: real_le_def)
+apply (simp add: mult_le_cancel_right)
+apply (blast intro: elim: order_less_asym) 
+done
 
 lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \<le> z*y) = (x\<le>y)"
   by (force elim: order_less_asym
@@ -894,110 +799,194 @@
 apply (simp add: real_add_commute)
 done
 
-
 lemma real_add_order: "[| 0 < x; 0 < y |] ==> (0::real) < x + y"
-apply (drule add_strict_mono [of concl: 0 0], assumption)
-apply simp 
-done
+by (drule add_strict_mono [of concl: 0 0], assumption, simp)
 
 lemma real_le_add_order: "[| 0 \<le> x; 0 \<le> y |] ==> (0::real) \<le> x + y"
 apply (drule order_le_imp_less_or_eq)+
 apply (auto intro: real_add_order order_less_imp_le)
 done
 
-
-subsection{*An Embedding of the Naturals in the Reals*}
+lemma real_inverse_unique: "x*y = (1::real) ==> y = inverse x"
+apply (case_tac "x \<noteq> 0")
+apply (rule_tac c1 = x in real_mult_left_cancel [THEN iffD1], auto)
+done
 
-lemma real_of_posnat_one: "real_of_posnat 0 = (1::real)"
-by (simp add: real_of_posnat_def pnat_one_iff [symmetric]
-              real_of_preal_def symmetric real_one_def)
+lemma real_inverse_gt_one: "[| (0::real) < x; x < 1 |] ==> 1 < inverse x"
+by (auto dest: less_imp_inverse_less)
 
-lemma real_of_posnat_two: "real_of_posnat (Suc 0) = (1::real) + (1::real)"
-by (simp add: real_of_posnat_def real_of_preal_def real_one_def pnat_two_eq
-                 real_add
-            prat_of_pnat_add [symmetric] preal_of_prat_add [symmetric]
-            pnat_add_ac)
+lemma real_mult_self_sum_ge_zero: "(0::real) \<le> x*x + y*y"
+proof -
+  have "0 + 0 \<le> x*x + y*y" by (blast intro: add_mono zero_le_square)
+  thus ?thesis by simp
+qed
+
 
-lemma real_of_posnat_add: 
-     "real_of_posnat n1 + real_of_posnat n2 = real_of_posnat (n1 + n2) + (1::real)"
-apply (unfold real_of_posnat_def)
-apply (simp (no_asm_use) add: real_of_posnat_one [symmetric] real_of_posnat_def real_of_preal_add [symmetric] preal_of_prat_add [symmetric] prat_of_pnat_add [symmetric] pnat_of_nat_add)
+subsection{*Embedding the Integers into the Reals*}
+
+lemma real_of_int_congruent: 
+  "congruent intrel (%p. (%(i,j). realrel ``  
+   {(preal_of_rat (rat (int(Suc i))), preal_of_rat (rat (int(Suc j))))}) p)"
+apply (simp add: congruent_def add_ac del: int_Suc, clarify)
+(*OPTION raised if only is changed to add?????????*)  
+apply (simp only: add_Suc_right zero_less_rat_of_int_iff zadd_int
+          preal_of_rat_add [symmetric] rat_of_int_add_distrib [symmetric], simp) 
 done
 
-lemma real_of_posnat_add_one:
-     "real_of_posnat (n + 1) = real_of_posnat n + (1::real)"
-apply (rule_tac a1 = " (1::real) " in add_right_cancel [THEN iffD1])
-apply (rule real_of_posnat_add [THEN subst])
-apply (simp (no_asm_use) add: real_of_posnat_two real_add_assoc)
+lemma real_of_int: 
+   "real (Abs_Integ (intrel `` {(i, j)})) =  
+      Abs_REAL(realrel ``  
+        {(preal_of_rat (rat (int(Suc i))),  
+          preal_of_rat (rat (int(Suc j))))})"
+apply (unfold real_of_int_def)
+apply (rule_tac f = Abs_REAL in arg_cong)
+apply (simp del: int_Suc
+            add: realrel_in_real [THEN Abs_REAL_inverse] 
+             UN_equiv_class [OF equiv_intrel real_of_int_congruent])
+done
+
+lemma inj_real_of_int: "inj(real :: int => real)"
+apply (rule inj_onI)
+apply (rule_tac z = x in eq_Abs_Integ)
+apply (rule_tac z = y in eq_Abs_Integ, clarify) 
+apply (simp del: int_Suc 
+            add: real_of_int zadd_int preal_of_rat_eq_iff
+               preal_of_rat_add [symmetric] rat_of_int_add_distrib [symmetric])
+done
+
+lemma real_of_int_int_zero: "real (int 0) = 0"  
+by (simp add: int_def real_zero_def real_of_int preal_add_commute)
+
+lemma real_of_int_zero [simp]: "real (0::int) = 0"  
+by (insert real_of_int_int_zero, simp)
+
+lemma real_of_one [simp]: "real (1::int) = (1::real)"
+apply (subst int_1 [symmetric])
+apply (simp add: int_def real_one_def)
+apply (simp add: real_of_int preal_of_rat_add [symmetric])  
 done
 
-lemma real_of_posnat_Suc:
-     "real_of_posnat (Suc n) = real_of_posnat n + (1::real)"
-by (subst real_of_posnat_add_one [symmetric], simp)
+lemma real_of_int_add: "real (x::int) + real y = real (x + y)"
+apply (rule_tac z = x in eq_Abs_Integ)
+apply (rule_tac z = y in eq_Abs_Integ, clarify) 
+apply (simp del: int_Suc
+            add: pos_add_strict real_of_int real_add zadd
+                 preal_of_rat_add [symmetric], simp) 
+done
+declare real_of_int_add [symmetric, simp]
+
+lemma real_of_int_minus: "-real (x::int) = real (-x)"
+apply (rule_tac z = x in eq_Abs_Integ)
+apply (auto simp add: real_of_int real_minus zminus)
+done
+declare real_of_int_minus [symmetric, simp]
+
+lemma real_of_int_diff: "real (x::int) - real y = real (x - y)"
+by (simp only: zdiff_def real_diff_def real_of_int_add real_of_int_minus)
+declare real_of_int_diff [symmetric, simp]
 
-lemma inj_real_of_posnat: "inj(real_of_posnat)"
-apply (rule inj_onI)
-apply (unfold real_of_posnat_def)
-apply (drule inj_real_of_preal [THEN injD])
-apply (drule inj_preal_of_prat [THEN injD])
-apply (drule inj_prat_of_pnat [THEN injD])
-apply (erule inj_pnat_of_nat [THEN injD])
+lemma real_of_int_mult: "real (x::int) * real y = real (x * y)"
+apply (rule_tac z = x in eq_Abs_Integ)
+apply (rule_tac z = y in eq_Abs_Integ)
+apply (rename_tac a b c d) 
+apply (simp del: int_Suc
+            add: pos_add_strict mult_pos real_of_int real_mult zmult
+                 preal_of_rat_add [symmetric] preal_of_rat_mult [symmetric])
+apply (rule_tac f=preal_of_rat in arg_cong) 
+apply (simp only: int_Suc right_distrib add_ac mult_ac zadd_int zmult_int
+        rat_of_int_add_distrib [symmetric] rat_of_int_mult_distrib [symmetric]
+        rat_inject)
+done
+declare real_of_int_mult [symmetric, simp]
+
+lemma real_of_int_Suc: "real (int (Suc n)) = real (int n) + (1::real)"
+by (simp only: real_of_one [symmetric] zadd_int add_ac int_Suc real_of_int_add)
+
+lemma real_of_int_zero_cancel [simp]: "(real x = 0) = (x = (0::int))"
+by (auto intro: inj_real_of_int [THEN injD])
+
+lemma zero_le_real_of_int: "0 \<le> real y ==> 0 \<le> (y::int)"
+apply (rule_tac z = y in eq_Abs_Integ)
+apply (simp add: real_le_def, clarify)  
+apply (rename_tac a b c d) 
+apply (simp del: int_Suc zdiff_def [symmetric]
+            add: real_zero_def real_of_int zle_def zless_def zdiff_def zadd
+                 zminus neg_def preal_add_ac preal_cancels)
+apply (drule sym, drule preal_eq_le_imp_le, assumption) 
+apply (simp del: int_Suc add: preal_of_rat_le_iff)
 done
 
+lemma real_of_int_le_cancel:
+  assumes le: "real (x::int) \<le> real y"
+  shows "x \<le> y"
+proof -
+  have "real x - real x \<le> real y - real x" using le
+    by (simp only: diff_minus add_le_cancel_right) 
+  hence "0 \<le> real y - real x" by simp
+  hence "0 \<le> y - x" by (simp only: real_of_int_diff zero_le_real_of_int) 
+  hence "0 + x \<le> (y - x) + x" by (simp only: add_le_cancel_right) 
+  thus  "x \<le> y" by simp 
+qed
+
+lemma real_of_int_inject [iff]: "(real (x::int) = real y) = (x = y)"
+by (blast dest!: inj_real_of_int [THEN injD])
+
+lemma real_of_int_less_cancel: "real (x::int) < real y ==> x < y"
+by (auto simp add: order_less_le real_of_int_le_cancel)
+
+lemma real_of_int_less_mono: "x < y ==> (real (x::int) < real y)"
+apply (simp add: linorder_not_le [symmetric])
+apply (auto dest!: real_of_int_less_cancel simp add: order_le_less)
+done
+
+lemma real_of_int_less_iff [iff]: "(real (x::int) < real y) = (x < y)"
+by (blast dest: real_of_int_less_cancel intro: real_of_int_less_mono)
+
+lemma real_of_int_le_iff [simp]: "(real (x::int) \<le> real y) = (x \<le> y)"
+by (simp add: linorder_not_less [symmetric])
+
+
+subsection{*Embedding the Naturals into the Reals*}
+
 lemma real_of_nat_zero [simp]: "real (0::nat) = 0"
-by (simp add: real_of_nat_def real_of_posnat_one)
+by (simp add: real_of_nat_def)
 
 lemma real_of_nat_one [simp]: "real (Suc 0) = (1::real)"
-by (simp add: real_of_nat_def real_of_posnat_two real_add_assoc)
+by (simp add: real_of_nat_def)
 
-lemma real_of_nat_add [simp]: 
-     "real (m + n) = real (m::nat) + real n"
-apply (simp add: real_of_nat_def add_ac)
-apply (simp add: real_of_posnat_add add_assoc [symmetric])
-apply (simp add: add_commute) 
-apply (simp add: add_assoc [symmetric])
-done
+lemma real_of_nat_add [simp]: "real (m + n) = real (m::nat) + real n"
+by (simp add: real_of_nat_def add_ac)
 
 (*Not for addsimps: often the LHS is used to represent a positive natural*)
 lemma real_of_nat_Suc: "real (Suc n) = real n + (1::real)"
-by (simp add: real_of_nat_def real_of_posnat_Suc add_ac)
+by (simp add: real_of_nat_def add_ac)
 
 lemma real_of_nat_less_iff [iff]: 
      "(real (n::nat) < real m) = (n < m)"
-by (auto simp add: real_of_nat_def real_of_posnat_def)
+by (simp add: real_of_nat_def)
 
 lemma real_of_nat_le_iff [iff]: "(real (n::nat) \<le> real m) = (n \<le> m)"
 by (simp add: linorder_not_less [symmetric])
 
 lemma inj_real_of_nat: "inj (real :: nat => real)"
 apply (rule inj_onI)
-apply (auto intro!: inj_real_of_posnat [THEN injD]
-            simp add: real_of_nat_def add_right_cancel)
+apply (simp add: real_of_nat_def)
 done
 
 lemma real_of_nat_ge_zero [iff]: "0 \<le> real (n::nat)"
-apply (induct_tac "n")
-apply (auto simp add: real_of_nat_Suc)
-apply (drule real_add_le_less_mono)
-apply (rule zero_less_one)
-apply (simp add: order_less_imp_le)
+apply (insert real_of_int_le_iff [of 0 "int n"]) 
+apply (simp add: real_of_nat_def) 
 done
 
+lemma real_of_nat_Suc_gt_zero: "0 < real (Suc n)"
+by (insert real_of_nat_less_iff [of 0 "Suc n"], simp) 
+
 lemma real_of_nat_mult [simp]: "real (m * n) = real (m::nat) * real n"
-apply (induct_tac "m")
-apply (auto simp add: real_of_nat_Suc left_distrib add_commute)
-done
+by (simp add: real_of_nat_def zmult_int [symmetric]) 
 
 lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)"
 by (auto dest: inj_real_of_nat [THEN injD])
 
-lemma real_of_nat_diff [rule_format]:
-     "n \<le> m --> real (m - n) = real (m::nat) - real n"
-apply (induct_tac "m", simp)
-apply (simp add: real_diff_def Suc_diff_le le_Suc_eq real_of_nat_Suc add_ac)
-apply (simp add: add_left_commute [of _ "- 1"]) 
-done
-
 lemma real_of_nat_zero_iff: "(real (n::nat) = 0) = (n = 0)"
   proof 
     assume "real n = 0"
@@ -1007,44 +996,33 @@
     show "n = 0 \<Longrightarrow> real n = 0" by simp
   qed
 
+lemma real_of_nat_diff: "n \<le> m ==> real (m - n) = real (m::nat) - real n"
+by (simp add: real_of_nat_def zdiff_int [symmetric])
+
 lemma real_of_nat_neg_int [simp]: "neg z ==> real (nat z) = 0"
-by (simp add: neg_nat real_of_nat_zero)
-
+by (simp add: neg_nat)
 
-lemma real_inverse_unique: "x*y = (1::real) ==> y = inverse x"
-apply (case_tac "x \<noteq> 0")
-apply (rule_tac c1 = x in real_mult_left_cancel [THEN iffD1], auto)
+lemma real_of_nat_gt_zero_cancel_iff [simp]: "(0 < real (n::nat)) = (0 < n)"
+by (rule real_of_nat_less_iff [THEN subst], auto)
+
+lemma real_of_nat_le_zero_cancel_iff [simp]: "(real (n::nat) \<le> 0) = (n = 0)"
+apply (rule real_of_nat_zero [THEN subst])
+apply (simp only: real_of_nat_le_iff, simp) 
 done
 
-lemma real_inverse_gt_one: "[| (0::real) < x; x < 1 |] ==> 1 < inverse x"
-by (auto dest: less_imp_inverse_less)
 
-lemma real_of_nat_gt_zero_cancel_iff: "(0 < real (n::nat)) = (0 < n)"
-by (rule real_of_nat_less_iff [THEN subst], auto)
-declare real_of_nat_gt_zero_cancel_iff [simp]
+lemma not_real_of_nat_less_zero [simp]: "~ real (n::nat) < 0"
+by (simp add: linorder_not_less real_of_nat_ge_zero)
 
-lemma real_of_nat_le_zero_cancel_iff: "(real (n::nat) <= 0) = (n = 0)"
-apply (rule real_of_nat_zero [THEN subst])
-apply (subst real_of_nat_le_iff, auto)
-done
-declare real_of_nat_le_zero_cancel_iff [simp]
+lemma real_of_nat_ge_zero_cancel_iff [simp]: "(0 \<le> real (n::nat)) = (0 \<le> n)"
+by (simp add: linorder_not_less)
 
-lemma not_real_of_nat_less_zero: "~ real (n::nat) < 0"
-apply (simp (no_asm) add: real_le_def [symmetric] real_of_nat_ge_zero)
-done
-declare not_real_of_nat_less_zero [simp]
+text{*Now obsolete, but used in Hyperreal/IntFloor???*}
+lemma real_of_int_real_of_nat: "real (int n) = real n"
+by (simp add: real_of_nat_def)
 
-lemma real_of_nat_ge_zero_cancel_iff [simp]: 
-      "(0 <= real (n::nat)) = (0 <= n)"
-apply (simp add: real_le_def le_def)
-done
-
-lemma real_mult_self_sum_ge_zero: "(0::real) \<le> x*x + y*y"
-proof -
-  have "0 + 0 \<le> x*x + y*y" by (blast intro: add_mono zero_le_square)
-  thus ?thesis by simp
-qed
-
+lemma real_of_nat_real_of_int: "~neg z ==> real (nat z) = real z"
+by (simp add: not_neg_eq_ge_0 real_of_nat_def)
 
 ML
 {*
@@ -1053,7 +1031,6 @@
 val real_le_def = thm "real_le_def";
 val real_diff_def = thm "real_diff_def";
 val real_divide_def = thm "real_divide_def";
-val real_of_nat_def = thm "real_of_nat_def";
 
 val preal_trans_lemma = thm"preal_trans_lemma";
 val realrel_iff = thm"realrel_iff";
@@ -1074,58 +1051,27 @@
 val real_add_zero_left = thm"real_add_zero_left";
 val real_add_zero_right = thm"real_add_zero_right";
 
-val real_less_eq_diff = thm "real_less_eq_diff";
-
 val real_mult = thm"real_mult";
 val real_mult_commute = thm"real_mult_commute";
 val real_mult_assoc = thm"real_mult_assoc";
 val real_mult_1 = thm"real_mult_1";
 val real_mult_1_right = thm"real_mult_1_right";
-val real_minus_mult_commute = thm"real_minus_mult_commute";
 val preal_le_linear = thm"preal_le_linear";
-val real_mult_inv_left = thm"real_mult_inv_left";
-val real_less_not_refl = thm"real_less_not_refl";
-val real_less_irrefl = thm"real_less_irrefl";
+val real_mult_inverse_left = thm"real_mult_inverse_left";
 val real_not_refl2 = thm"real_not_refl2";
-val preal_lemma_trans = thm"preal_lemma_trans";
-val real_less_trans = thm"real_less_trans";
-val real_less_not_sym = thm"real_less_not_sym";
-val real_less_asym = thm"real_less_asym";
 val real_of_preal_add = thm"real_of_preal_add";
 val real_of_preal_mult = thm"real_of_preal_mult";
-val real_of_preal_ExI = thm"real_of_preal_ExI";
-val real_of_preal_ExD = thm"real_of_preal_ExD";
-val real_of_preal_iff = thm"real_of_preal_iff";
 val real_of_preal_trichotomy = thm"real_of_preal_trichotomy";
-val real_of_preal_trichotomyE = thm"real_of_preal_trichotomyE";
-val real_of_preal_lessD = thm"real_of_preal_lessD";
-val real_of_preal_lessI = thm"real_of_preal_lessI";
-val real_of_preal_less_iff1 = thm"real_of_preal_less_iff1";
-val real_of_preal_minus_less_self = thm"real_of_preal_minus_less_self";
 val real_of_preal_minus_less_zero = thm"real_of_preal_minus_less_zero";
 val real_of_preal_not_minus_gt_zero = thm"real_of_preal_not_minus_gt_zero";
 val real_of_preal_zero_less = thm"real_of_preal_zero_less";
-val real_of_preal_not_less_zero = thm"real_of_preal_not_less_zero";
-val real_minus_minus_zero_less = thm"real_minus_minus_zero_less";
-val real_of_preal_sum_zero_less = thm"real_of_preal_sum_zero_less";
-val real_of_preal_minus_less_all = thm"real_of_preal_minus_less_all";
-val real_of_preal_not_minus_gt_all = thm"real_of_preal_not_minus_gt_all";
-val real_of_preal_minus_less_rev1 = thm"real_of_preal_minus_less_rev1";
-val real_of_preal_minus_less_rev2 = thm"real_of_preal_minus_less_rev2";
-val real_linear = thm"real_linear";
-val real_neq_iff = thm"real_neq_iff";
-val real_linear_less2 = thm"real_linear_less2";
 val real_le_imp_less_or_eq = thm"real_le_imp_less_or_eq";
-val real_less_or_eq_imp_le = thm"real_less_or_eq_imp_le";
-val real_le_less = thm"real_le_less";
 val real_le_refl = thm"real_le_refl";
 val real_le_linear = thm"real_le_linear";
 val real_le_trans = thm"real_le_trans";
 val real_le_anti_sym = thm"real_le_anti_sym";
 val real_less_le = thm"real_less_le";
 val real_less_sum_gt_zero = thm"real_less_sum_gt_zero";
-val real_sum_gt_zero_less = thm"real_sum_gt_zero_less";
-
 val real_gt_zero_preal_Ex = thm "real_gt_zero_preal_Ex";
 val real_gt_preal_preal_Ex = thm "real_gt_preal_preal_Ex";
 val real_ge_preal_preal_Ex = thm "real_ge_preal_preal_Ex";
@@ -1151,10 +1097,25 @@
 
 val real_mult_left_cancel = thm"real_mult_left_cancel";
 val real_mult_right_cancel = thm"real_mult_right_cancel";
-val real_of_posnat_one = thm "real_of_posnat_one";
-val real_of_posnat_two = thm "real_of_posnat_two";
-val real_of_posnat_add = thm "real_of_posnat_add";
-val real_of_posnat_add_one = thm "real_of_posnat_add_one";
+val real_inverse_unique = thm "real_inverse_unique";
+val real_inverse_gt_one = thm "real_inverse_gt_one";
+
+val real_of_int = thm"real_of_int";
+val inj_real_of_int = thm"inj_real_of_int";
+val real_of_int_zero = thm"real_of_int_zero";
+val real_of_one = thm"real_of_one";
+val real_of_int_add = thm"real_of_int_add";
+val real_of_int_minus = thm"real_of_int_minus";
+val real_of_int_diff = thm"real_of_int_diff";
+val real_of_int_mult = thm"real_of_int_mult";
+val real_of_int_Suc = thm"real_of_int_Suc";
+val real_of_int_real_of_nat = thm"real_of_int_real_of_nat";
+val real_of_nat_real_of_int = thm"real_of_nat_real_of_int";
+val real_of_int_less_cancel = thm"real_of_int_less_cancel";
+val real_of_int_inject = thm"real_of_int_inject";
+val real_of_int_less_mono = thm"real_of_int_less_mono";
+val real_of_int_less_iff = thm"real_of_int_less_iff";
+val real_of_int_le_iff = thm"real_of_int_le_iff";
 val real_of_nat_zero = thm "real_of_nat_zero";
 val real_of_nat_one = thm "real_of_nat_one";
 val real_of_nat_add = thm "real_of_nat_add";
@@ -1163,13 +1124,12 @@
 val real_of_nat_le_iff = thm "real_of_nat_le_iff";
 val inj_real_of_nat = thm "inj_real_of_nat";
 val real_of_nat_ge_zero = thm "real_of_nat_ge_zero";
+val real_of_nat_Suc_gt_zero = thm "real_of_nat_Suc_gt_zero";
 val real_of_nat_mult = thm "real_of_nat_mult";
 val real_of_nat_inject = thm "real_of_nat_inject";
 val real_of_nat_diff = thm "real_of_nat_diff";
 val real_of_nat_zero_iff = thm "real_of_nat_zero_iff";
 val real_of_nat_neg_int = thm "real_of_nat_neg_int";
-val real_inverse_unique = thm "real_inverse_unique";
-val real_inverse_gt_one = thm "real_inverse_gt_one";
 val real_of_nat_gt_zero_cancel_iff = thm "real_of_nat_gt_zero_cancel_iff";
 val real_of_nat_le_zero_cancel_iff = thm "real_of_nat_le_zero_cancel_iff";
 val not_real_of_nat_less_zero = thm "not_real_of_nat_less_zero";
--- a/src/HOL/Real/RealInt.thy	Tue Jan 27 09:44:14 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,150 +0,0 @@
-(*  Title:       RealInt.thy
-    ID:         $Id$
-    Author:      Jacques D. Fleuriot
-    Copyright:   1999 University of Edinburgh
-*)
-
-header{*Embedding the Integers into the Reals*}
-
-theory RealInt = RealDef:
-
-defs (overloaded)
-  real_of_int_def:
-   "real z == Abs_REAL(UN (i,j): Rep_Integ z. realrel ``
-		       {(preal_of_prat(prat_of_pnat(pnat_of_nat i)),
-			 preal_of_prat(prat_of_pnat(pnat_of_nat j)))})"
-
-
-lemma real_of_int_congruent: 
-  "congruent intrel (%p. (%(i,j). realrel ``  
-   {(preal_of_prat (prat_of_pnat (pnat_of_nat i)),  
-     preal_of_prat (prat_of_pnat (pnat_of_nat j)))}) p)"
-apply (unfold congruent_def)
-apply (auto simp add: pnat_of_nat_add prat_of_pnat_add [symmetric] preal_of_prat_add [symmetric])
-done
-
-lemma real_of_int: 
-   "real (Abs_Integ (intrel `` {(i, j)})) =  
-      Abs_REAL(realrel ``  
-        {(preal_of_prat (prat_of_pnat (pnat_of_nat i)),  
-          preal_of_prat (prat_of_pnat (pnat_of_nat j)))})"
-apply (unfold real_of_int_def)
-apply (rule_tac f = Abs_REAL in arg_cong)
-apply (simp add: realrel_in_real [THEN Abs_REAL_inverse] 
-             UN_equiv_class [OF equiv_intrel real_of_int_congruent])
-done
-
-lemma inj_real_of_int: "inj(real :: int => real)"
-apply (rule inj_onI)
-apply (rule_tac z = x in eq_Abs_Integ)
-apply (rule_tac z = y in eq_Abs_Integ)
-apply (auto dest!: inj_preal_of_prat [THEN injD] inj_prat_of_pnat [THEN injD]
-                   inj_pnat_of_nat [THEN injD]
-      simp add: real_of_int preal_of_prat_add [symmetric] prat_of_pnat_add [symmetric] pnat_of_nat_add)
-done
-
-lemma real_of_int_zero: "real (int 0) = 0"
-apply (unfold int_def real_zero_def)
-apply (simp add: real_of_int preal_add_commute)
-done
-
-lemma real_of_one: "real (1::int) = (1::real)"
-apply (subst int_1 [symmetric])
-apply (auto simp add: int_def real_one_def real_of_int preal_of_prat_add [symmetric] pnat_two_eq prat_of_pnat_add [symmetric] pnat_one_iff [symmetric])
-done
-
-lemma real_of_int_add: "real (x::int) + real y = real (x + y)"
-apply (rule_tac z = x in eq_Abs_Integ)
-apply (rule_tac z = y in eq_Abs_Integ)
-apply (auto simp add: real_of_int preal_of_prat_add [symmetric]
-            prat_of_pnat_add [symmetric] zadd real_add pnat_of_nat_add pnat_add_ac)
-done
-declare real_of_int_add [symmetric, simp]
-
-lemma real_of_int_minus: "-real (x::int) = real (-x)"
-apply (rule_tac z = x in eq_Abs_Integ)
-apply (auto simp add: real_of_int real_minus zminus)
-done
-declare real_of_int_minus [symmetric, simp]
-
-lemma real_of_int_diff: 
-  "real (x::int) - real y = real (x - y)"
-apply (unfold zdiff_def real_diff_def)
-apply (simp only: real_of_int_add real_of_int_minus)
-done
-declare real_of_int_diff [symmetric, simp]
-
-lemma real_of_int_mult: "real (x::int) * real y = real (x * y)"
-apply (rule_tac z = x in eq_Abs_Integ)
-apply (rule_tac z = y in eq_Abs_Integ)
-apply (auto simp add: real_of_int real_mult zmult
-         preal_of_prat_mult [symmetric] pnat_of_nat_mult 
-        prat_of_pnat_mult [symmetric] preal_of_prat_add [symmetric]
-        prat_of_pnat_add [symmetric] pnat_of_nat_add mult_ac add_ac pnat_add_ac)
-done
-declare real_of_int_mult [symmetric, simp]
-
-lemma real_of_int_Suc: "real (int (Suc n)) = real (int n) + (1::real)"
-by (simp add: real_of_one [symmetric] zadd_int add_ac)
-
-declare real_of_one [simp]
-
-(* relating two of the embeddings *)
-lemma real_of_int_real_of_nat: "real (int n) = real n"
-apply (induct_tac "n")
-apply (simp_all only: real_of_int_zero real_of_nat_zero real_of_int_Suc real_of_nat_Suc)
-done
-
-lemma real_of_nat_real_of_int: "~neg z ==> real (nat z) = real z"
-by (simp add: not_neg_eq_ge_0 real_of_int_real_of_nat [symmetric])
-
-lemma real_of_int_zero_cancel [simp]: "(real x = 0) = (x = int 0)"
-by (auto intro: inj_real_of_int [THEN injD] 
-         simp only: real_of_int_zero)
-
-lemma real_of_int_less_cancel: "real (x::int) < real y ==> x < y"
-apply (rule ccontr, drule linorder_not_less [THEN iffD1])
-apply (auto simp add: zle_iff_zadd real_of_int_add [symmetric] real_of_int_real_of_nat linorder_not_le [symmetric])
-apply (subgoal_tac "~ real y + 0 \<le> real y + real n") 
- prefer 2 apply simp 
-apply (simp only: add_le_cancel_left, simp) 
-done
-
-lemma real_of_int_inject [iff]: "(real (x::int) = real y) = (x = y)"
-by (blast dest!: inj_real_of_int [THEN injD])
-
-lemma real_of_int_less_mono: "x < y ==> (real (x::int) < real y)"
-apply (simp add: linorder_not_le [symmetric])
-apply (auto dest!: real_of_int_less_cancel simp add: order_le_less)
-done
-
-lemma real_of_int_less_iff [iff]: "(real (x::int) < real y) = (x < y)"
-by (blast dest: real_of_int_less_cancel intro: real_of_int_less_mono)
-
-lemma real_of_int_le_iff [simp]: "(real (x::int) \<le> real y) = (x \<le> y)"
-by (simp add: linorder_not_less [symmetric])
-
-ML
-{*
-val real_of_int_congruent = thm"real_of_int_congruent";
-val real_of_int = thm"real_of_int";
-val inj_real_of_int = thm"inj_real_of_int";
-val real_of_int_zero = thm"real_of_int_zero";
-val real_of_one = thm"real_of_one";
-val real_of_int_add = thm"real_of_int_add";
-val real_of_int_minus = thm"real_of_int_minus";
-val real_of_int_diff = thm"real_of_int_diff";
-val real_of_int_mult = thm"real_of_int_mult";
-val real_of_int_Suc = thm"real_of_int_Suc";
-val real_of_int_real_of_nat = thm"real_of_int_real_of_nat";
-val real_of_nat_real_of_int = thm"real_of_nat_real_of_int";
-val real_of_int_zero_cancel = thm"real_of_int_zero_cancel";
-val real_of_int_less_cancel = thm"real_of_int_less_cancel";
-val real_of_int_inject = thm"real_of_int_inject";
-val real_of_int_less_mono = thm"real_of_int_less_mono";
-val real_of_int_less_iff = thm"real_of_int_less_iff";
-val real_of_int_le_iff = thm"real_of_int_le_iff";
-*}
-
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/rat_arith.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -0,0 +1,662 @@
+(*  Title:      HOL/Real/rat_arith0.ML
+    ID:         $Id$
+    Author:     Lawrence C Paulson
+    Copyright   2004 University of Cambridge
+
+Simprocs for common factor cancellation & Rational coefficient handling
+
+Instantiation of the generic linear arithmetic package for type rat.
+*)
+
+(*FIXME DELETE*)
+val rat_mult_strict_left_mono =
+    read_instantiate_sg(sign_of (the_context())) [("a","?a::rat")] mult_strict_left_mono;
+
+val rat_mult_left_mono =
+    read_instantiate_sg(sign_of (the_context())) [("a","?a::rat")] mult_left_mono;
+
+
+val rat_number_of_def = thm "rat_number_of_def";
+val diff_rat_def = thm "diff_rat_def";
+
+val rat_of_int_zero = thm "rat_of_int_zero";
+val rat_of_int_one = thm "rat_of_int_one";
+val rat_of_int_add_distrib = thm "rat_of_int_add_distrib";
+val rat_of_int_minus_distrib = thm "rat_of_int_minus_distrib";
+val rat_of_int_diff_distrib = thm "rat_of_int_diff_distrib";
+val rat_of_int_mult_distrib = thm "rat_of_int_mult_distrib";
+val rat_inject = thm "rat_inject";
+val rat_of_int_zero_cancel = thm "rat_of_int_zero_cancel";
+val rat_of_int_less_iff = thm "rat_of_int_less_iff";
+val rat_of_int_le_iff = thm "rat_of_int_le_iff";
+
+val number_of_rat = thm "number_of_rat";
+val rat_numeral_0_eq_0 = thm "rat_numeral_0_eq_0";
+val rat_numeral_1_eq_1 = thm "rat_numeral_1_eq_1";
+val add_rat_number_of = thm "add_rat_number_of";
+val minus_rat_number_of = thm "minus_rat_number_of";
+val diff_rat_number_of = thm "diff_rat_number_of";
+val mult_rat_number_of = thm "mult_rat_number_of";
+val rat_mult_2 = thm "rat_mult_2";
+val rat_mult_2_right = thm "rat_mult_2_right";
+val eq_rat_number_of = thm "eq_rat_number_of";
+val less_rat_number_of = thm "less_rat_number_of";
+val rat_minus_1_eq_m1 = thm "rat_minus_1_eq_m1";
+val rat_mult_minus1 = thm "rat_mult_minus1";
+val rat_mult_minus1_right = thm "rat_mult_minus1_right";
+val rat_add_number_of_left = thm "rat_add_number_of_left";
+val rat_mult_number_of_left = thm "rat_mult_number_of_left";
+val rat_add_number_of_diff1 = thm "rat_add_number_of_diff1";
+val rat_add_number_of_diff2 = thm "rat_add_number_of_diff2";
+
+val rat_add_0_left = thm "rat_add_0_left";
+val rat_add_0_right = thm "rat_add_0_right";
+val rat_mult_1_left = thm "rat_mult_1_left";
+val rat_mult_1_right = thm "rat_mult_1_right";
+
+(*Maps 0 to Numeral0 and 1 to Numeral1 and -(Numeral1) to -1*)
+val rat_numeral_ss =
+    HOL_ss addsimps [rat_numeral_0_eq_0 RS sym, rat_numeral_1_eq_1 RS sym,
+                     rat_minus_1_eq_m1];
+
+fun rename_numerals th =
+    asm_full_simplify rat_numeral_ss (Thm.transfer (the_context ()) th);
+
+
+structure Rat_Numeral_Simprocs =
+struct
+
+(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in simprocs
+  isn't complicated by the abstract 0 and 1.*)
+val numeral_syms = [rat_numeral_0_eq_0 RS sym, rat_numeral_1_eq_1 RS sym];
+
+(*Utilities*)
+
+val ratT = Type("Rational.rat", []);
+
+fun mk_numeral n = HOLogic.number_of_const ratT $ HOLogic.mk_bin n;
+
+(*Decodes a binary rat constant, or 0, 1*)
+val dest_numeral = Int_Numeral_Simprocs.dest_numeral;
+val find_first_numeral = Int_Numeral_Simprocs.find_first_numeral;
+
+val zero = mk_numeral 0;
+val mk_plus = HOLogic.mk_binop "op +";
+
+val uminus_const = Const ("uminus", ratT --> ratT);
+
+(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
+fun mk_sum []        = zero
+  | mk_sum [t,u]     = mk_plus (t, u)
+  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
+
+(*this version ALWAYS includes a trailing zero*)
+fun long_mk_sum []        = zero
+  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
+
+val dest_plus = HOLogic.dest_bin "op +" ratT;
+
+(*decompose additions AND subtractions as a sum*)
+fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
+        dest_summing (pos, t, dest_summing (pos, u, ts))
+  | dest_summing (pos, Const ("op -", _) $ t $ u, ts) =
+        dest_summing (pos, t, dest_summing (not pos, u, ts))
+  | dest_summing (pos, t, ts) =
+        if pos then t::ts else uminus_const$t :: ts;
+
+fun dest_sum t = dest_summing (true, t, []);
+
+val mk_diff = HOLogic.mk_binop "op -";
+val dest_diff = HOLogic.dest_bin "op -" ratT;
+
+val one = mk_numeral 1;
+val mk_times = HOLogic.mk_binop "op *";
+
+fun mk_prod [] = one
+  | mk_prod [t] = t
+  | mk_prod (t :: ts) = if t = one then mk_prod ts
+                        else mk_times (t, mk_prod ts);
+
+val dest_times = HOLogic.dest_bin "op *" ratT;
+
+fun dest_prod t =
+      let val (t,u) = dest_times t
+      in  dest_prod t @ dest_prod u  end
+      handle TERM _ => [t];
+
+(*DON'T do the obvious simplifications; that would create special cases*)
+fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts);
+
+(*Express t as a product of (possibly) a numeral with other sorted terms*)
+fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t
+  | dest_coeff sign t =
+    let val ts = sort Term.term_ord (dest_prod t)
+        val (n, ts') = find_first_numeral [] ts
+                          handle TERM _ => (1, ts)
+    in (sign*n, mk_prod ts') end;
+
+(*Find first coefficient-term THAT MATCHES u*)
+fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
+  | find_first_coeff past u (t::terms) =
+        let val (n,u') = dest_coeff 1 t
+        in  if u aconv u' then (n, rev past @ terms)
+                          else find_first_coeff (t::past) u terms
+        end
+        handle TERM _ => find_first_coeff (t::past) u terms;
+
+
+(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
+val add_0s  = map rename_numerals [rat_add_0_left, rat_add_0_right];
+val mult_1s = map rename_numerals [rat_mult_1_left, rat_mult_1_right] @
+              [rat_mult_minus1, rat_mult_minus1_right];
+
+(*To perform binary arithmetic*)
+val bin_simps =
+    [rat_numeral_0_eq_0 RS sym, rat_numeral_1_eq_1 RS sym,
+     add_rat_number_of, rat_add_number_of_left, minus_rat_number_of,
+     diff_rat_number_of, mult_rat_number_of, rat_mult_number_of_left] @
+    bin_arith_simps @ bin_rel_simps;
+
+(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
+  during re-arrangement*)
+val non_add_bin_simps = 
+    bin_simps \\ [rat_add_number_of_left, add_rat_number_of];
+
+(*To evaluate binary negations of coefficients*)
+val rat_minus_simps = NCons_simps @
+                   [rat_minus_1_eq_m1, minus_rat_number_of,
+                    bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min,
+                    bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min];
+
+(*To let us treat subtraction as addition*)
+val diff_simps = [diff_rat_def, minus_add_distrib, minus_minus];
+
+(*to extract again any uncancelled minuses*)
+val rat_minus_from_mult_simps =
+    [minus_minus, minus_mult_left RS sym, minus_mult_right RS sym];
+
+(*combine unary minus with numeric literals, however nested within a product*)
+val rat_mult_minus_simps =
+    [mult_assoc, minus_mult_left, minus_mult_commute];
+
+(*Apply the given rewrite (if present) just once*)
+fun trans_tac None      = all_tac
+  | trans_tac (Some th) = ALLGOALS (rtac (th RS trans));
+
+(*Final simplification: cancel + and *  *)
+val simplify_meta_eq =
+    Int_Numeral_Simprocs.simplify_meta_eq
+         [add_0, add_0_right,
+          mult_zero_left, mult_zero_right, mult_1, mult_1_right];
+
+fun prep_simproc (name, pats, proc) =
+  Simplifier.simproc (Theory.sign_of (the_context ())) name pats proc;
+
+structure CancelNumeralsCommon =
+  struct
+  val mk_sum            = mk_sum
+  val dest_sum          = dest_sum
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff 1
+  val find_first_coeff  = find_first_coeff []
+  val trans_tac         = trans_tac
+  val norm_tac =
+     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
+                                         rat_minus_simps@add_ac))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@rat_mult_minus_simps))
+     THEN ALLGOALS
+              (simp_tac (HOL_ss addsimps rat_minus_from_mult_simps@
+                                         add_ac@mult_ac))
+  val numeral_simp_tac  = ALLGOALS (simp_tac (HOL_ss addsimps add_0s@bin_simps))
+  val simplify_meta_eq  = simplify_meta_eq
+  end;
+
+
+structure EqCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Bin_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" ratT
+  val bal_add1 = eq_add_iff1 RS trans
+  val bal_add2 = eq_add_iff2 RS trans
+);
+
+structure LessCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Bin_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel "op <"
+  val dest_bal = HOLogic.dest_bin "op <" ratT
+  val bal_add1 = less_add_iff1 RS trans
+  val bal_add2 = less_add_iff2 RS trans
+);
+
+structure LeCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Bin_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel "op <="
+  val dest_bal = HOLogic.dest_bin "op <=" ratT
+  val bal_add1 = le_add_iff1 RS trans
+  val bal_add2 = le_add_iff2 RS trans
+);
+
+val cancel_numerals =
+  map prep_simproc
+   [("rateq_cancel_numerals",
+     ["(l::rat) + m = n", "(l::rat) = m + n",
+      "(l::rat) - m = n", "(l::rat) = m - n",
+      "(l::rat) * m = n", "(l::rat) = m * n"],
+     EqCancelNumerals.proc),
+    ("ratless_cancel_numerals",
+     ["(l::rat) + m < n", "(l::rat) < m + n",
+      "(l::rat) - m < n", "(l::rat) < m - n",
+      "(l::rat) * m < n", "(l::rat) < m * n"],
+     LessCancelNumerals.proc),
+    ("ratle_cancel_numerals",
+     ["(l::rat) + m <= n", "(l::rat) <= m + n",
+      "(l::rat) - m <= n", "(l::rat) <= m - n",
+      "(l::rat) * m <= n", "(l::rat) <= m * n"],
+     LeCancelNumerals.proc)];
+
+
+structure CombineNumeralsData =
+  struct
+  val add               = op + : int*int -> int
+  val mk_sum            = long_mk_sum    (*to work for e.g. 2*x + 3*x *)
+  val dest_sum          = dest_sum
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff 1
+  val left_distrib      = combine_common_factor RS trans
+  val prove_conv        = Bin_Simprocs.prove_conv_nohyps
+  val trans_tac         = trans_tac
+  val norm_tac =
+     ALLGOALS (simp_tac (HOL_ss addsimps numeral_syms@add_0s@mult_1s@
+                                   diff_simps@rat_minus_simps@add_ac))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@rat_mult_minus_simps))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps rat_minus_from_mult_simps@
+                                              add_ac@mult_ac))
+  val numeral_simp_tac  = ALLGOALS
+                    (simp_tac (HOL_ss addsimps add_0s@bin_simps))
+  val simplify_meta_eq  =
+        Int_Numeral_Simprocs.simplify_meta_eq (add_0s@mult_1s)
+  end;
+
+structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
+
+val combine_numerals =
+  prep_simproc ("rat_combine_numerals", ["(i::rat) + j", "(i::rat) - j"], CombineNumerals.proc);
+
+
+(** Declarations for ExtractCommonTerm **)
+
+(*this version ALWAYS includes a trailing one*)
+fun long_mk_prod []        = one
+  | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts);
+
+(*Find first term that matches u*)
+fun find_first past u []         = raise TERM("find_first", [])
+  | find_first past u (t::terms) =
+        if u aconv t then (rev past @ terms)
+        else find_first (t::past) u terms
+        handle TERM _ => find_first (t::past) u terms;
+
+(*Final simplification: cancel + and *  *)
+fun cancel_simplify_meta_eq cancel_th th =
+    Int_Numeral_Simprocs.simplify_meta_eq
+        [rat_mult_1_left, rat_mult_1_right]
+        (([th, cancel_th]) MRS trans);
+
+(*** Making constant folding work for 0 and 1 too ***)
+
+structure RatAbstractNumeralsData =
+  struct
+  val dest_eq         = HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of
+  val is_numeral      = Bin_Simprocs.is_numeral
+  val numeral_0_eq_0  = rat_numeral_0_eq_0
+  val numeral_1_eq_1  = rat_numeral_1_eq_1
+  val prove_conv      = Bin_Simprocs.prove_conv_nohyps_novars
+  fun norm_tac simps  = ALLGOALS (simp_tac (HOL_ss addsimps simps))
+  val simplify_meta_eq = Bin_Simprocs.simplify_meta_eq
+  end
+
+structure RatAbstractNumerals = AbstractNumeralsFun (RatAbstractNumeralsData)
+
+(*For addition, we already have rules for the operand 0.
+  Multiplication is omitted because there are already special rules for
+  both 0 and 1 as operands.  Unary minus is trivial, just have - 1 = -1.
+  For the others, having three patterns is a compromise between just having
+  one (many spurious calls) and having nine (just too many!) *)
+val eval_numerals =
+  map prep_simproc
+   [("rat_add_eval_numerals",
+     ["(m::rat) + 1", "(m::rat) + number_of v"],
+     RatAbstractNumerals.proc add_rat_number_of),
+    ("rat_diff_eval_numerals",
+     ["(m::rat) - 1", "(m::rat) - number_of v"],
+     RatAbstractNumerals.proc diff_rat_number_of),
+    ("rat_eq_eval_numerals",
+     ["(m::rat) = 0", "(m::rat) = 1", "(m::rat) = number_of v"],
+     RatAbstractNumerals.proc eq_rat_number_of),
+    ("rat_less_eval_numerals",
+     ["(m::rat) < 0", "(m::rat) < 1", "(m::rat) < number_of v"],
+     RatAbstractNumerals.proc less_rat_number_of),
+    ("rat_le_eval_numerals",
+     ["(m::rat) <= 0", "(m::rat) <= 1", "(m::rat) <= number_of v"],
+     RatAbstractNumerals.proc le_number_of_eq_not_less)]
+
+end;
+
+
+Addsimprocs Rat_Numeral_Simprocs.eval_numerals;
+Addsimprocs Rat_Numeral_Simprocs.cancel_numerals;
+Addsimprocs [Rat_Numeral_Simprocs.combine_numerals];
+
+
+
+(** Constant folding for rat plus and times **)
+
+(*We do not need
+    structure Rat_Plus_Assoc = Assoc_Fold (Rat_Plus_Assoc_Data);
+  because combine_numerals does the same thing*)
+
+structure Rat_Times_Assoc_Data : ASSOC_FOLD_DATA =
+struct
+  val ss                = HOL_ss
+  val eq_reflection     = eq_reflection
+  val sg_ref    = Sign.self_ref (Theory.sign_of (the_context ()))
+  val T      = Rat_Numeral_Simprocs.ratT
+  val plus   = Const ("op *", [Rat_Numeral_Simprocs.ratT,Rat_Numeral_Simprocs.ratT] ---> Rat_Numeral_Simprocs.ratT)
+  val add_ac = mult_ac
+end;
+
+structure Rat_Times_Assoc = Assoc_Fold (Rat_Times_Assoc_Data);
+
+Addsimprocs [Rat_Times_Assoc.conv];
+
+
+(****Common factor cancellation****)
+
+(*To quote from Provers/Arith/cancel_numeral_factor.ML:
+
+This simproc Cancels common coefficients in balanced expressions:
+
+     u*#m ~~ u'*#m'  ==  #n*u ~~ #n'*u'
+
+where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /)
+and d = gcd(m,m') and n=m/d and n'=m'/d.
+*)
+
+local
+  open Rat_Numeral_Simprocs
+in
+
+val rel_rat_number_of = [eq_rat_number_of, less_rat_number_of,
+                          le_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 rat_minus_from_mult_simps @ mult_1s))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@rat_mult_minus_simps))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps mult_ac))
+  val numeral_simp_tac  =
+         ALLGOALS (simp_tac (HOL_ss addsimps rel_rat_number_of@bin_simps))
+  val simplify_meta_eq  = simplify_meta_eq
+  end
+
+structure DivCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Bin_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binop "HOL.divide"
+  val dest_bal = HOLogic.dest_bin "HOL.divide" Rat_Numeral_Simprocs.ratT
+  val cancel = mult_divide_cancel_left RS trans
+  val neg_exchanges = false
+)
+
+structure EqCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Bin_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" Rat_Numeral_Simprocs.ratT
+  val cancel = mult_cancel_left RS trans
+  val neg_exchanges = false
+)
+
+structure LessCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Bin_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel "op <"
+  val dest_bal = HOLogic.dest_bin "op <" Rat_Numeral_Simprocs.ratT
+  val cancel = mult_less_cancel_left RS trans
+  val neg_exchanges = true
+)
+
+structure LeCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Bin_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel "op <="
+  val dest_bal = HOLogic.dest_bin "op <=" Rat_Numeral_Simprocs.ratT
+  val cancel = mult_le_cancel_left RS trans
+  val neg_exchanges = true
+)
+
+val rat_cancel_numeral_factors_relations =
+  map prep_simproc
+   [("rateq_cancel_numeral_factor",
+     ["(l::rat) * m = n", "(l::rat) = m * n"],
+     EqCancelNumeralFactor.proc),
+    ("ratless_cancel_numeral_factor",
+     ["(l::rat) * m < n", "(l::rat) < m * n"],
+     LessCancelNumeralFactor.proc),
+    ("ratle_cancel_numeral_factor",
+     ["(l::rat) * m <= n", "(l::rat) <= m * n"],
+     LeCancelNumeralFactor.proc)]
+
+val rat_cancel_numeral_factors_divide = prep_simproc
+        ("ratdiv_cancel_numeral_factor",
+         ["((l::rat) * m) / n", "(l::rat) / (m * n)",
+          "((number_of v)::rat) / (number_of w)"],
+         DivCancelNumeralFactor.proc)
+
+val rat_cancel_numeral_factors =
+    rat_cancel_numeral_factors_relations @
+    [rat_cancel_numeral_factors_divide]
+
+end;
+
+Addsimprocs rat_cancel_numeral_factors;
+
+
+(*examples:
+print_depth 22;
+set timing;
+set trace_simp;
+fun test s = (Goal s; by (Simp_tac 1));
+
+test "0 <= (y::rat) * -2";
+test "9*x = 12 * (y::rat)";
+test "(9*x) / (12 * (y::rat)) = z";
+test "9*x < 12 * (y::rat)";
+test "9*x <= 12 * (y::rat)";
+
+test "-99*x = 132 * (y::rat)";
+test "(-99*x) / (132 * (y::rat)) = z";
+test "-99*x < 132 * (y::rat)";
+test "-99*x <= 132 * (y::rat)";
+
+test "999*x = -396 * (y::rat)";
+test "(999*x) / (-396 * (y::rat)) = z";
+test "999*x < -396 * (y::rat)";
+test "999*x <= -396 * (y::rat)";
+
+test  "(- ((2::rat) * x) <= 2 * y)";
+test "-99*x = -81 * (y::rat)";
+test "(-99*x) / (-81 * (y::rat)) = z";
+test "-99*x <= -81 * (y::rat)";
+test "-99*x < -81 * (y::rat)";
+
+test "-2 * x = -1 * (y::rat)";
+test "-2 * x = -(y::rat)";
+test "(-2 * x) / (-1 * (y::rat)) = z";
+test "-2 * x < -(y::rat)";
+test "-2 * x <= -1 * (y::rat)";
+test "-x < -23 * (y::rat)";
+test "-x <= -23 * (y::rat)";
+*)
+
+
+(** Declarations for ExtractCommonTerm **)
+
+local
+  open Rat_Numeral_Simprocs
+in
+
+structure CancelFactorCommon =
+  struct
+  val mk_sum            = long_mk_prod
+  val dest_sum          = dest_prod
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff
+  val find_first        = find_first []
+  val trans_tac         = trans_tac
+  val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@mult_ac))
+  end;
+
+structure EqCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Bin_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" Rat_Numeral_Simprocs.ratT
+  val simplify_meta_eq  = cancel_simplify_meta_eq mult_cancel_left
+);
+
+
+structure DivideCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Bin_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binop "HOL.divide"
+  val dest_bal = HOLogic.dest_bin "HOL.divide" Rat_Numeral_Simprocs.ratT
+  val simplify_meta_eq  = cancel_simplify_meta_eq mult_divide_cancel_eq_if
+);
+
+val rat_cancel_factor =
+  map prep_simproc
+   [("rat_eq_cancel_factor", ["(l::rat) * m = n", "(l::rat) = m * n"], EqCancelFactor.proc),
+    ("rat_divide_cancel_factor", ["((l::rat) * m) / n", "(l::rat) / (m * n)"],
+     DivideCancelFactor.proc)];
+
+end;
+
+Addsimprocs rat_cancel_factor;
+
+
+(*examples:
+print_depth 22;
+set timing;
+set trace_simp;
+fun test s = (Goal s; by (Asm_simp_tac 1));
+
+test "x*k = k*(y::rat)";
+test "k = k*(y::rat)";
+test "a*(b*c) = (b::rat)";
+test "a*(b*c) = d*(b::rat)*(x*a)";
+
+
+test "(x*k) / (k*(y::rat)) = (uu::rat)";
+test "(k) / (k*(y::rat)) = (uu::rat)";
+test "(a*(b*c)) / ((b::rat)) = (uu::rat)";
+test "(a*(b*c)) / (d*(b::rat)*(x*a)) = (uu::rat)";
+
+(*FIXME: what do we do about this?*)
+test "a*(b*c)/(y*z) = d*(b::rat)*(x*a)/z";
+*)
+
+
+
+(****Instantiation of the generic linear arithmetic package****)
+
+val add_zero_left = thm"Ring_and_Field.add_0";
+val add_zero_right = thm"Ring_and_Field.add_0_right";
+
+
+local
+
+(* reduce contradictory <= to False *)
+val add_rules = 
+    [order_less_irrefl, rat_numeral_0_eq_0, rat_numeral_1_eq_1,
+     rat_minus_1_eq_m1, 
+     add_rat_number_of, minus_rat_number_of, diff_rat_number_of,
+     mult_rat_number_of, eq_rat_number_of, less_rat_number_of,
+     le_number_of_eq_not_less, diff_minus,
+     minus_add_distrib, minus_minus, mult_assoc, minus_zero,
+     add_zero_left, add_zero_right, left_minus, right_minus,
+     mult_zero_left, mult_zero_right, mult_1, mult_1_right,
+     minus_mult_left RS sym, minus_mult_right RS sym];
+
+val simprocs = [Rat_Times_Assoc.conv, Rat_Numeral_Simprocs.combine_numerals,
+                rat_cancel_numeral_factors_divide]@
+               Rat_Numeral_Simprocs.cancel_numerals @
+               Rat_Numeral_Simprocs.eval_numerals;
+
+val mono_ss = simpset() addsimps
+                [add_mono,add_strict_mono,add_less_le_mono,add_le_less_mono];
+
+val add_mono_thms_rat =
+  map (fn s => prove_goal (the_context ()) s
+                 (fn prems => [cut_facts_tac prems 1, asm_simp_tac mono_ss 1]))
+    ["(i <= j) & (k <= l) ==> i + k <= j + (l::rat)",
+     "(i  = j) & (k <= l) ==> i + k <= j + (l::rat)",
+     "(i <= j) & (k  = l) ==> i + k <= j + (l::rat)",
+     "(i  = j) & (k  = l) ==> i + k  = j + (l::rat)",
+     "(i < j) & (k = l)   ==> i + k < j + (l::rat)",
+     "(i = j) & (k < l)   ==> i + k < j + (l::rat)",
+     "(i < j) & (k <= l)  ==> i + k < j + (l::rat)",
+     "(i <= j) & (k < l)  ==> i + k < j + (l::rat)",
+     "(i < j) & (k < l)   ==> i + k < j + (l::rat)"];
+
+fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
+
+val rat_mult_mono_thms =
+ [(rat_mult_strict_left_mono,
+   cvar(rat_mult_strict_left_mono, hd(tl(prems_of rat_mult_strict_left_mono)))),
+  (rat_mult_left_mono,
+   cvar(rat_mult_left_mono, hd(tl(prems_of rat_mult_left_mono))))]
+
+(* reduce contradictory <= to False *)
+val simps = [True_implies_equals,
+             inst "a" "(number_of ?v)::rat" right_distrib,
+             divide_1,times_divide_eq_right,times_divide_eq_left,
+         rat_of_int_zero, rat_of_int_one, rat_of_int_add_distrib,
+         rat_of_int_minus_distrib, rat_of_int_diff_distrib,
+         rat_of_int_mult_distrib, number_of_rat RS sym];
+
+in
+
+val fast_rat_arith_simproc = Simplifier.simproc (Theory.sign_of (the_context ()))
+  "fast_rat_arith" ["(m::rat) < n","(m::rat) <= n", "(m::rat) = n"]
+  Fast_Arith.lin_arith_prover;
+
+val int_inj_thms = [rat_of_int_le_iff RS iffD2, rat_of_int_less_iff RS iffD2,
+                    rat_inject RS iffD2];
+
+val rat_arith_setup =
+ [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
+   {add_mono_thms = add_mono_thms @ add_mono_thms_rat,
+    mult_mono_thms = mult_mono_thms @ rat_mult_mono_thms,
+    inj_thms = (***int_inj_thms @*???**)  inj_thms,
+    lessD = lessD,  (*Can't change LA_Data_Ref.lessD: the rats are dense!*)
+    simpset = simpset addsimps add_rules
+                      addsimps simps
+                      addsimprocs simprocs}),
+(*???
+  arith_inj_const ("Rational.rat", HOLogic.intT --> Rat_Numeral_Simprocs.ratT),
+???*)
+  arith_discrete ("Rational.rat",false),
+  Simplifier.change_simpset_of (op addsimprocs) [fast_rat_arith_simproc]];
+
+
+end;
+
+
--- a/src/HOL/Real/real_arith.ML	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Real/real_arith.ML	Tue Jan 27 15:39:51 2004 +0100
@@ -8,6 +8,384 @@
 Instantiation of the generic linear arithmetic package for type real.
 *)
 
+(*FIXME DELETE*)
+val real_mult_left_mono =
+    read_instantiate_sg(sign_of (the_context())) [("a","?a::real")] mult_left_mono;
+
+val real_numeral_0_eq_0 = thm "real_numeral_0_eq_0";
+val real_numeral_1_eq_1 = thm "real_numeral_1_eq_1";
+val real_number_of_def = thm "real_number_of_def";
+val add_real_number_of = thm "add_real_number_of";
+val minus_real_number_of = thm "minus_real_number_of";
+val diff_real_number_of = thm "diff_real_number_of";
+val mult_real_number_of = thm "mult_real_number_of";
+val real_mult_2 = thm "real_mult_2";
+val real_mult_2_right = thm "real_mult_2_right";
+val eq_real_number_of = thm "eq_real_number_of";
+val less_real_number_of = thm "less_real_number_of";
+val real_minus_1_eq_m1 = thm "real_minus_1_eq_m1";
+val real_mult_minus1 = thm "real_mult_minus1";
+val real_mult_minus1_right = thm "real_mult_minus1_right";
+val zero_less_real_of_nat_iff = thm "zero_less_real_of_nat_iff";
+val zero_le_real_of_nat_iff = thm "zero_le_real_of_nat_iff";
+val real_add_number_of_left = thm "real_add_number_of_left";
+val real_mult_number_of_left = thm "real_mult_number_of_left";
+val real_add_number_of_diff1 = thm "real_add_number_of_diff1";
+val real_add_number_of_diff2 = thm "real_add_number_of_diff2";
+val real_of_nat_number_of = thm "real_of_nat_number_of";
+
+(*Maps 0 to Numeral0 and 1 to Numeral1 and -(Numeral1) to -1*)
+val real_numeral_ss =
+    HOL_ss addsimps [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym,
+                     real_minus_1_eq_m1];
+
+fun rename_numerals th =
+    asm_full_simplify real_numeral_ss (Thm.transfer (the_context ()) th);
+
+
+structure Real_Numeral_Simprocs =
+struct
+
+(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in simprocs
+  isn't complicated by the abstract 0 and 1.*)
+val numeral_syms = [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym];
+
+(*Utilities*)
+
+fun mk_numeral n = HOLogic.number_of_const HOLogic.realT $ HOLogic.mk_bin n;
+
+(*Decodes a binary real constant, or 0, 1*)
+val dest_numeral = Int_Numeral_Simprocs.dest_numeral;
+val find_first_numeral = Int_Numeral_Simprocs.find_first_numeral;
+
+val zero = mk_numeral 0;
+val mk_plus = HOLogic.mk_binop "op +";
+
+val uminus_const = Const ("uminus", HOLogic.realT --> HOLogic.realT);
+
+(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
+fun mk_sum []        = zero
+  | mk_sum [t,u]     = mk_plus (t, u)
+  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
+
+(*this version ALWAYS includes a trailing zero*)
+fun long_mk_sum []        = zero
+  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
+
+val dest_plus = HOLogic.dest_bin "op +" HOLogic.realT;
+
+(*decompose additions AND subtractions as a sum*)
+fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
+        dest_summing (pos, t, dest_summing (pos, u, ts))
+  | dest_summing (pos, Const ("op -", _) $ t $ u, ts) =
+        dest_summing (pos, t, dest_summing (not pos, u, ts))
+  | dest_summing (pos, t, ts) =
+        if pos then t::ts else uminus_const$t :: ts;
+
+fun dest_sum t = dest_summing (true, t, []);
+
+val mk_diff = HOLogic.mk_binop "op -";
+val dest_diff = HOLogic.dest_bin "op -" HOLogic.realT;
+
+val one = mk_numeral 1;
+val mk_times = HOLogic.mk_binop "op *";
+
+fun mk_prod [] = one
+  | mk_prod [t] = t
+  | mk_prod (t :: ts) = if t = one then mk_prod ts
+                        else mk_times (t, mk_prod ts);
+
+val dest_times = HOLogic.dest_bin "op *" HOLogic.realT;
+
+fun dest_prod t =
+      let val (t,u) = dest_times t
+      in  dest_prod t @ dest_prod u  end
+      handle TERM _ => [t];
+
+(*DON'T do the obvious simplifications; that would create special cases*)
+fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts);
+
+(*Express t as a product of (possibly) a numeral with other sorted terms*)
+fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t
+  | dest_coeff sign t =
+    let val ts = sort Term.term_ord (dest_prod t)
+        val (n, ts') = find_first_numeral [] ts
+                          handle TERM _ => (1, ts)
+    in (sign*n, mk_prod ts') end;
+
+(*Find first coefficient-term THAT MATCHES u*)
+fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
+  | find_first_coeff past u (t::terms) =
+        let val (n,u') = dest_coeff 1 t
+        in  if u aconv u' then (n, rev past @ terms)
+                          else find_first_coeff (t::past) u terms
+        end
+        handle TERM _ => find_first_coeff (t::past) u terms;
+
+
+(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
+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_minus1, real_mult_minus1_right];
+
+(*To perform binary arithmetic*)
+val bin_simps =
+    [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym,
+     add_real_number_of, real_add_number_of_left, minus_real_number_of,
+     diff_real_number_of, mult_real_number_of, real_mult_number_of_left] @
+    bin_arith_simps @ bin_rel_simps;
+
+(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
+  during re-arrangement*)
+val non_add_bin_simps = 
+    bin_simps \\ [real_add_number_of_left, add_real_number_of];
+
+(*To evaluate binary negations of coefficients*)
+val real_minus_simps = NCons_simps @
+                   [real_minus_1_eq_m1, minus_real_number_of,
+                    bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min,
+                    bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min];
+
+(*To let us treat subtraction as addition*)
+val diff_simps = [real_diff_def, minus_add_distrib, minus_minus];
+
+(*to extract again any uncancelled minuses*)
+val real_minus_from_mult_simps =
+    [minus_minus, minus_mult_left RS sym, minus_mult_right RS sym];
+
+(*combine unary minus with numeric literals, however nested within a product*)
+val real_mult_minus_simps =
+    [mult_assoc, minus_mult_left, minus_mult_commute];
+
+(*Apply the given rewrite (if present) just once*)
+fun trans_tac None      = all_tac
+  | trans_tac (Some th) = ALLGOALS (rtac (th RS trans));
+
+(*Final simplification: cancel + and *  *)
+val simplify_meta_eq =
+    Int_Numeral_Simprocs.simplify_meta_eq
+         [add_0, add_0_right,
+          mult_zero_left, mult_zero_right, mult_1, mult_1_right];
+
+fun prep_simproc (name, pats, proc) =
+  Simplifier.simproc (Theory.sign_of (the_context ())) name pats proc;
+
+structure CancelNumeralsCommon =
+  struct
+  val mk_sum            = mk_sum
+  val dest_sum          = dest_sum
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff 1
+  val find_first_coeff  = find_first_coeff []
+  val trans_tac         = trans_tac
+  val norm_tac =
+     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
+                                         real_minus_simps@add_ac))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@real_mult_minus_simps))
+     THEN ALLGOALS
+              (simp_tac (HOL_ss addsimps real_minus_from_mult_simps@
+                                         add_ac@mult_ac))
+  val numeral_simp_tac  = ALLGOALS (simp_tac (HOL_ss addsimps add_0s@bin_simps))
+  val simplify_meta_eq  = simplify_meta_eq
+  end;
+
+
+structure EqCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Bin_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" HOLogic.realT
+  val bal_add1 = eq_add_iff1 RS trans
+  val bal_add2 = eq_add_iff2 RS trans
+);
+
+structure LessCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Bin_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel "op <"
+  val dest_bal = HOLogic.dest_bin "op <" HOLogic.realT
+  val bal_add1 = less_add_iff1 RS trans
+  val bal_add2 = less_add_iff2 RS trans
+);
+
+structure LeCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Bin_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel "op <="
+  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.realT
+  val bal_add1 = le_add_iff1 RS trans
+  val bal_add2 = le_add_iff2 RS trans
+);
+
+val cancel_numerals =
+  map prep_simproc
+   [("realeq_cancel_numerals",
+     ["(l::real) + m = n", "(l::real) = m + n",
+      "(l::real) - m = n", "(l::real) = m - n",
+      "(l::real) * m = n", "(l::real) = m * n"],
+     EqCancelNumerals.proc),
+    ("realless_cancel_numerals",
+     ["(l::real) + m < n", "(l::real) < m + n",
+      "(l::real) - m < n", "(l::real) < m - n",
+      "(l::real) * m < n", "(l::real) < m * n"],
+     LessCancelNumerals.proc),
+    ("realle_cancel_numerals",
+     ["(l::real) + m <= n", "(l::real) <= m + n",
+      "(l::real) - m <= n", "(l::real) <= m - n",
+      "(l::real) * m <= n", "(l::real) <= m * n"],
+     LeCancelNumerals.proc)];
+
+
+structure CombineNumeralsData =
+  struct
+  val add               = op + : int*int -> int
+  val mk_sum            = long_mk_sum    (*to work for e.g. 2*x + 3*x *)
+  val dest_sum          = dest_sum
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff 1
+  val left_distrib      = combine_common_factor RS trans
+  val prove_conv        = Bin_Simprocs.prove_conv_nohyps
+  val trans_tac         = trans_tac
+  val norm_tac =
+     ALLGOALS (simp_tac (HOL_ss addsimps numeral_syms@add_0s@mult_1s@
+                                   diff_simps@real_minus_simps@add_ac))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@real_mult_minus_simps))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps real_minus_from_mult_simps@
+                                              add_ac@mult_ac))
+  val numeral_simp_tac  = ALLGOALS
+                    (simp_tac (HOL_ss addsimps add_0s@bin_simps))
+  val simplify_meta_eq  =
+        Int_Numeral_Simprocs.simplify_meta_eq (add_0s@mult_1s)
+  end;
+
+structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
+
+val combine_numerals =
+  prep_simproc ("real_combine_numerals", ["(i::real) + j", "(i::real) - j"], CombineNumerals.proc);
+
+
+(** Declarations for ExtractCommonTerm **)
+
+(*this version ALWAYS includes a trailing one*)
+fun long_mk_prod []        = one
+  | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts);
+
+(*Find first term that matches u*)
+fun find_first past u []         = raise TERM("find_first", [])
+  | find_first past u (t::terms) =
+        if u aconv t then (rev past @ terms)
+        else find_first (t::past) u terms
+        handle TERM _ => find_first (t::past) u terms;
+
+(*Final simplification: cancel + and *  *)
+fun cancel_simplify_meta_eq cancel_th th =
+    Int_Numeral_Simprocs.simplify_meta_eq
+        [real_mult_1, real_mult_1_right]
+        (([th, cancel_th]) MRS trans);
+
+(*** Making constant folding work for 0 and 1 too ***)
+
+structure RealAbstractNumeralsData =
+  struct
+  val dest_eq         = HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of
+  val is_numeral      = Bin_Simprocs.is_numeral
+  val numeral_0_eq_0  = real_numeral_0_eq_0
+  val numeral_1_eq_1  = real_numeral_1_eq_1
+  val prove_conv      = Bin_Simprocs.prove_conv_nohyps_novars
+  fun norm_tac simps  = ALLGOALS (simp_tac (HOL_ss addsimps simps))
+  val simplify_meta_eq = Bin_Simprocs.simplify_meta_eq
+  end
+
+structure RealAbstractNumerals = AbstractNumeralsFun (RealAbstractNumeralsData)
+
+(*For addition, we already have rules for the operand 0.
+  Multiplication is omitted because there are already special rules for
+  both 0 and 1 as operands.  Unary minus is trivial, just have - 1 = -1.
+  For the others, having three patterns is a compromise between just having
+  one (many spurious calls) and having nine (just too many!) *)
+val eval_numerals =
+  map prep_simproc
+   [("real_add_eval_numerals",
+     ["(m::real) + 1", "(m::real) + number_of v"],
+     RealAbstractNumerals.proc add_real_number_of),
+    ("real_diff_eval_numerals",
+     ["(m::real) - 1", "(m::real) - number_of v"],
+     RealAbstractNumerals.proc diff_real_number_of),
+    ("real_eq_eval_numerals",
+     ["(m::real) = 0", "(m::real) = 1", "(m::real) = number_of v"],
+     RealAbstractNumerals.proc eq_real_number_of),
+    ("real_less_eval_numerals",
+     ["(m::real) < 0", "(m::real) < 1", "(m::real) < number_of v"],
+     RealAbstractNumerals.proc less_real_number_of),
+    ("real_le_eval_numerals",
+     ["(m::real) <= 0", "(m::real) <= 1", "(m::real) <= number_of v"],
+     RealAbstractNumerals.proc le_number_of_eq_not_less)]
+
+end;
+
+
+Addsimprocs Real_Numeral_Simprocs.eval_numerals;
+Addsimprocs Real_Numeral_Simprocs.cancel_numerals;
+Addsimprocs [Real_Numeral_Simprocs.combine_numerals];
+
+(*examples:
+print_depth 22;
+set timing;
+set trace_simp;
+fun test s = (Goal s; by (Simp_tac 1));
+
+test "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::real)";
+
+test "2*u = (u::real)";
+test "(i + j + 12 + (k::real)) - 15 = y";
+test "(i + j + 12 + (k::real)) - 5 = y";
+
+test "y - b < (b::real)";
+test "y - (3*b + c) < (b::real) - 2*c";
+
+test "(2*x - (u*v) + y) - v*3*u = (w::real)";
+test "(2*x*u*v + (u*v)*4 + y) - v*u*4 = (w::real)";
+test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::real)";
+test "u*v - (x*u*v + (u*v)*4 + y) = (w::real)";
+
+test "(i + j + 12 + (k::real)) = u + 15 + y";
+test "(i + j*2 + 12 + (k::real)) = j + 5 + y";
+
+test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::real)";
+
+test "a + -(b+c) + b = (d::real)";
+test "a + -(b+c) - b = (d::real)";
+
+(*negative numerals*)
+test "(i + j + -2 + (k::real)) - (u + 5 + y) = zz";
+test "(i + j + -3 + (k::real)) < u + 5 + y";
+test "(i + j + 3 + (k::real)) < u + -6 + y";
+test "(i + j + -12 + (k::real)) - 15 = y";
+test "(i + j + 12 + (k::real)) - -15 = y";
+test "(i + j + -12 + (k::real)) - -15 = y";
+*)
+
+
+(** Constant folding for real plus and times **)
+
+(*We do not need
+    structure Real_Plus_Assoc = Assoc_Fold (Real_Plus_Assoc_Data);
+  because combine_numerals does the same thing*)
+
+structure Real_Times_Assoc_Data : ASSOC_FOLD_DATA =
+struct
+  val ss                = HOL_ss
+  val eq_reflection     = eq_reflection
+  val sg_ref    = Sign.self_ref (Theory.sign_of (the_context ()))
+  val T      = HOLogic.realT
+  val plus   = Const ("op *", [HOLogic.realT,HOLogic.realT] ---> HOLogic.realT)
+  val add_ac = mult_ac
+end;
+
+structure Real_Times_Assoc = Assoc_Fold (Real_Times_Assoc_Data);
+
+Addsimprocs [Real_Times_Assoc.conv];
+
 
 (****Common factor cancellation****)
 
@@ -26,7 +404,7 @@
 in
 
 val rel_real_number_of = [eq_real_number_of, less_real_number_of,
-                          le_real_number_of_eq_not_less]
+                          le_number_of_eq_not_less]
 
 structure CancelNumeralFactorCommon =
   struct
@@ -216,10 +594,6 @@
 val add_zero_left = thm"Ring_and_Field.add_0";
 val add_zero_right = thm"Ring_and_Field.add_0_right";
 
-val real_mult_left_mono =
-    read_instantiate_sg(sign_of RealBin.thy) [("a","?a::real")] mult_left_mono;
-
-
 local
 
 (* reduce contradictory <= to False *)
@@ -228,7 +602,7 @@
      real_minus_1_eq_m1, 
      add_real_number_of, minus_real_number_of, diff_real_number_of,
      mult_real_number_of, eq_real_number_of, less_real_number_of,
-     le_real_number_of_eq_not_less, real_diff_def,
+     le_number_of_eq_not_less, diff_minus,
      minus_add_distrib, minus_minus, mult_assoc, minus_zero,
      add_zero_left, add_zero_right, left_minus, right_minus,
      mult_zero_left, mult_zero_right, mult_1, mult_1_right,
@@ -240,8 +614,7 @@
                Real_Numeral_Simprocs.eval_numerals;
 
 val mono_ss = simpset() addsimps
-                [add_mono,add_strict_mono,
-                 real_add_less_le_mono,real_add_le_less_mono];
+                [add_mono,add_strict_mono,add_less_le_mono,add_le_less_mono];
 
 val add_mono_thms_real =
   map (fn s => prove_goal (the_context ()) s
@@ -271,7 +644,7 @@
          real_of_nat_zero, real_of_nat_Suc, real_of_nat_add, real_of_nat_mult,
          real_of_int_zero, real_of_one, real_of_int_add RS sym,
          real_of_int_minus RS sym, real_of_int_diff RS sym,
-         real_of_int_mult RS sym, real_number_of];
+         real_of_int_mult RS sym, symmetric real_number_of_def];
 
 val int_inj_thms = [real_of_int_le_iff RS iffD2, real_of_int_less_iff RS iffD2,
                     real_of_int_inject RS iffD2];
@@ -290,7 +663,7 @@
    {add_mono_thms = add_mono_thms @ add_mono_thms_real,
     mult_mono_thms = mult_mono_thms @ real_mult_mono_thms,
     inj_thms = int_inj_thms @ nat_inj_thms @ inj_thms,
-    lessD = lessD,  (*Can't change lessD: the reals are dense!*)
+    lessD = lessD,  (*Can't change LA_Data_Ref.lessD: the reals are dense!*)
     simpset = simpset addsimps add_rules
                       addsimps simps
                       addsimprocs simprocs}),
--- a/src/HOL/Ring_and_Field.thy	Tue Jan 27 09:44:14 2004 +0100
+++ b/src/HOL/Ring_and_Field.thy	Tue Jan 27 15:39:51 2004 +0100
@@ -210,6 +210,9 @@
 lemma minus_mult_minus [simp]: "(- a) * (- b) = a * (b::'a::ring)"
   by (simp add: minus_mult_left [symmetric] minus_mult_right [symmetric])
 
+lemma minus_mult_commute: "(- a) * b = a * (- b::'a::ring)"
+  by (simp add: minus_mult_left [symmetric] minus_mult_right [symmetric])
+
 lemma right_diff_distrib: "a * (b - c) = a * b - a * (c::'a::ring)"
 by (simp add: right_distrib diff_minus 
               minus_mult_left [symmetric] minus_mult_right [symmetric]) 
@@ -889,6 +892,10 @@
 apply (simp add: mult_assoc [symmetric] add_commute)
 done
 
+lemma inverse_divide [simp]:
+      "inverse (a/b) = b / (a::'a::{field,division_by_zero})"
+  by (simp add: divide_inverse_zero mult_commute)
+
 lemma nonzero_mult_divide_cancel_left:
   assumes [simp]: "b\<noteq>0" and [simp]: "c\<noteq>0" 
     shows "(c*a)/(c*b) = a/(b::'a::field)"
@@ -1139,6 +1146,43 @@
 by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) 
 
 
+subsection{*Inverses and the Number One*}
+
+lemma one_less_inverse_iff:
+    "(1 < inverse x) = (0 < x & x < (1::'a::{ordered_field,division_by_zero}))"proof cases
+  assume "0 < x"
+    with inverse_less_iff_less [OF zero_less_one, of x]
+    show ?thesis by simp
+next
+  assume notless: "~ (0 < x)"
+  have "~ (1 < inverse x)"
+  proof
+    assume "1 < inverse x"
+    also with notless have "... \<le> 0" by (simp add: linorder_not_less)
+    also have "... < 1" by (rule zero_less_one) 
+    finally show False by auto
+  qed
+  with notless show ?thesis by simp
+qed
+
+lemma inverse_eq_1_iff [simp]:
+    "(inverse x = 1) = (x = (1::'a::{field,division_by_zero}))"
+by (insert inverse_eq_iff_eq [of x 1], simp) 
+
+lemma one_le_inverse_iff:
+   "(1 \<le> inverse x) = (0 < x & x \<le> (1::'a::{ordered_field,division_by_zero}))"
+by (force simp add: order_le_less one_less_inverse_iff zero_less_one 
+                    eq_commute [of 1]) 
+
+lemma inverse_less_1_iff:
+   "(inverse x < 1) = (x \<le> 0 | 1 < (x::'a::{ordered_field,division_by_zero}))"
+by (simp add: linorder_not_le [symmetric] one_le_inverse_iff) 
+
+lemma inverse_le_1_iff:
+   "(inverse x \<le> 1) = (x \<le> 0 | 1 \<le> (x::'a::{ordered_field,division_by_zero}))"
+by (simp add: linorder_not_less [symmetric] one_less_inverse_iff) 
+
+
 subsection{*Division and Signs*}
 
 lemma zero_less_divide_iff:
@@ -1412,13 +1456,16 @@
 
 subsection {* Ordered Fields are Dense *}
 
-lemma zero_less_two: "0 < (1+1::'a::ordered_field)"
+lemma less_add_one: "a < (a+1::'a::ordered_semiring)"
 proof -
-  have "0 + 0 <  (1+1::'a::ordered_field)"
-    by (blast intro: zero_less_one add_strict_mono) 
+  have "a+0 < (a+1::'a::ordered_semiring)"
+    by (blast intro: zero_less_one add_strict_left_mono) 
   thus ?thesis by simp
 qed
 
+lemma zero_less_two: "0 < (1+1::'a::ordered_semiring)"
+  by (blast intro: order_less_trans zero_less_one less_add_one) 
+
 lemma less_half_sum: "a < b ==> a < (a+b) / (1+1::'a::ordered_field)"
 by (simp add: zero_less_two pos_less_divide_eq right_distrib) 
 
@@ -1610,6 +1657,7 @@
 val minus_mult_left = thm"minus_mult_left";
 val minus_mult_right = thm"minus_mult_right";
 val minus_mult_minus = thm"minus_mult_minus";
+val minus_mult_commute = thm"minus_mult_commute";
 val right_diff_distrib = thm"right_diff_distrib";
 val left_diff_distrib = thm"left_diff_distrib";
 val minus_diff_eq = thm"minus_diff_eq";
@@ -1700,6 +1748,7 @@
 val nonzero_inverse_eq_divide = thm"nonzero_inverse_eq_divide";
 val divide_self = thm"divide_self";
 val divide_inverse_zero = thm"divide_inverse_zero";
+val inverse_divide = thm"inverse_divide";
 val divide_zero_left = thm"divide_zero_left";
 val inverse_eq_divide = thm"inverse_eq_divide";
 val nonzero_add_divide_distrib = thm"nonzero_add_divide_distrib";