Polymorphic treatment of binary arithmetic using axclasses
authorpaulson
Sun, 15 Feb 2004 10:46:37 +0100
changeset 14387 e96d5c42c4b0
parent 14386 ad1ffcc90162
child 14388 04f04408b99b
Polymorphic treatment of binary arithmetic using axclasses
doc-src/TutorialI/Types/document/Numbers.tex
doc-src/TutorialI/Types/document/Pairs.tex
src/HOL/Complex/CLim.ML
src/HOL/Complex/Complex.thy
src/HOL/Complex/ComplexArith0.ML
src/HOL/Complex/ComplexArith0.thy
src/HOL/Complex/ComplexBin.ML
src/HOL/Complex/ComplexBin.thy
src/HOL/Complex/NSCA.ML
src/HOL/Complex/NSCA.thy
src/HOL/Complex/NSComplex.thy
src/HOL/Complex/NSComplexArith.thy
src/HOL/Complex/NSComplexBin.ML
src/HOL/Complex/NSComplexBin.thy
src/HOL/Complex/NSInduct.thy
src/HOL/Complex/hcomplex_arith.ML
src/HOL/Hyperreal/HSeries.ML
src/HOL/Hyperreal/HyperArith.thy
src/HOL/Hyperreal/HyperDef.thy
src/HOL/Hyperreal/HyperPow.thy
src/HOL/Hyperreal/IntFloor.ML
src/HOL/Hyperreal/Integration.ML
src/HOL/Hyperreal/Lim.ML
src/HOL/Hyperreal/Lim.thy
src/HOL/Hyperreal/NSA.thy
src/HOL/Hyperreal/SEQ.ML
src/HOL/Hyperreal/Transcendental.ML
src/HOL/Hyperreal/hypreal_arith.ML
src/HOL/Integ/Bin.thy
src/HOL/Integ/IntArith.thy
src/HOL/Integ/IntDef.thy
src/HOL/Integ/IntDiv.thy
src/HOL/Integ/IntDiv_setup.ML
src/HOL/Integ/NatBin.thy
src/HOL/Integ/NatSimprocs.thy
src/HOL/Integ/int_arith1.ML
src/HOL/Integ/int_factor_simprocs.ML
src/HOL/Integ/nat_simprocs.ML
src/HOL/IsaMakefile
src/HOL/NumberTheory/Int2.thy
src/HOL/NumberTheory/IntPrimes.thy
src/HOL/NumberTheory/Quadratic_Reciprocity.thy
src/HOL/Numeral.thy
src/HOL/ROOT.ML
src/HOL/Real/PReal.thy
src/HOL/Real/RComplete.thy
src/HOL/Real/RatArith.thy
src/HOL/Real/Rational.thy
src/HOL/Real/RealArith.thy
src/HOL/Real/RealDef.thy
src/HOL/Real/RealPow.thy
src/HOL/Real/rat_arith.ML
src/HOL/Real/real_arith.ML
src/HOL/Ring_and_Field.thy
src/HOL/hologic.ML
src/Provers/Arith/assoc_fold.ML
src/Provers/Arith/cancel_numerals.ML
src/Provers/Arith/combine_numerals.ML
src/Provers/Arith/extract_common_term.ML
src/Pure/drule.ML
src/ZF/Integ/int_arith.ML
src/ZF/arith_data.ML
--- a/doc-src/TutorialI/Types/document/Numbers.tex	Sat Feb 14 02:06:12 2004 +0100
+++ b/doc-src/TutorialI/Types/document/Numbers.tex	Sun Feb 15 10:46:37 2004 +0100
@@ -37,12 +37,12 @@
 %
 \begin{isamarkuptext}%
 \begin{isabelle}%
-Numeral{\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}%
+Numeral{\isadigit{0}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}{\isacharprime}a{\isacharparenright}%
 \end{isabelle}
 \rulename{numeral_0_eq_0}
 
 \begin{isabelle}%
-Numeral{\isadigit{1}}\ {\isacharequal}\ {\isadigit{1}}%
+Numeral{\isadigit{1}}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{1}}{\isasymColon}{\isacharprime}a{\isacharparenright}%
 \end{isabelle}
 \rulename{numeral_1_eq_1}
 
--- a/doc-src/TutorialI/Types/document/Pairs.tex	Sat Feb 14 02:06:12 2004 +0100
+++ b/doc-src/TutorialI/Types/document/Pairs.tex	Sun Feb 15 10:46:37 2004 +0100
@@ -33,7 +33,7 @@
 \isa{case\ xs\ of\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ {\isadigit{0}}\ {\isacharbar}\ {\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isacharhash}\ zs\ {\isasymRightarrow}\ x\ {\isacharplus}\ y}\\
 \isa{{\isasymforall}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}{\isasymin}A{\isachardot}\ x{\isacharequal}y}\\
 \isa{{\isacharbraceleft}{\isacharparenleft}x{\isacharcomma}y{\isacharcomma}z{\isacharparenright}{\isachardot}\ x{\isacharequal}z{\isacharbraceright}}\\
-\isa{{\isasymUnion}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isasymin}A{\isachardot}\ {\isacharbraceleft}x\ {\isacharplus}\ y{\isacharbraceright}}
+\isa{{\isasymUnion}\isactrlbsub {\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isasymin}A\isactrlesub \ {\isacharbraceleft}x\ {\isacharplus}\ y{\isacharbraceright}}
 \end{quote}
 The intuitive meanings of these expressions should be obvious.
 Unfortunately, we need to know in more detail what the notation really stands
--- a/src/HOL/Complex/CLim.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Complex/CLim.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -5,6 +5,20 @@
                   differentiation for complex functions
 *)
 
+(*FIXME: MOVE these two to Complex.thy*)
+Goal "(x + - a = (0::complex)) = (x=a)";
+by (simp_tac (simpset() addsimps [diff_eq_eq,symmetric complex_diff_def]) 1);
+qed "complex_add_minus_iff";
+Addsimps [complex_add_minus_iff];
+
+Goal "(x+y = (0::complex)) = (y = -x)";
+by Auto_tac;
+by (dtac (sym RS (diff_eq_eq RS iffD2)) 1);
+by Auto_tac;  
+qed "complex_add_eq_0_iff";
+AddIffs [complex_add_eq_0_iff];
+
+
 (*-----------------------------------------------------------------------*)
 (* Limit of complex to complex function                                               *)
 (*-----------------------------------------------------------------------*)
@@ -1175,7 +1189,7 @@
 Goal "(ALL z. f z - f x = g z * (z - x)) & isNSContc g x & g x = l \
 \     ==> NSCDERIV f x :> l";
 by (auto_tac (claset(), 
-              simpset() delsimprocs complex_cancel_factor
+              simpset() delsimprocs field_cancel_factor
                         addsimps [NSCDERIV_iff2]));
 by (asm_full_simp_tac (simpset() addsimps [isNSContc_def]) 1);
 qed "CARAT_CDERIVD";
--- a/src/HOL/Complex/Complex.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Complex/Complex.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -1,6 +1,7 @@
 (*  Title:       Complex.thy
     Author:      Jacques D. Fleuriot
     Copyright:   2001 University of Edinburgh
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
 *)
 
 header {* Complex Numbers: Rectangular and Polar Representations *}
@@ -830,6 +831,143 @@
 done
 
 
+subsection{*Numerals and Arithmetic*}
+
+instance complex :: number ..
+
+primrec (*the type constraint is essential!*)
+  number_of_Pls: "number_of bin.Pls = 0"
+  number_of_Min: "number_of bin.Min = - (1::complex)"
+  number_of_BIT: "number_of(w BIT x) = (if x then 1 else 0) +
+	                               (number_of w) + (number_of w)"
+
+declare number_of_Pls [simp del]
+        number_of_Min [simp del]
+        number_of_BIT [simp del]
+
+instance complex :: number_ring
+proof
+  show "Numeral0 = (0::complex)" by (rule number_of_Pls)
+  show "-1 = - (1::complex)" by (rule number_of_Min)
+  fix w :: bin and x :: bool
+  show "(number_of (w BIT x) :: complex) =
+        (if x then 1 else 0) + number_of w + number_of w"
+    by (rule number_of_BIT)
+qed
+
+
+text{*Collapse applications of @{term complex_of_real} to @{term number_of}*}
+lemma complex_number_of [simp]: "complex_of_real (number_of w) = number_of w"
+apply (induct w) 
+apply (simp_all only: number_of complex_of_real_add [symmetric] 
+                      complex_of_real_minus, simp_all) 
+done
+
+text{*This theorem is necessary because theorems such as
+   @{text iszero_number_of_0} only hold for ordered rings. They cannot
+   be generalized to fields in general because they fail for finite fields.
+   They work for type complex because the reals can be embedded in them.*}
+lemma iszero_complex_number_of [simp]:
+     "iszero (number_of w :: complex) = iszero (number_of w :: real)"
+by (simp only: complex_of_real_zero_iff complex_number_of [symmetric] 
+               iszero_def)  
+
+
+(*These allow simplification of expressions involving mixed numbers.
+  Convert???
+Goalw [complex_number_of_def] 
+  "((number_of xa :: complex) + ii * number_of ya =  
+        number_of xb) =  
+   (((number_of xa :: complex) = number_of xb) &  
+    ((number_of ya :: complex) = 0))"
+by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff2,
+    complex_of_real_zero_iff]));
+qed "complex_number_of_eq_cancel_iff2";
+Addsimps [complex_number_of_eq_cancel_iff2];
+
+Goalw [complex_number_of_def] 
+  "((number_of xa :: complex) + number_of ya * ii = \
+\       number_of xb) = \
+\  (((number_of xa :: complex) = number_of xb) & \
+\   ((number_of ya :: complex) = 0))";
+by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff2a,
+    complex_of_real_zero_iff]));
+qed "complex_number_of_eq_cancel_iff2a";
+Addsimps [complex_number_of_eq_cancel_iff2a];
+
+Goalw [complex_number_of_def] 
+  "((number_of xa :: complex) + ii * number_of ya = \
+\    ii * number_of yb) = \
+\  (((number_of xa :: complex) = 0) & \
+\   ((number_of ya :: complex) = number_of yb))";
+by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff3,
+    complex_of_real_zero_iff]));
+qed "complex_number_of_eq_cancel_iff3";
+Addsimps [complex_number_of_eq_cancel_iff3];
+
+Goalw [complex_number_of_def] 
+  "((number_of xa :: complex) + number_of ya * ii= \
+\    ii * number_of yb) = \
+\  (((number_of xa :: complex) = 0) & \
+\   ((number_of ya :: complex) = number_of yb))";
+by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff3a,
+    complex_of_real_zero_iff]));
+qed "complex_number_of_eq_cancel_iff3a";
+Addsimps [complex_number_of_eq_cancel_iff3a];
+*)
+
+lemma complex_number_of_cnj [simp]: "cnj(number_of v :: complex) = number_of v"
+apply (subst complex_number_of [symmetric])
+apply (rule complex_cnj_complex_of_real)
+done
+
+lemma complex_number_of_cmod: 
+      "cmod(number_of v :: complex) = abs (number_of v :: real)"
+by (simp only: complex_number_of [symmetric] complex_mod_complex_of_real)
+
+lemma complex_number_of_Re [simp]: "Re(number_of v :: complex) = number_of v"
+by (simp only: complex_number_of [symmetric] Re_complex_of_real)
+
+lemma complex_number_of_Im [simp]: "Im(number_of v :: complex) = 0"
+by (simp only: complex_number_of [symmetric] Im_complex_of_real)
+
+lemma expi_two_pi_i [simp]: "expi((2::complex) * complex_of_real pi * ii) = 1"
+by (simp add: expi_def complex_Re_mult_eq complex_Im_mult_eq cis_def)
+
+
+(*examples:
+print_depth 22
+set timing;
+set trace_simp;
+fun test s = (Goal s, by (Simp_tac 1)); 
+
+test "23 * ii + 45 * ii= (x::complex)";
+
+test "5 * ii + 12 - 45 * ii= (x::complex)";
+test "5 * ii + 40 - 12 * ii + 9 = (x::complex) + 89 * ii";
+test "5 * ii + 40 - 12 * ii + 9 - 78 = (x::complex) + 89 * ii";
+
+test "l + 10 * ii + 90 + 3*l +  9 + 45 * ii= (x::complex)";
+test "87 + 10 * ii + 90 + 3*7 +  9 + 45 * ii= (x::complex)";
+
+
+fun test s = (Goal s; by (Asm_simp_tac 1)); 
+
+test "x*k = k*(y::complex)";
+test "k = k*(y::complex)"; 
+test "a*(b*c) = (b::complex)";
+test "a*(b*c) = d*(b::complex)*(x*a)";
+
+
+test "(x*k) / (k*(y::complex)) = (uu::complex)";
+test "(k) / (k*(y::complex)) = (uu::complex)"; 
+test "(a*(b*c)) / ((b::complex)) = (uu::complex)";
+test "(a*(b*c)) / (d*(b::complex)*(x*a)) = (uu::complex)";
+
+(*FIXME: what do we do about this?*)
+test "a*(b*c)/(y*z) = d*(b::complex)*(x*a)/z";
+*)
+
 
 ML
 {*
--- a/src/HOL/Complex/ComplexArith0.ML	Sat Feb 14 02:06:12 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,187 +0,0 @@
-(*  Title:       ComplexArith0.ML
-    Author:      Jacques D. Fleuriot
-    Copyright:   2001  University of Edinburgh
-    Description: Assorted facts that need binary literals 
-		 Also, common factor cancellation (see e.g. HyperArith0)
-*)
-
-local
-  open Complex_Numeral_Simprocs
-in
-
-val rel_complex_number_of = [eq_complex_number_of];
-
-
-structure CancelNumeralFactorCommon =
-  struct
-  val mk_coeff		= mk_coeff
-  val dest_coeff	= dest_coeff 1
-  val trans_tac         = Real_Numeral_Simprocs.trans_tac
-  val norm_tac =  ALLGOALS (simp_tac (HOL_ss addsimps complex_minus_from_mult_simps @ mult_1s)) 
-                  THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@complex_mult_minus_simps))
-                  THEN ALLGOALS (simp_tac (HOL_ss addsimps mult_ac))
-  val numeral_simp_tac	=  ALLGOALS (simp_tac (HOL_ss addsimps rel_complex_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" complexT
-  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 =" complexT
-  val cancel = field_mult_cancel_left RS trans
-  val neg_exchanges = false
-)
-
-val complex_cancel_numeral_factors_relations = 
-  map prep_simproc
-   [("complexeq_cancel_numeral_factor",
-     ["(l::complex) * m = n", "(l::complex) = m * n"], 
-     EqCancelNumeralFactor.proc)];
-
-val complex_cancel_numeral_factors_divide = prep_simproc
-	("complexdiv_cancel_numeral_factor", 
-	 ["((l::complex) * m) / n", "(l::complex) / (m * n)", 
-                     "((number_of v)::complex) / (number_of w)"], 
-	 DivCancelNumeralFactor.proc);
-
-val complex_cancel_numeral_factors = 
-    complex_cancel_numeral_factors_relations @ 
-    [complex_cancel_numeral_factors_divide];
-
-end;
-
-
-Addsimprocs complex_cancel_numeral_factors;
-
-
-(*examples:
-print_depth 22;
-set timing;
-set trace_simp;
-fun test s = (Goal s; by (Simp_tac 1)); 
-
-
-test "9*x = 12 * (y::complex)";
-test "(9*x) / (12 * (y::complex)) = z";
-
-test "-99*x = 132 * (y::complex)";
-
-test "999*x = -396 * (y::complex)";
-test "(999*x) / (-396 * (y::complex)) = z";
-
-test "-99*x = -81 * (y::complex)";
-test "(-99*x) / (-81 * (y::complex)) = z";
-
-test "-2 * x = -1 * (y::complex)";
-test "-2 * x = -(y::complex)";
-test "(-2 * x) / (-1 * (y::complex)) = z";
-
-*)
-
-
-(** Declarations for ExtractCommonTerm **)
-
-local
-  open Complex_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         = Real_Numeral_Simprocs.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 =" complexT
-  val simplify_meta_eq  = cancel_simplify_meta_eq field_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" complexT
-  val simplify_meta_eq  = cancel_simplify_meta_eq mult_divide_cancel_eq_if
-);
-
-val complex_cancel_factor = 
-  map prep_simproc
-   [("complex_eq_cancel_factor", ["(l::complex) * m = n", "(l::complex) = m * n"], 
-     EqCancelFactor.proc),
-    ("complex_divide_cancel_factor", ["((l::complex) * m) / n", "(l::complex) / (m * n)"], 
-     DivideCancelFactor.proc)];
-
-end;
-
-Addsimprocs complex_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::complex)";
-test "k = k*(y::complex)"; 
-test "a*(b*c) = (b::complex)";
-test "a*(b*c) = d*(b::complex)*(x*a)";
-
-
-test "(x*k) / (k*(y::complex)) = (uu::complex)";
-test "(k) / (k*(y::complex)) = (uu::complex)"; 
-test "(a*(b*c)) / ((b::complex)) = (uu::complex)";
-test "(a*(b*c)) / (d*(b::complex)*(x*a)) = (uu::complex)";
-
-(*FIXME: what do we do about this?*)
-test "a*(b*c)/(y*z) = d*(b::complex)*(x*a)/z";
-*)
-
-
-(** Division by 1, -1 **)
-
-Goal "x/-1 = -(x::complex)";
-by (Simp_tac 1); 
-qed "complex_divide_minus1";
-Addsimps [complex_divide_minus1];
-
-Goal "-1/(x::complex) = - (1/x)";
-by (simp_tac (simpset() addsimps [complex_divide_def, inverse_minus_eq]) 1); 
-qed "complex_minus1_divide";
-Addsimps [complex_minus1_divide];
-
-Goal "(x + - a = (0::complex)) = (x=a)";
-by (simp_tac (simpset() addsimps [diff_eq_eq,symmetric complex_diff_def]) 1);
-qed "complex_add_minus_iff";
-Addsimps [complex_add_minus_iff];
-
-Goal "(x+y = (0::complex)) = (y = -x)";
-by Auto_tac;
-by (dtac (sym RS (diff_eq_eq RS iffD2)) 1);
-by Auto_tac;  
-qed "complex_add_eq_0_iff";
-AddIffs [complex_add_eq_0_iff];
-
-
--- a/src/HOL/Complex/ComplexArith0.thy	Sat Feb 14 02:06:12 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2 +0,0 @@
-ComplexArith0 = ComplexBin
-
--- a/src/HOL/Complex/ComplexBin.ML	Sat Feb 14 02:06:12 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,538 +0,0 @@
-(*  Title:      ComplexBin.ML
-    Author:     Jacques D. Fleuriot
-    Copyright:  2001 University of Edinburgh
-    Descrition: Binary arithmetic for the complex numbers
-*)
-
-(** complex_of_real (coercion from real to complex) **)
-
-Goal "complex_of_real (number_of w) = number_of w";
-by (simp_tac (simpset() addsimps [complex_number_of_def]) 1);
-qed "complex_number_of";
-Addsimps [complex_number_of];
- 
-Goalw [complex_number_of_def] "Numeral0 = (0::complex)";
-by (Simp_tac 1);
-qed "complex_numeral_0_eq_0";
- 
-Goalw [complex_number_of_def] "Numeral1 = (1::complex)";
-by (Simp_tac 1);
-qed "complex_numeral_1_eq_1";
-
-(** Addition **)
-
-Goal "(number_of v :: complex) + number_of v' = number_of (bin_add v v')";
-by (simp_tac
-    (HOL_ss addsimps [complex_number_of_def, 
-                      complex_of_real_add, add_real_number_of]) 1);
-qed "add_complex_number_of";
-Addsimps [add_complex_number_of];
-
-
-(** Subtraction **)
-
-Goalw [complex_number_of_def]
-     "- (number_of w :: complex) = number_of (bin_minus w)";
-by (simp_tac
-    (HOL_ss addsimps [minus_real_number_of, complex_of_real_minus RS sym]) 1);
-qed "minus_complex_number_of";
-Addsimps [minus_complex_number_of];
-
-Goalw [complex_number_of_def, complex_diff_def]
-     "(number_of v :: complex) - number_of w = number_of (bin_add v (bin_minus w))";
-by (Simp_tac 1); 
-qed "diff_complex_number_of";
-Addsimps [diff_complex_number_of];
-
-
-(** Multiplication **)
-
-Goal "(number_of v :: complex) * number_of v' = number_of (bin_mult v v')";
-by (simp_tac
-    (HOL_ss addsimps [complex_number_of_def, 
-	              complex_of_real_mult, mult_real_number_of]) 1);
-qed "mult_complex_number_of";
-Addsimps [mult_complex_number_of];
-
-Goal "(2::complex) = 1 + 1";
-by (simp_tac (simpset() addsimps [complex_numeral_1_eq_1 RS sym]) 1);
-val lemma = result();
-
-(*For specialist use: NOT as default simprules*)
-Goal "2 * z = (z+z::complex)";
-by (simp_tac (simpset () addsimps [lemma, left_distrib]) 1);
-qed "complex_mult_2";
-
-Goal "z * 2 = (z+z::complex)";
-by (stac mult_commute 1 THEN rtac complex_mult_2 1);
-qed "complex_mult_2_right";
-
-(** Equals (=) **)
-
-Goal "((number_of v :: complex) = number_of v') = \
-\     iszero (number_of (bin_add v (bin_minus v')) :: int)";
-by (simp_tac
-    (HOL_ss addsimps [complex_number_of_def, 
-	              complex_of_real_eq_iff, eq_real_number_of]) 1);
-qed "eq_complex_number_of";
-Addsimps [eq_complex_number_of];
-
-(*** New versions of existing theorems involving 0, 1 ***)
-
-Goal "- 1 = (-1::complex)";
-by (simp_tac (simpset() addsimps [complex_numeral_1_eq_1 RS sym]) 1);
-qed "complex_minus_1_eq_m1";
-
-Goal "-1 * z = -(z::complex)";
-by (simp_tac (simpset() addsimps [complex_minus_1_eq_m1 RS sym]) 1);
-qed "complex_mult_minus1";
-
-Goal "z * -1 = -(z::complex)";
-by (stac mult_commute 1 THEN rtac complex_mult_minus1 1);
-qed "complex_mult_minus1_right";
-
-Addsimps [complex_mult_minus1,complex_mult_minus1_right];
-
-
-(*Maps 0 to Numeral0 and 1 to Numeral1 and -Numeral1 to -1*)
-val complex_numeral_ss = 
-    hypreal_numeral_ss addsimps [complex_numeral_0_eq_0 RS sym, complex_numeral_1_eq_1 RS sym, 
-		                 complex_minus_1_eq_m1];
-
-fun rename_numerals th = 
-    asm_full_simplify complex_numeral_ss (Thm.transfer (the_context ()) th);
-
-(*Now insert some identities previously stated for 0 and 1c*)
-
-Addsimps [complex_numeral_0_eq_0,complex_numeral_1_eq_1];
-
-Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::complex)";
-by (auto_tac (claset(),simpset() addsimps [complex_add_assoc RS sym]));
-qed "complex_add_number_of_left";
-
-Goal "number_of v *(number_of w * z) = (number_of(bin_mult v w) * z::complex)";
-by (simp_tac (simpset() addsimps [mult_assoc RS sym]) 1);
-qed "complex_mult_number_of_left";
-
-Goalw [complex_diff_def]
-    "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::complex)";
-by (rtac complex_add_number_of_left 1);
-qed "complex_add_number_of_diff1";
-
-Goal "number_of v + (c - number_of w) = \
-\     number_of (bin_add v (bin_minus w)) + (c::complex)";
-by (auto_tac (claset(),simpset() addsimps [complex_diff_def]@ add_ac));
-qed "complex_add_number_of_diff2";
-
-Addsimps [complex_add_number_of_left, complex_mult_number_of_left,
-	  complex_add_number_of_diff1, complex_add_number_of_diff2]; 
-
-
-(**** Simprocs for numeric literals ****)
-
-(** Combining of literal coefficients in sums of products **)
-
-Goal "(x = y) = (x-y = (0::complex))";
-by (simp_tac (simpset() addsimps [diff_eq_eq]) 1);   
-qed "complex_eq_iff_diff_eq_0";
-
-
-
-structure Complex_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 = [complex_numeral_0_eq_0 RS sym, complex_numeral_1_eq_1 RS sym];
-
-
-(*Utilities*)
-
-val complexT = Type("Complex.complex",[]);
-
-fun mk_numeral n = HOLogic.number_of_const complexT $ HOLogic.mk_bin n;
-
-val dest_numeral = Real_Numeral_Simprocs.dest_numeral;
-val find_first_numeral = Real_Numeral_Simprocs.find_first_numeral;
-
-val zero = mk_numeral 0;
-val mk_plus = HOLogic.mk_binop "op +";
-
-val uminus_const = Const ("uminus", complexT --> complexT);
-
-(*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 +" complexT;
-
-(*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 -" complexT;
-
-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 *" complexT;
-
-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 [complex_add_zero_left, complex_add_zero_right];
-val mult_plus_1s = map rename_numerals [complex_mult_one_left, complex_mult_one_right];
-val mult_minus_1s = map rename_numerals
-                      [complex_mult_minus1, complex_mult_minus1_right];
-val mult_1s = mult_plus_1s @ mult_minus_1s;
-
-(*To perform binary arithmetic*)
-val bin_simps =
-    [complex_numeral_0_eq_0 RS sym, complex_numeral_1_eq_1 RS sym,
-     add_complex_number_of, complex_add_number_of_left, 
-     minus_complex_number_of, diff_complex_number_of, mult_complex_number_of, 
-     complex_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 \\ [complex_add_number_of_left, add_complex_number_of];
-
-(*To evaluate binary negations of coefficients*)
-val complex_minus_simps = NCons_simps @
-                   [complex_minus_1_eq_m1,minus_complex_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 = [complex_diff_def, minus_add_distrib, minus_minus];
-
-(* push the unary minus down: - x * y = x * - y *)
-val complex_minus_mult_eq_1_to_2 = 
-    [minus_mult_left RS sym, minus_mult_right] MRS trans 
-    |> standard;
-
-(*to extract again any uncancelled minuses*)
-val complex_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 complex_mult_minus_simps =
-    [mult_assoc, minus_mult_left, complex_minus_mult_eq_1_to_2];
-
-(*Final simplification: cancel + and *  *)
-val simplify_meta_eq = 
-    Int_Numeral_Simprocs.simplify_meta_eq
-         [add_zero_left, add_zero_right,
- 	  mult_zero_left, mult_zero_right, mult_1, mult_1_right];
-
-val prep_simproc = Real_Numeral_Simprocs.prep_simproc;
-
-
-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         = Real_Numeral_Simprocs.trans_tac
-  val norm_tac = 
-     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
-                                         complex_minus_simps@add_ac))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@complex_mult_minus_simps))
-     THEN ALLGOALS
-              (simp_tac (HOL_ss addsimps complex_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 =" complexT
-  val bal_add1 = eq_add_iff1 RS trans
-  val bal_add2 = eq_add_iff2 RS trans
-);
-
-
-val cancel_numerals = 
-  map prep_simproc
-   [("complexeq_cancel_numerals",
-               ["(l::complex) + m = n", "(l::complex) = m + n", 
-		"(l::complex) - m = n", "(l::complex) = m - n", 
-		"(l::complex) * m = n", "(l::complex) = m * n"], 
-     EqCancelNumerals.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         = Real_Numeral_Simprocs.trans_tac
-  val norm_tac = 
-     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
-                                         complex_minus_simps@add_ac))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@complex_mult_minus_simps))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps complex_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 CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
-
-val combine_numerals = 
-    prep_simproc ("complex_combine_numerals",
-		  ["(i::complex) + j", "(i::complex) - 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 
-        [complex_mult_one_left, complex_mult_one_right] 
-        (([th, cancel_th]) MRS trans);
-
-(*** Making constant folding work for 0 and 1 too ***)
-
-structure ComplexAbstractNumeralsData =
-  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  = complex_numeral_0_eq_0
-  val numeral_1_eq_1  = complex_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 ComplexAbstractNumerals = AbstractNumeralsFun (ComplexAbstractNumeralsData);
-
-(*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
-   [("complex_add_eval_numerals",
-     ["(m::complex) + 1", "(m::complex) + number_of v"],
-     ComplexAbstractNumerals.proc add_complex_number_of),
-    ("complex_diff_eval_numerals",
-     ["(m::complex) - 1", "(m::complex) - number_of v"],
-     ComplexAbstractNumerals.proc diff_complex_number_of),
-    ("complex_eq_eval_numerals",
-     ["(m::complex) = 0", "(m::complex) = 1", "(m::complex) = number_of v"],
-     ComplexAbstractNumerals.proc eq_complex_number_of)];
-
-end;
-
-Addsimprocs Complex_Numeral_Simprocs.eval_numerals;
-Addsimprocs Complex_Numeral_Simprocs.cancel_numerals;
-Addsimprocs [Complex_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::complex)";
-test " 2*u = (u::complex)";
-test "(i + j +  12 + (k::complex)) -  15 = y";
-test "(i + j +  12 + (k::complex)) -  5 = y";
-
-test "( 2*x - (u*v) + y) - v* 3*u = (w::complex)";
-test "( 2*x*u*v + (u*v)* 4 + y) - v*u* 4 = (w::complex)";
-test "( 2*x*u*v + (u*v)* 4 + y) - v*u = (w::complex)";
-test "u*v - (x*u*v + (u*v)* 4 + y) = (w::complex)";
-
-test "(i + j +  12 + (k::complex)) = u +  15 + y";
-test "(i + j* 2 +  12 + (k::complex)) = 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::complex)";
-
-test "a + -(b+c) + b = (d::complex)";
-test "a + -(b+c) - b = (d::complex)";
-
-(*negative numerals*)
-test "(i + j +  -2 + (k::complex)) - (u +  5 + y) = zz";
-
-test "(i + j +  -12 + (k::complex)) -  15 = y";
-test "(i + j +  12 + (k::complex)) -  -15 = y";
-test "(i + j +  -12 + (k::complex)) -  -15 = y";
-
-*)
-
-
-(** Constant folding for complex plus and times **)
-
-structure Complex_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	     = Complex_Numeral_Simprocs.complexT
-  val plus   = Const ("op *", [T,T] ---> T)
-  val add_ac = mult_ac
-end;
-
-structure Complex_Times_Assoc = Assoc_Fold (Complex_Times_Assoc_Data);
-
-Addsimprocs [Complex_Times_Assoc.conv];
-
-Addsimps [complex_of_real_zero_iff];
-
-
-(*Convert???
-Goalw [complex_number_of_def] 
-  "((number_of xa :: complex) + ii * number_of ya = \
-\       number_of xb) = \
-\  (((number_of xa :: complex) = number_of xb) & \
-\   ((number_of ya :: complex) = 0))";
-by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff2,
-    complex_of_real_zero_iff]));
-qed "complex_number_of_eq_cancel_iff2";
-Addsimps [complex_number_of_eq_cancel_iff2];
-
-Goalw [complex_number_of_def] 
-  "((number_of xa :: complex) + number_of ya * ii = \
-\       number_of xb) = \
-\  (((number_of xa :: complex) = number_of xb) & \
-\   ((number_of ya :: complex) = 0))";
-by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff2a,
-    complex_of_real_zero_iff]));
-qed "complex_number_of_eq_cancel_iff2a";
-Addsimps [complex_number_of_eq_cancel_iff2a];
-
-Goalw [complex_number_of_def] 
-  "((number_of xa :: complex) + ii * number_of ya = \
-\    ii * number_of yb) = \
-\  (((number_of xa :: complex) = 0) & \
-\   ((number_of ya :: complex) = number_of yb))";
-by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff3,
-    complex_of_real_zero_iff]));
-qed "complex_number_of_eq_cancel_iff3";
-Addsimps [complex_number_of_eq_cancel_iff3];
-
-Goalw [complex_number_of_def] 
-  "((number_of xa :: complex) + number_of ya * ii= \
-\    ii * number_of yb) = \
-\  (((number_of xa :: complex) = 0) & \
-\   ((number_of ya :: complex) = number_of yb))";
-by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff3a,
-    complex_of_real_zero_iff]));
-qed "complex_number_of_eq_cancel_iff3a";
-Addsimps [complex_number_of_eq_cancel_iff3a];
-*)
-
-Goalw [complex_number_of_def] "cnj (number_of v :: complex) = number_of v";
-by (rtac complex_cnj_complex_of_real 1);
-qed "complex_number_of_cnj";
-Addsimps [complex_number_of_cnj];
-
-Goalw [complex_number_of_def] 
-      "cmod(number_of v :: complex) = abs (number_of v :: real)";
-by (auto_tac (claset(), HOL_ss addsimps [complex_mod_complex_of_real]));
-qed "complex_number_of_cmod";
-Addsimps [complex_number_of_cmod];
-
-Goalw [complex_number_of_def] 
-      "Re(number_of v :: complex) = number_of v";
-by (auto_tac (claset(), HOL_ss addsimps [Re_complex_of_real]));
-qed "complex_number_of_Re";
-Addsimps [complex_number_of_Re];
-
-Goalw [complex_number_of_def] 
-      "Im(number_of v :: complex) = 0";
-by (auto_tac (claset(), HOL_ss addsimps [Im_complex_of_real]));
-qed "complex_number_of_Im";
-Addsimps [complex_number_of_Im];
-
-Goalw [expi_def] 
-   "expi((2::complex) * complex_of_real pi * ii) = 1";
-by (auto_tac (claset(),simpset() addsimps [complex_Re_mult_eq,
-    complex_Im_mult_eq,cis_def]));
-qed "expi_two_pi_i";
-Addsimps [expi_two_pi_i];
-
-(*examples:
-print_depth 22;
-set timing;
-set trace_simp;
-fun test s = (Goal s, by (Simp_tac 1)); 
-
-test "23 * ii + 45 * ii= (x::complex)";
-
-test "5 * ii + 12 - 45 * ii= (x::complex)";
-test "5 * ii + 40 - 12 * ii + 9 = (x::complex) + 89 * ii";
-test "5 * ii + 40 - 12 * ii + 9 - 78 = (x::complex) + 89 * ii";
-
-test "l + 10 * ii + 90 + 3*l +  9 + 45 * ii= (x::complex)";
-test "87 + 10 * ii + 90 + 3*7 +  9 + 45 * ii= (x::complex)";
-
-
-*)
--- a/src/HOL/Complex/ComplexBin.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Complex/ComplexBin.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -5,18 +5,5 @@
                 This case is reduced to that for the reals.
 *)
 
-ComplexBin = Complex + 
-
-
-instance
-  complex :: number 
-
-instance complex :: plus_ac0(complex_add_commute,complex_add_assoc,complex_add_zero_left)
+theory ComplexBin = Complex:
 
-
-defs
-  complex_number_of_def
-    "number_of v == complex_of_real (number_of v)"
-     (*::bin=>complex               ::bin=>complex*)
-
-end
--- a/src/HOL/Complex/NSCA.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Complex/NSCA.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -5,7 +5,7 @@
 *)
 
 val complex_induct = thm"complex.induct";
-
+val hcomplex_number_of = thm"hcomplex_number_of";
 
 (*--------------------------------------------------------------------------------------*)
 (* Closure laws for members of (embedded) set standard complex SComplex                 *)
@@ -58,8 +58,8 @@
 qed "SReal_hcmod_hcomplex_of_complex";
 Addsimps [SReal_hcmod_hcomplex_of_complex];
 
-Goalw [hcomplex_number_of_def]
-    "hcmod (number_of w ::hcomplex) : Reals";
+Goal "hcmod (number_of w ::hcomplex) : Reals";
+by (stac (hcomplex_number_of RS sym) 1); 
 by (rtac SReal_hcmod_hcomplex_of_complex 1);
 qed "SReal_hcmod_number_of";
 Addsimps [SReal_hcmod_number_of];
@@ -73,7 +73,8 @@
 qed "SComplex_hcomplex_of_complex";
 Addsimps [SComplex_hcomplex_of_complex];
 
-Goalw [hcomplex_number_of_def] "(number_of w ::hcomplex) : SComplex";
+Goal "(number_of w ::hcomplex) : SComplex";
+by (stac (hcomplex_number_of RS sym) 1); 
 by (rtac SComplex_hcomplex_of_complex 1);
 qed "SComplex_number_of";
 Addsimps [SComplex_number_of];
@@ -122,7 +123,7 @@
 qed "SComplex_hcmod_SReal";
 
 Goal "0 : SComplex";
-by (auto_tac (claset(),simpset() addsimps [SComplex_def]));
+by (auto_tac ((claset(),simpset() addsimps [SComplex_def]) addIffs [hcomplex_of_complex_zero_iff]));
 qed "SComplex_zero";
 Addsimps [SComplex_zero];
 
@@ -206,7 +207,7 @@
 AddIffs [CInfinitesimal_zero];
 
 Goal "x/(2::hcomplex) + x/(2::hcomplex) = x";
-by Auto_tac;  
+by Auto_tac;
 qed "hcomplex_sum_of_halves";
 
 Goalw [CInfinitesimal_def,Infinitesimal_def] 
@@ -954,6 +955,8 @@
     Reals_Re_Im_SComplex]) 1);
 qed "SComplex_SReal_iff";
 
+val hcomplex_zero_num = thm"hcomplex_zero_num";
+
 Goal "(Abs_hcomplex(hcomplexrel ``{%n. X n}) : CInfinitesimal) = \
 \     (Abs_hypreal(hyprel `` {%n. Re(X n)}) : Infinitesimal & \
 \      Abs_hypreal(hyprel `` {%n. Im(X n)}) : Infinitesimal)";
--- a/src/HOL/Complex/NSCA.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Complex/NSCA.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -4,7 +4,7 @@
     Description : Infinite, infinitesimal complex number etc! 
 *)
 
-NSCA = NSComplexArith + 
+NSCA = NSComplex + 
 
 consts   
 
--- a/src/HOL/Complex/NSComplex.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Complex/NSComplex.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -1368,8 +1368,10 @@
 lemma hcomplex_of_complex_zero [simp]: "hcomplex_of_complex 0 = 0"
 by (simp add: hcomplex_of_complex_def hcomplex_zero_def)
 
-lemma hcomplex_of_complex_zero_iff: "(hcomplex_of_complex r = 0) = (r = 0)"
-by (auto intro: FreeUltrafilterNat_P simp add: hcomplex_of_complex_def hcomplex_zero_def)
+lemma hcomplex_of_complex_zero_iff [simp]:
+     "(hcomplex_of_complex r = 0) = (r = 0)"
+by (auto intro: FreeUltrafilterNat_P 
+         simp add: hcomplex_of_complex_def hcomplex_zero_def)
 
 lemma hcomplex_of_complex_inverse [simp]:
      "hcomplex_of_complex (inverse r) = inverse (hcomplex_of_complex r)"
@@ -1398,6 +1400,199 @@
      "hcmod (hcomplex_of_complex x) = hypreal_of_real (cmod x)"
 by (simp add: hypreal_of_real_def hcomplex_of_complex_def hcmod)
 
+
+subsection{*Numerals and Arithmetic*}
+
+instance hcomplex :: number ..
+
+primrec (*the type constraint is essential!*)
+  number_of_Pls: "number_of bin.Pls = 0"
+  number_of_Min: "number_of bin.Min = - (1::hcomplex)"
+  number_of_BIT: "number_of(w BIT x) = (if x then 1 else 0) +
+	                               (number_of w) + (number_of w)"
+
+declare number_of_Pls [simp del]
+        number_of_Min [simp del]
+        number_of_BIT [simp del]
+
+instance hcomplex :: number_ring
+proof
+  show "Numeral0 = (0::hcomplex)" by (rule number_of_Pls)
+  show "-1 = - (1::hcomplex)" by (rule number_of_Min)
+  fix w :: bin and x :: bool
+  show "(number_of (w BIT x) :: hcomplex) =
+        (if x then 1 else 0) + number_of w + number_of w"
+    by (rule number_of_BIT)
+qed
+
+
+text{*Collapse applications of @{term hcomplex_of_complex} to @{term number_of}*}
+lemma hcomplex_number_of [simp]:
+     "hcomplex_of_complex (number_of w) = number_of w"
+apply (induct w) 
+apply (simp_all only: number_of hcomplex_of_complex_add 
+                      hcomplex_of_complex_minus, simp_all) 
+done
+
+lemma hcomplex_of_hypreal_eq_hcomplex_of_complex: 
+     "hcomplex_of_hypreal (hypreal_of_real x) =  
+      hcomplex_of_complex(complex_of_real x)"
+by (simp add: hypreal_of_real_def hcomplex_of_hypreal hcomplex_of_complex_def 
+              complex_of_real_def)
+
+lemma hcomplex_hypreal_number_of: 
+  "hcomplex_of_complex (number_of w) = hcomplex_of_hypreal(number_of w)"
+by (simp only: complex_number_of [symmetric] hypreal_number_of [symmetric] 
+               hcomplex_of_hypreal_eq_hcomplex_of_complex)
+
+text{*This theorem is necessary because theorems such as
+   @{text iszero_number_of_0} only hold for ordered rings. They cannot
+   be generalized to fields in general because they fail for finite fields.
+   They work for type complex because the reals can be embedded in them.*}
+lemma iszero_hcomplex_number_of [simp]:
+     "iszero (number_of w :: hcomplex) = iszero (number_of w :: real)"
+apply (simp only: iszero_complex_number_of [symmetric])  
+apply (simp only: hcomplex_of_complex_zero_iff hcomplex_number_of [symmetric] 
+                  iszero_def)  
+done
+
+
+(*
+Goal "z + hcnj z =  
+      hcomplex_of_hypreal (2 * hRe(z))"
+by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
+by (auto_tac (claset(),HOL_ss addsimps [hRe,hcnj,hcomplex_add,
+    hypreal_mult,hcomplex_of_hypreal,complex_add_cnj]));
+qed "hcomplex_add_hcnj";
+
+Goal "z - hcnj z = \
+\     hcomplex_of_hypreal (hypreal_of_real #2 * hIm(z)) * iii";
+by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
+by (auto_tac (claset(),simpset() addsimps [hIm,hcnj,hcomplex_diff,
+    hypreal_of_real_def,hypreal_mult,hcomplex_of_hypreal,
+    complex_diff_cnj,iii_def,hcomplex_mult]));
+qed "hcomplex_diff_hcnj";
+*)
+
+
+lemma hcomplex_hcnj_num_zero_iff: "(hcnj z = 0) = (z = 0)"
+apply (auto simp add: hcomplex_hcnj_zero_iff)
+done
+declare hcomplex_hcnj_num_zero_iff [simp]
+
+lemma hcomplex_zero_num: "0 = Abs_hcomplex (hcomplexrel `` {%n. 0})"
+apply (simp add: hcomplex_zero_def)
+done
+
+lemma hcomplex_one_num: "1 =  Abs_hcomplex (hcomplexrel `` {%n. 1})"
+apply (simp add: hcomplex_one_def)
+done
+
+(*** Real and imaginary stuff ***)
+
+(*Convert???
+Goalw [hcomplex_number_of_def] 
+  "((number_of xa :: hcomplex) + iii * number_of ya =  
+        number_of xb + iii * number_of yb) =  
+   (((number_of xa :: hcomplex) = number_of xb) &  
+    ((number_of ya :: hcomplex) = number_of yb))"
+by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff,
+     hcomplex_hypreal_number_of]));
+qed "hcomplex_number_of_eq_cancel_iff";
+Addsimps [hcomplex_number_of_eq_cancel_iff];
+
+Goalw [hcomplex_number_of_def] 
+  "((number_of xa :: hcomplex) + number_of ya * iii = \
+\       number_of xb + number_of yb * iii) = \
+\  (((number_of xa :: hcomplex) = number_of xb) & \
+\   ((number_of ya :: hcomplex) = number_of yb))";
+by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffA,
+    hcomplex_hypreal_number_of]));
+qed "hcomplex_number_of_eq_cancel_iffA";
+Addsimps [hcomplex_number_of_eq_cancel_iffA];
+
+Goalw [hcomplex_number_of_def] 
+  "((number_of xa :: hcomplex) + number_of ya * iii = \
+\       number_of xb + iii * number_of yb) = \
+\  (((number_of xa :: hcomplex) = number_of xb) & \
+\   ((number_of ya :: hcomplex) = number_of yb))";
+by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffB,
+    hcomplex_hypreal_number_of]));
+qed "hcomplex_number_of_eq_cancel_iffB";
+Addsimps [hcomplex_number_of_eq_cancel_iffB];
+
+Goalw [hcomplex_number_of_def] 
+  "((number_of xa :: hcomplex) + iii * number_of ya = \
+\       number_of xb + number_of yb * iii) = \
+\  (((number_of xa :: hcomplex) = number_of xb) & \
+\   ((number_of ya :: hcomplex) = number_of yb))";
+by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffC,
+     hcomplex_hypreal_number_of]));
+qed "hcomplex_number_of_eq_cancel_iffC";
+Addsimps [hcomplex_number_of_eq_cancel_iffC];
+
+Goalw [hcomplex_number_of_def] 
+  "((number_of xa :: hcomplex) + iii * number_of ya = \
+\       number_of xb) = \
+\  (((number_of xa :: hcomplex) = number_of xb) & \
+\   ((number_of ya :: hcomplex) = 0))";
+by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff2,
+    hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff]));
+qed "hcomplex_number_of_eq_cancel_iff2";
+Addsimps [hcomplex_number_of_eq_cancel_iff2];
+
+Goalw [hcomplex_number_of_def] 
+  "((number_of xa :: hcomplex) + number_of ya * iii = \
+\       number_of xb) = \
+\  (((number_of xa :: hcomplex) = number_of xb) & \
+\   ((number_of ya :: hcomplex) = 0))";
+by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff2a,
+    hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff]));
+qed "hcomplex_number_of_eq_cancel_iff2a";
+Addsimps [hcomplex_number_of_eq_cancel_iff2a];
+
+Goalw [hcomplex_number_of_def] 
+  "((number_of xa :: hcomplex) + iii * number_of ya = \
+\    iii * number_of yb) = \
+\  (((number_of xa :: hcomplex) = 0) & \
+\   ((number_of ya :: hcomplex) = number_of yb))";
+by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff3,
+    hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff]));
+qed "hcomplex_number_of_eq_cancel_iff3";
+Addsimps [hcomplex_number_of_eq_cancel_iff3];
+
+Goalw [hcomplex_number_of_def] 
+  "((number_of xa :: hcomplex) + number_of ya * iii= \
+\    iii * number_of yb) = \
+\  (((number_of xa :: hcomplex) = 0) & \
+\   ((number_of ya :: hcomplex) = number_of yb))";
+by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff3a,
+    hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff]));
+qed "hcomplex_number_of_eq_cancel_iff3a";
+Addsimps [hcomplex_number_of_eq_cancel_iff3a];
+*)
+
+lemma hcomplex_number_of_hcnj [simp]:
+     "hcnj (number_of v :: hcomplex) = number_of v"
+by (simp only: hcomplex_number_of [symmetric] hcomplex_hypreal_number_of
+               hcomplex_hcnj_hcomplex_of_hypreal)
+
+lemma hcomplex_number_of_hcmod [simp]: 
+      "hcmod(number_of v :: hcomplex) = abs (number_of v :: hypreal)"
+by (simp only: hcomplex_number_of [symmetric] hcomplex_hypreal_number_of
+               hcmod_hcomplex_of_hypreal)
+
+lemma hcomplex_number_of_hRe [simp]: 
+      "hRe(number_of v :: hcomplex) = number_of v"
+by (simp only: hcomplex_number_of [symmetric] hcomplex_hypreal_number_of
+               hRe_hcomplex_of_hypreal)
+
+lemma hcomplex_number_of_hIm [simp]: 
+      "hIm(number_of v :: hcomplex) = 0"
+by (simp only: hcomplex_number_of [symmetric] hcomplex_hypreal_number_of
+               hIm_hcomplex_of_hypreal)
+
+
 ML
 {*
 val hcomplex_zero_def = thm"hcomplex_zero_def";
--- a/src/HOL/Complex/NSComplexArith.thy	Sat Feb 14 02:06:12 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,19 +0,0 @@
-(*  Title:       NSComplexArith
-    Author:      Lawrence C. Paulson
-    Copyright:   2003  University of Cambridge
-
-Common factor cancellation
-*)
-
-theory NSComplexArith = NSComplexBin
-files "hcomplex_arith.ML":
-
-subsubsection{*Division By @{term "-1"}*}
-
-lemma hcomplex_divide_minus1 [simp]: "x/-1 = -(x::hcomplex)"
-by simp
-
-lemma hcomplex_minus1_divide [simp]: "-1/(x::hcomplex) = - (1/x)"
-by (simp add: hcomplex_divide_def inverse_minus_eq)
-
-end
--- a/src/HOL/Complex/NSComplexBin.ML	Sat Feb 14 02:06:12 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,595 +0,0 @@
-(*  Title:      NSComplexBin.ML
-    Author:     Jacques D. Fleuriot
-    Copyright:  2001 University of Edinburgh
-    Descrition: Binary arithmetic for the nonstandard complex numbers
-*)
-
-(** hcomplex_of_complex (coercion from complex to nonstandard complex) **)
-
-Goal "hcomplex_of_complex (number_of w) = number_of w";
-by (simp_tac (simpset() addsimps [hcomplex_number_of_def]) 1);
-qed "hcomplex_number_of";
-Addsimps [hcomplex_number_of];
-
-Goalw [hypreal_of_real_def]
-     "hcomplex_of_hypreal (hypreal_of_real x) = \
-\     hcomplex_of_complex(complex_of_real x)";
-by (simp_tac (simpset() addsimps [hcomplex_of_hypreal,
-    hcomplex_of_complex_def,complex_of_real_def]) 1);
-qed "hcomplex_of_hypreal_eq_hcomplex_of_complex";
-
-Goalw [complex_number_of_def,hypreal_number_of_def] 
-  "hcomplex_of_complex (number_of w) = hcomplex_of_hypreal(number_of w)";
-by (rtac (hcomplex_of_hypreal_eq_hcomplex_of_complex RS sym) 1);
-qed "hcomplex_hypreal_number_of";
-
-Goalw [hcomplex_number_of_def] "Numeral0 = (0::hcomplex)";
-by(Simp_tac 1);
-qed "hcomplex_numeral_0_eq_0";
-
-Goalw [hcomplex_number_of_def] "Numeral1 = (1::hcomplex)";
-by(Simp_tac 1);
-qed "hcomplex_numeral_1_eq_1";
-
-(*
-Goal "z + hcnj z = \
-\     hcomplex_of_hypreal (2 * hRe(z))";
-by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
-by (auto_tac (claset(),HOL_ss addsimps [hRe,hcnj,hcomplex_add,
-    hypreal_mult,hcomplex_of_hypreal,complex_add_cnj]));
-qed "hcomplex_add_hcnj";
-
-Goal "z - hcnj z = \
-\     hcomplex_of_hypreal (hypreal_of_real #2 * hIm(z)) * iii";
-by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
-by (auto_tac (claset(),simpset() addsimps [hIm,hcnj,hcomplex_diff,
-    hypreal_of_real_def,hypreal_mult,hcomplex_of_hypreal,
-    complex_diff_cnj,iii_def,hcomplex_mult]));
-qed "hcomplex_diff_hcnj";
-*)
-
-(** Addition **)
-
-Goal "(number_of v :: hcomplex) + number_of v' = number_of (bin_add v v')";
-by (simp_tac
-    (HOL_ss addsimps [hcomplex_number_of_def, 
-                      hcomplex_of_complex_add RS sym, add_complex_number_of]) 1);
-qed "add_hcomplex_number_of";
-Addsimps [add_hcomplex_number_of];
-
-
-(** Subtraction **)
-
-Goalw [hcomplex_number_of_def]
-     "- (number_of w :: hcomplex) = number_of (bin_minus w)";
-by (simp_tac
-    (HOL_ss addsimps [minus_complex_number_of, hcomplex_of_complex_minus RS sym]) 1);
-qed "minus_hcomplex_number_of";
-Addsimps [minus_hcomplex_number_of];
-
-Goalw [hcomplex_number_of_def, hcomplex_diff_def]
-     "(number_of v :: hcomplex) - number_of w = \
-\     number_of (bin_add v (bin_minus w))";
-by (Simp_tac 1); 
-qed "diff_hcomplex_number_of";
-Addsimps [diff_hcomplex_number_of];
-
-
-(** Multiplication **)
-
-Goal "(number_of v :: hcomplex) * number_of v' = number_of (bin_mult v v')";
-by (simp_tac
-    (HOL_ss addsimps [hcomplex_number_of_def, 
-	              hcomplex_of_complex_mult RS sym, mult_complex_number_of]) 1);
-qed "mult_hcomplex_number_of";
-Addsimps [mult_hcomplex_number_of];
-
-Goal "(2::hcomplex) = 1 + 1";
-by (simp_tac (simpset() addsimps [hcomplex_numeral_1_eq_1 RS sym]) 1);
-val lemma = result();
-
-(*For specialist use: NOT as default simprules*)
-Goal "2 * z = (z+z::hcomplex)";
-by (simp_tac (simpset() addsimps [lemma, hcomplex_add_mult_distrib]) 1);
-qed "hcomplex_mult_2";
-
-Goal "z * 2 = (z+z::hcomplex)";
-by (stac hcomplex_mult_commute 1 THEN rtac hcomplex_mult_2 1);
-qed "hcomplex_mult_2_right";
-
-(** Equals (=) **)
-
-Goal "((number_of v :: hcomplex) = number_of v') = \
-\     iszero (number_of (bin_add v (bin_minus v')) :: int)";
-by (simp_tac
-    (HOL_ss addsimps [hcomplex_number_of_def, 
-	              hcomplex_of_complex_eq_iff, eq_complex_number_of]) 1);
-qed "eq_hcomplex_number_of";
-Addsimps [eq_hcomplex_number_of];
-
-(*** New versions of existing theorems involving 0, 1hc ***)
-
-Goal "- 1 = (-1::hcomplex)";
-by (simp_tac (simpset() addsimps [hcomplex_numeral_1_eq_1 RS sym]) 1);
-qed "hcomplex_minus_1_eq_m1";
-
-Goal "-1 * z = -(z::hcomplex)";
-by (simp_tac (simpset() addsimps [hcomplex_minus_1_eq_m1 RS sym]) 1);
-qed "hcomplex_mult_minus1";
-
-Goal "z * -1 = -(z::hcomplex)";
-by (stac hcomplex_mult_commute 1 THEN rtac hcomplex_mult_minus1 1);
-qed "hcomplex_mult_minus1_right";
-
-Addsimps [hcomplex_mult_minus1,hcomplex_mult_minus1_right];
-
-(*Maps 0 to Numeral0 and 1 to Numeral1 and -Numeral1 to -1*)
-val hcomplex_numeral_ss = 
-    complex_numeral_ss addsimps [hcomplex_numeral_0_eq_0 RS sym, hcomplex_numeral_1_eq_1 RS sym, 
-		                 hcomplex_minus_1_eq_m1];
-
-fun rename_numerals th = 
-    asm_full_simplify hcomplex_numeral_ss (Thm.transfer (the_context ()) th);
-
-
-(*Now insert some identities previously stated for 0 and 1hc*)
-
-
-Addsimps [hcomplex_numeral_0_eq_0,hcomplex_numeral_1_eq_1];
-
-Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::hcomplex)";
-by (auto_tac (claset(),simpset() addsimps [hcomplex_add_assoc RS sym]));
-qed "hcomplex_add_number_of_left";
-
-Goal "number_of v *(number_of w * z) = (number_of(bin_mult v w) * z::hcomplex)";
-by (simp_tac (simpset() addsimps [hcomplex_mult_assoc RS sym]) 1);
-qed "hcomplex_mult_number_of_left";
-
-Goalw [hcomplex_diff_def]
-    "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::hcomplex)";
-by (rtac hcomplex_add_number_of_left 1);
-qed "hcomplex_add_number_of_diff1";
-
-Goal "number_of v + (c - number_of w) = \
-\     number_of (bin_add v (bin_minus w)) + (c::hcomplex)";
-by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_def]@ add_ac));
-qed "hcomplex_add_number_of_diff2";
-
-Addsimps [hcomplex_add_number_of_left, hcomplex_mult_number_of_left,
-	  hcomplex_add_number_of_diff1, hcomplex_add_number_of_diff2]; 
-
-
-(**** Simprocs for numeric literals ****)
-
-structure HComplex_Numeral_Simprocs =
-struct
-
-(*Utilities*)
-
-val hcomplexT = Type("NSComplex.hcomplex",[]);
-
-fun mk_numeral n = HOLogic.number_of_const hcomplexT $ HOLogic.mk_bin n;
-
-val dest_numeral = Complex_Numeral_Simprocs.dest_numeral;
-
-val find_first_numeral = Complex_Numeral_Simprocs.find_first_numeral;
-
-val zero = mk_numeral 0;
-val mk_plus = HOLogic.mk_binop "op +";
-
-val uminus_const = Const ("uminus", hcomplexT --> hcomplexT);
-
-(*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 +" hcomplexT;
-
-(*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 -" hcomplexT;
-
-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 *" hcomplexT;
-
-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 [hcomplex_add_zero_left, hcomplex_add_zero_right];
-val mult_plus_1s = map rename_numerals [hcomplex_mult_one_left, hcomplex_mult_one_right];
-val mult_minus_1s = map rename_numerals [hcomplex_mult_minus1, hcomplex_mult_minus1_right];
-val mult_1s = mult_plus_1s @ mult_minus_1s;
-
-(*To perform binary arithmetic*)
-val bin_simps =
-    [hcomplex_numeral_0_eq_0 RS sym, hcomplex_numeral_1_eq_1 RS sym,
-     add_hcomplex_number_of, hcomplex_add_number_of_left, 
-     minus_hcomplex_number_of, diff_hcomplex_number_of, mult_hcomplex_number_of, 
-     hcomplex_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 \\ [hcomplex_add_number_of_left, add_hcomplex_number_of];
-
-(*To evaluate binary negations of coefficients*)
-val hcomplex_minus_simps = NCons_simps @
-                   [hcomplex_minus_1_eq_m1,minus_hcomplex_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 = [hcomplex_diff_def, minus_add_distrib, minus_minus];
-
-(*push the unary minus down: - x * y = x * - y *)
-val hcomplex_minus_mult_eq_1_to_2 = 
-    [minus_mult_left RS sym, minus_mult_right] MRS trans 
-    |> standard;
-
-(*to extract again any uncancelled minuses*)
-val hcomplex_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 hcomplex_mult_minus_simps =
-    [hcomplex_mult_assoc, minus_mult_left, hcomplex_minus_mult_eq_1_to_2];
-
-(*Final simplification: cancel + and *  *)
-val simplify_meta_eq = 
-    Int_Numeral_Simprocs.simplify_meta_eq
-         [add_zero_left, add_zero_right,
- 	  mult_zero_left, mult_zero_right, mult_1, mult_1_right];
-
-val prep_simproc = Complex_Numeral_Simprocs.prep_simproc;
-
-
-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         = Real_Numeral_Simprocs.trans_tac
-  val norm_tac = 
-     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
-                                         hcomplex_minus_simps@add_ac))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@hcomplex_mult_minus_simps))
-     THEN ALLGOALS
-              (simp_tac (HOL_ss addsimps hcomplex_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 =" hcomplexT
-  val bal_add1 = eq_add_iff1 RS trans
-  val bal_add2 = eq_add_iff2 RS trans
-);
-
-
-val cancel_numerals = 
-  map prep_simproc
-   [("hcomplexeq_cancel_numerals",
-      ["(l::hcomplex) + m = n", "(l::hcomplex) = m + n", 
-		"(l::hcomplex) - m = n", "(l::hcomplex) = m - n", 
-		"(l::hcomplex) * m = n", "(l::hcomplex) = m * n"], 
-     EqCancelNumerals.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         = Real_Numeral_Simprocs.trans_tac
-  val norm_tac = 
-     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
-                                         hcomplex_minus_simps@add_ac))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@hcomplex_mult_minus_simps))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps hcomplex_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 CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
-
-val combine_numerals = 
-    prep_simproc ("hcomplex_combine_numerals",
-		  ["(i::hcomplex) + j", "(i::hcomplex) - 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 
-        [hcomplex_mult_one_left, hcomplex_mult_one_right] 
-        (([th, cancel_th]) MRS trans);
-
-(*** Making constant folding work for 0 and 1 too ***)
-
-structure HComplexAbstractNumeralsData =
-  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  = hcomplex_numeral_0_eq_0
-  val numeral_1_eq_1  = hcomplex_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 HComplexAbstractNumerals = AbstractNumeralsFun (HComplexAbstractNumeralsData)
-
-(*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
-   [("hcomplex_add_eval_numerals",
-     ["(m::hcomplex) + 1", "(m::hcomplex) + number_of v"],
-     HComplexAbstractNumerals.proc add_hcomplex_number_of),
-    ("hcomplex_diff_eval_numerals",
-     ["(m::hcomplex) - 1", "(m::hcomplex) - number_of v"],
-     HComplexAbstractNumerals.proc diff_hcomplex_number_of),
-    ("hcomplex_eq_eval_numerals",
-     ["(m::hcomplex) = 0", "(m::hcomplex) = 1", "(m::hcomplex) = number_of v"],
-     HComplexAbstractNumerals.proc eq_hcomplex_number_of)]
-
-end;
-
-Addsimprocs HComplex_Numeral_Simprocs.eval_numerals;
-Addsimprocs HComplex_Numeral_Simprocs.cancel_numerals;
-Addsimprocs [HComplex_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::hcomplex)";
-test " 2*u = (u::hcomplex)";
-test "(i + j + 12 + (k::hcomplex)) - 15 = y";
-test "(i + j + 12 + (k::hcomplex)) -  5 = y";
-
-test "( 2*x - (u*v) + y) - v* 3*u = (w::hcomplex)";
-test "( 2*x*u*v + (u*v)* 4 + y) - v*u* 4 = (w::hcomplex)";
-test "( 2*x*u*v + (u*v)* 4 + y) - v*u = (w::hcomplex)";
-test "u*v - (x*u*v + (u*v)* 4 + y) = (w::hcomplex)";
-
-test "(i + j + 12 + (k::hcomplex)) = u + 15 + y";
-test "(i + j* 2 + 12 + (k::hcomplex)) = 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::hcomplex)";
-
-test "a + -(b+c) + b = (d::hcomplex)";
-test "a + -(b+c) - b = (d::hcomplex)";
-
-(*negative numerals*)
-test "(i + j +  -2 + (k::hcomplex)) - (u +  5 + y) = zz";
-
-test "(i + j +  -12 + (k::hcomplex)) - 15 = y";
-test "(i + j + 12 + (k::hcomplex)) -  -15 = y";
-test "(i + j +  -12 + (k::hcomplex)) - -15 = y";
-*)
-
-(** Constant folding for hcomplex plus and times **)
-
-structure HComplex_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	     = HComplex_Numeral_Simprocs.hcomplexT
-  val plus   = Const ("op *", [T,T] ---> T)
-  val add_ac = mult_ac
-end;
-
-structure HComplex_Times_Assoc = Assoc_Fold (HComplex_Times_Assoc_Data);
-
-Addsimprocs [HComplex_Times_Assoc.conv];
-
-Addsimps [hcomplex_of_complex_zero_iff];
-
-
-(** extra thms **)
-
-Goal "(hcnj z = 0) = (z = 0)";
-by (auto_tac (claset(),simpset() addsimps [hcomplex_hcnj_zero_iff]));
-qed "hcomplex_hcnj_num_zero_iff";
-Addsimps [hcomplex_hcnj_num_zero_iff];
-
-Goal "0 = Abs_hcomplex (hcomplexrel `` {%n. 0})";
-by (simp_tac (simpset() addsimps [hcomplex_zero_def RS meta_eq_to_obj_eq RS sym]) 1);
-qed "hcomplex_zero_num";
-
-Goal "1 =  Abs_hcomplex (hcomplexrel `` {%n. 1})";
-by (simp_tac (simpset() addsimps [hcomplex_one_def RS meta_eq_to_obj_eq RS sym]) 1);
-qed "hcomplex_one_num";
-
-(*** Real and imaginary stuff ***)
-
-(*Convert???
-Goalw [hcomplex_number_of_def] 
-  "((number_of xa :: hcomplex) + iii * number_of ya = \
-\       number_of xb + iii * number_of yb) = \
-\  (((number_of xa :: hcomplex) = number_of xb) & \
-\   ((number_of ya :: hcomplex) = number_of yb))";
-by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff,
-     hcomplex_hypreal_number_of]));
-qed "hcomplex_number_of_eq_cancel_iff";
-Addsimps [hcomplex_number_of_eq_cancel_iff];
-
-Goalw [hcomplex_number_of_def] 
-  "((number_of xa :: hcomplex) + number_of ya * iii = \
-\       number_of xb + number_of yb * iii) = \
-\  (((number_of xa :: hcomplex) = number_of xb) & \
-\   ((number_of ya :: hcomplex) = number_of yb))";
-by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffA,
-    hcomplex_hypreal_number_of]));
-qed "hcomplex_number_of_eq_cancel_iffA";
-Addsimps [hcomplex_number_of_eq_cancel_iffA];
-
-Goalw [hcomplex_number_of_def] 
-  "((number_of xa :: hcomplex) + number_of ya * iii = \
-\       number_of xb + iii * number_of yb) = \
-\  (((number_of xa :: hcomplex) = number_of xb) & \
-\   ((number_of ya :: hcomplex) = number_of yb))";
-by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffB,
-    hcomplex_hypreal_number_of]));
-qed "hcomplex_number_of_eq_cancel_iffB";
-Addsimps [hcomplex_number_of_eq_cancel_iffB];
-
-Goalw [hcomplex_number_of_def] 
-  "((number_of xa :: hcomplex) + iii * number_of ya = \
-\       number_of xb + number_of yb * iii) = \
-\  (((number_of xa :: hcomplex) = number_of xb) & \
-\   ((number_of ya :: hcomplex) = number_of yb))";
-by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffC,
-     hcomplex_hypreal_number_of]));
-qed "hcomplex_number_of_eq_cancel_iffC";
-Addsimps [hcomplex_number_of_eq_cancel_iffC];
-
-Goalw [hcomplex_number_of_def] 
-  "((number_of xa :: hcomplex) + iii * number_of ya = \
-\       number_of xb) = \
-\  (((number_of xa :: hcomplex) = number_of xb) & \
-\   ((number_of ya :: hcomplex) = 0))";
-by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff2,
-    hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff]));
-qed "hcomplex_number_of_eq_cancel_iff2";
-Addsimps [hcomplex_number_of_eq_cancel_iff2];
-
-Goalw [hcomplex_number_of_def] 
-  "((number_of xa :: hcomplex) + number_of ya * iii = \
-\       number_of xb) = \
-\  (((number_of xa :: hcomplex) = number_of xb) & \
-\   ((number_of ya :: hcomplex) = 0))";
-by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff2a,
-    hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff]));
-qed "hcomplex_number_of_eq_cancel_iff2a";
-Addsimps [hcomplex_number_of_eq_cancel_iff2a];
-
-Goalw [hcomplex_number_of_def] 
-  "((number_of xa :: hcomplex) + iii * number_of ya = \
-\    iii * number_of yb) = \
-\  (((number_of xa :: hcomplex) = 0) & \
-\   ((number_of ya :: hcomplex) = number_of yb))";
-by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff3,
-    hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff]));
-qed "hcomplex_number_of_eq_cancel_iff3";
-Addsimps [hcomplex_number_of_eq_cancel_iff3];
-
-Goalw [hcomplex_number_of_def] 
-  "((number_of xa :: hcomplex) + number_of ya * iii= \
-\    iii * number_of yb) = \
-\  (((number_of xa :: hcomplex) = 0) & \
-\   ((number_of ya :: hcomplex) = number_of yb))";
-by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff3a,
-    hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff]));
-qed "hcomplex_number_of_eq_cancel_iff3a";
-Addsimps [hcomplex_number_of_eq_cancel_iff3a];
-*)
-
-Goalw [hcomplex_number_of_def] "hcnj (number_of v :: hcomplex) = number_of v";
-by (rtac (hcomplex_hypreal_number_of RS ssubst) 1);
-by (rtac hcomplex_hcnj_hcomplex_of_hypreal 1);
-qed "hcomplex_number_of_hcnj";
-Addsimps [hcomplex_number_of_hcnj];
-
-Goalw [hcomplex_number_of_def] 
-      "hcmod(number_of v :: hcomplex) = abs (number_of v :: hypreal)";
-by (rtac (hcomplex_hypreal_number_of RS ssubst) 1);
-by (auto_tac (claset(), HOL_ss addsimps [hcmod_hcomplex_of_hypreal]));
-qed "hcomplex_number_of_hcmod";
-Addsimps [hcomplex_number_of_hcmod];
-
-Goalw [hcomplex_number_of_def] 
-      "hRe(number_of v :: hcomplex) = number_of v";
-by (rtac (hcomplex_hypreal_number_of RS ssubst) 1);
-by (auto_tac (claset(), HOL_ss addsimps [hRe_hcomplex_of_hypreal]));
-qed "hcomplex_number_of_hRe";
-Addsimps [hcomplex_number_of_hRe];
-
-Goalw [hcomplex_number_of_def] 
-      "hIm(number_of v :: hcomplex) = 0";
-by (rtac (hcomplex_hypreal_number_of RS ssubst) 1);
-by (auto_tac (claset(), HOL_ss addsimps [hIm_hcomplex_of_hypreal]));
-qed "hcomplex_number_of_hIm";
-Addsimps [hcomplex_number_of_hIm];
-
-
-
--- a/src/HOL/Complex/NSComplexBin.thy	Sat Feb 14 02:06:12 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,19 +0,0 @@
-(*  Title:      NSComplexBin.thy
-    Author:     Jacques D. Fleuriot
-    Copyright:  2001 University of Edinburgh
-    Descrition: Binary arithmetic for the nonstandard complex numbers
-                This case is reduced to that for the complexes (hence reals).
-*)
-
-NSComplexBin = NSComplex + 
-
-
-instance
-  hcomplex :: number 
-
-defs
-  hcomplex_number_of_def
-    "number_of v == hcomplex_of_complex (number_of v)"
-     (*::bin=>complex               ::bin=>complex*)
-
-end
\ No newline at end of file
--- a/src/HOL/Complex/NSInduct.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Complex/NSInduct.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -4,7 +4,7 @@
     Description: Nonstandard characterization of induction etc.
 *)
 
-NSInduct =  ComplexArith0 + 
+NSInduct =  Complex + 
 
 constdefs
 
--- a/src/HOL/Complex/hcomplex_arith.ML	Sat Feb 14 02:06:12 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,164 +0,0 @@
-(*  Title:       hcomplex_arith.ML
-    Author:      Jacques D. Fleuriot
-    Copyright:   2001  University of Edinburgh
-
-Common factor cancellation
-*)
-
-local
-  open HComplex_Numeral_Simprocs
-in
-
-val rel_hcomplex_number_of = [eq_hcomplex_number_of];
-
-
-structure CancelNumeralFactorCommon =
-  struct
-  val mk_coeff		= mk_coeff
-  val dest_coeff	= dest_coeff 1
-  val trans_tac         = Real_Numeral_Simprocs.trans_tac
-  val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps hcomplex_minus_from_mult_simps @ mult_1s))
-                 THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@hcomplex_mult_minus_simps))
-                 THEN ALLGOALS (simp_tac (HOL_ss addsimps mult_ac))
-  val numeral_simp_tac	= 
-         ALLGOALS (simp_tac (HOL_ss addsimps rel_hcomplex_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" hcomplexT
-  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 =" hcomplexT
-  val cancel = field_mult_cancel_left RS trans
-  val neg_exchanges = false
-)
-
-
-val hcomplex_cancel_numeral_factors_relations = 
-  map prep_simproc
-   [("hcomplexeq_cancel_numeral_factor",
-    ["(l::hcomplex) * m = n", "(l::hcomplex) = m * n"], 
-     EqCancelNumeralFactor.proc)];
-
-val hcomplex_cancel_numeral_factors_divide = prep_simproc
-	("hcomplexdiv_cancel_numeral_factor", 
-	 ["((l::hcomplex) * m) / n", "(l::hcomplex) / (m * n)", 
-                     "((number_of v)::hcomplex) / (number_of w)"], 
-	 DivCancelNumeralFactor.proc);
-
-val hcomplex_cancel_numeral_factors = 
-    hcomplex_cancel_numeral_factors_relations @ 
-    [hcomplex_cancel_numeral_factors_divide];
-
-end;
-
-
-Addsimprocs hcomplex_cancel_numeral_factors;
-
-
-(*examples:
-print_depth 22;
-set timing;
-set trace_simp;
-fun test s = (Goal s; by (Simp_tac 1)); 
-
-
-test "#9*x = #12 * (y::hcomplex)";
-test "(#9*x) / (#12 * (y::hcomplex)) = z";
-
-test "#-99*x = #132 * (y::hcomplex)";
-
-test "#999*x = #-396 * (y::hcomplex)";
-test "(#999*x) / (#-396 * (y::hcomplex)) = z";
-
-test "#-99*x = #-81 * (y::hcomplex)";
-test "(#-99*x) / (#-81 * (y::hcomplex)) = z";
-
-test "#-2 * x = #-1 * (y::hcomplex)";
-test "#-2 * x = -(y::hcomplex)";
-test "(#-2 * x) / (#-1 * (y::hcomplex)) = z";
-
-*)
-
-
-(** Declarations for ExtractCommonTerm **)
-
-local
-  open HComplex_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         = Real_Numeral_Simprocs.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 =" hcomplexT
-  val simplify_meta_eq  = cancel_simplify_meta_eq field_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" hcomplexT
-  val simplify_meta_eq  = cancel_simplify_meta_eq mult_divide_cancel_eq_if
-);
-
-val hcomplex_cancel_factor = 
-  map prep_simproc
-   [("hcomplex_eq_cancel_factor", ["(l::hcomplex) * m = n", "(l::hcomplex) = m * n"], 
-     EqCancelFactor.proc),
-    ("hcomplex_divide_cancel_factor", ["((l::hcomplex) * m) / n", "(l::hcomplex) / (m * n)"], 
-     DivideCancelFactor.proc)];
-
-end;
-
-Addsimprocs hcomplex_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::hcomplex)";
-test "k = k*(y::hcomplex)"; 
-test "a*(b*c) = (b::hcomplex)";
-test "a*(b*c) = d*(b::hcomplex)*(x*a)";
-
-
-test "(x*k) / (k*(y::hcomplex)) = (uu::hcomplex)";
-test "(k) / (k*(y::hcomplex)) = (uu::hcomplex)"; 
-test "(a*(b*c)) / ((b::hcomplex)) = (uu::hcomplex)";
-test "(a*(b*c)) / (d*(b::hcomplex)*(x*a)) = (uu::hcomplex)";
-
-(*FIXME: what do we do about this?*)
-test "a*(b*c)/(y*z) = d*(b::hcomplex)*(x*a)/z";
-*)
-
-
--- a/src/HOL/Hyperreal/HSeries.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Hyperreal/HSeries.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -183,8 +183,7 @@
 
 Goalw [hypnat_omega_def,hypnat_zero_def,omega_def]  
      "sumhr(0, whn, %i. 1) = omega - 1";
-by (simp_tac (HOL_ss addsimps
-             [hypreal_numeral_1_eq_1, hypreal_one_def]) 1); 
+by (simp_tac (HOL_ss addsimps [numeral_1_eq_1, hypreal_one_def]) 1); 
 by (auto_tac (claset(),
               simpset() addsimps [sumhr, hypreal_diff, real_of_nat_Suc]));
 qed "sumhr_hypreal_omega_minus_one";
--- a/src/HOL/Hyperreal/HyperArith.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Hyperreal/HyperArith.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -9,132 +9,39 @@
 theory HyperArith = HyperDef
 files ("hypreal_arith.ML"):
 
-subsection{*Binary Arithmetic for the Hyperreals*}
+
+subsection{*Numerals and Arithmetic*}
 
 instance hypreal :: number ..
 
-defs (overloaded)
-  hypreal_number_of_def:
-    "number_of v == hypreal_of_real (number_of v)"
-     (*::bin=>hypreal               ::bin=>real*)
-     --{*This case is reduced to that for the reals.*}
-
-
-
-subsubsection{*Embedding the Reals into the Hyperreals*}
-
-lemma hypreal_number_of [simp]: "hypreal_of_real (number_of w) = number_of w"
-by (simp add: hypreal_number_of_def)
+primrec (*the type constraint is essential!*)
+  number_of_Pls: "number_of bin.Pls = 0"
+  number_of_Min: "number_of bin.Min = - (1::hypreal)"
+  number_of_BIT: "number_of(w BIT x) = (if x then 1 else 0) +
+	                               (number_of w) + (number_of w)"
 
-lemma hypreal_numeral_0_eq_0: "Numeral0 = (0::hypreal)"
-by (simp add: hypreal_number_of_def)
-
-lemma hypreal_numeral_1_eq_1: "Numeral1 = (1::hypreal)"
-by (simp add: hypreal_number_of_def)
-
-subsubsection{*Addition*}
-
-lemma add_hypreal_number_of [simp]:
-     "(number_of v :: hypreal) + number_of v' = number_of (bin_add v v')"
-by (simp only: hypreal_number_of_def hypreal_of_real_add [symmetric]
-               add_real_number_of)
-
+declare number_of_Pls [simp del]
+        number_of_Min [simp del]
+        number_of_BIT [simp del]
 
-subsubsection{*Subtraction*}
-
-lemma minus_hypreal_number_of [simp]:
-     "- (number_of w :: hypreal) = number_of (bin_minus w)"
-by (simp only: hypreal_number_of_def minus_real_number_of
-               hypreal_of_real_minus [symmetric])
-
-lemma diff_hypreal_number_of [simp]:
-     "(number_of v :: hypreal) - number_of w =
-      number_of (bin_add v (bin_minus w))"
-by (unfold hypreal_number_of_def hypreal_diff_def, simp)
-
-
-subsubsection{*Multiplication*}
-
-lemma mult_hypreal_number_of [simp]:
-     "(number_of v :: hypreal) * number_of v' = number_of (bin_mult v v')"
-by (simp only: hypreal_number_of_def hypreal_of_real_mult [symmetric] mult_real_number_of)
-
-text{*Lemmas for specialist use, NOT as default simprules*}
-lemma hypreal_mult_2: "2 * z = (z+z::hypreal)"
-proof -
-  have eq: "(2::hypreal) = 1 + 1"
-    by (simp add: hypreal_numeral_1_eq_1 [symmetric])
-  thus ?thesis by (simp add: eq left_distrib)
+instance hypreal :: number_ring
+proof
+  show "Numeral0 = (0::hypreal)" by (rule number_of_Pls)
+  show "-1 = - (1::hypreal)" by (rule number_of_Min)
+  fix w :: bin and x :: bool
+  show "(number_of (w BIT x) :: hypreal) =
+        (if x then 1 else 0) + number_of w + number_of w"
+    by (rule number_of_BIT)
 qed
 
-lemma hypreal_mult_2_right: "z * 2 = (z+z::hypreal)"
-by (subst hypreal_mult_commute, rule hypreal_mult_2)
-
 
-subsubsection{*Comparisons*}
-
-(** Equals (=) **)
-
-lemma eq_hypreal_number_of [simp]:
-     "((number_of v :: hypreal) = number_of v') =
-      iszero (number_of (bin_add v (bin_minus v')) :: int)"
-apply (simp only: hypreal_number_of_def hypreal_of_real_eq_iff eq_real_number_of)
+text{*Collapse applications of @{term hypreal_of_real} to @{term number_of}*}
+lemma hypreal_number_of [simp]: "hypreal_of_real (number_of w) = number_of w"
+apply (induct w) 
+apply (simp_all only: number_of hypreal_of_real_add hypreal_of_real_minus, simp_all) 
 done
 
 
-(** Less-than (<) **)
-
-(*"neg" is used in rewrite rules for binary comparisons*)
-lemma less_hypreal_number_of [simp]:
-     "((number_of v :: hypreal) < number_of v') =
-      neg (number_of (bin_add v (bin_minus v')) :: int)"
-by (simp only: hypreal_number_of_def hypreal_of_real_less_iff less_real_number_of)
-
-
-
-text{*New versions of existing theorems involving 0, 1*}
-
-lemma hypreal_minus_1_eq_m1 [simp]: "- 1 = (-1::hypreal)"
-by (simp add: hypreal_numeral_1_eq_1 [symmetric])
-
-lemma hypreal_mult_minus1 [simp]: "-1 * z = -(z::hypreal)"
-proof -
-  have  "-1 * z = (- 1) * z" by (simp add: hypreal_minus_1_eq_m1)
-  also have "... = - (1 * z)" by (simp only: minus_mult_left)
-  also have "... = -z" by simp
-  finally show ?thesis .
-qed
-
-lemma hypreal_mult_minus1_right [simp]: "(z::hypreal) * -1 = -z"
-by (subst hypreal_mult_commute, rule hypreal_mult_minus1)
-
-
-subsection{*Simplification of Arithmetic when Nested to the Right*}
-
-lemma hypreal_add_number_of_left [simp]:
-     "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::hypreal)"
-by (simp add: add_assoc [symmetric])
-
-lemma hypreal_mult_number_of_left [simp]:
-     "number_of v *(number_of w * z) = (number_of(bin_mult v w) * z::hypreal)"
-by (simp add: hypreal_mult_assoc [symmetric])
-
-lemma hypreal_add_number_of_diff1:
-    "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::hypreal)"
-by (simp add: hypreal_diff_def hypreal_add_number_of_left)
-
-lemma hypreal_add_number_of_diff2 [simp]:
-     "number_of v + (c - number_of w) =
-      number_of (bin_add v (bin_minus w)) + (c::hypreal)"
-apply (subst diff_hypreal_number_of [symmetric])
-apply (simp only: hypreal_diff_def add_ac)
-done
-
-
-declare hypreal_numeral_0_eq_0 [simp] hypreal_numeral_1_eq_1 [simp]
-
-
-
 use "hypreal_arith.ML"
 
 setup hypreal_arith_setup
@@ -143,15 +50,6 @@
 by arith
 
 
-subsubsection{*Division By @{term 1} and @{term "-1"}*}
-
-lemma hypreal_divide_minus1 [simp]: "x/-1 = -(x::hypreal)"
-by simp
-
-lemma hypreal_minus1_divide [simp]: "-1/(x::hypreal) = - (1/x)"
-by (simp add: hypreal_divide_def hypreal_minus_inverse)
-
-
 subsection{*The Function @{term hypreal_of_real}*}
 
 lemma number_of_less_hypreal_of_real_iff [simp]:
@@ -161,7 +59,7 @@
 done
 
 lemma number_of_le_hypreal_of_real_iff [simp]:
-     "(number_of w <= hypreal_of_real z) = (number_of w <= z)"
+     "(number_of w \<le> hypreal_of_real z) = (number_of w \<le> z)"
 apply (subst hypreal_of_real_le_iff [symmetric])
 apply (simp (no_asm))
 done
@@ -179,20 +77,13 @@
 done
 
 lemma hypreal_of_real_le_number_of_iff [simp]:
-     "(hypreal_of_real z <= number_of w) = (z <= number_of w)"
+     "(hypreal_of_real z \<le> number_of w) = (z \<le> number_of w)"
 apply (subst hypreal_of_real_le_iff [symmetric])
 apply (simp (no_asm))
 done
 
 subsection{*Absolute Value Function for the Hyperreals*}
 
-lemma hrabs_number_of [simp]:
-     "abs (number_of v :: hypreal) =
-        (if neg (number_of v :: int) then number_of (bin_minus v)
-         else number_of v)"
-by (simp add: hrabs_def)
-
-
 declare abs_mult [simp]
 
 lemma hrabs_add_less: "[| abs x < r; abs y < s |] ==> abs(x+y) < r + (s::hypreal)"
@@ -200,6 +91,7 @@
 apply (simp split add: split_if_asm)
 done
 
+text{*used once in NSA*}
 lemma hrabs_less_gt_zero: "abs x < r ==> (0::hypreal) < r"
 by (blast intro!: order_le_less_trans abs_ge_zero)
 
@@ -210,10 +102,6 @@
 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)
 
-
-(*----------------------------------------------------------
-    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)
@@ -301,9 +189,7 @@
 
 val hypreal_of_nat_def = thm"hypreal_of_nat_def";
 
-val hrabs_number_of = thm "hrabs_number_of";
 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_add_lemma_disj = thm "hrabs_add_lemma_disj";
 val hypreal_of_real_hrabs = thm "hypreal_of_real_hrabs";
--- a/src/HOL/Hyperreal/HyperDef.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Hyperreal/HyperDef.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -504,8 +504,7 @@
       "(Abs_hypreal(hyprel``{%n. X n}) \<le> Abs_hypreal(hyprel``{%n. Y n})) =  
        ({n. X n \<le> Y n} \<in> FreeUltrafilterNat)"
 apply (unfold hypreal_le_def)
-apply (auto intro!: lemma_hyprel_refl)
-apply (ultra)
+apply (auto intro!: lemma_hyprel_refl, ultra)
 done
 
 lemma hypreal_le_refl: "w \<le> (w::hypreal)"
@@ -517,21 +516,18 @@
 apply (rule eq_Abs_hypreal [of i])
 apply (rule eq_Abs_hypreal [of j])
 apply (rule eq_Abs_hypreal [of k])
-apply (simp add: hypreal_le) 
-apply ultra
+apply (simp add: hypreal_le, ultra)
 done
 
 lemma hypreal_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::hypreal)"
 apply (rule eq_Abs_hypreal [of z])
 apply (rule eq_Abs_hypreal [of w])
-apply (simp add: hypreal_le) 
-apply ultra
+apply (simp add: hypreal_le, ultra)
 done
 
 (* Axiom 'order_less_le' of class 'order': *)
 lemma hypreal_less_le: "((w::hypreal) < z) = (w \<le> z & w \<noteq> z)"
-apply (simp add: hypreal_less_def)
-done
+by (simp add: hypreal_less_def)
 
 instance hypreal :: order
 proof qed
@@ -543,8 +539,7 @@
 lemma hypreal_le_linear: "(z::hypreal) \<le> w | w \<le> z"
 apply (rule eq_Abs_hypreal [of z])
 apply (rule eq_Abs_hypreal [of w])
-apply (auto simp add: hypreal_le) 
-apply ultra
+apply (auto simp add: hypreal_le, ultra)
 done
 
 instance hypreal :: linorder 
@@ -565,8 +560,7 @@
 apply (rule eq_Abs_hypreal [of y])
 apply (rule eq_Abs_hypreal [of z])
 apply (auto simp add: hypreal_zero_def hypreal_le hypreal_mult 
-                      linorder_not_le [symmetric])
-apply ultra 
+                      linorder_not_le [symmetric], ultra) 
 done
 
 
@@ -590,7 +584,7 @@
   by (rule Ring_and_Field.mult_1_right)
 
 lemma hypreal_mult_minus_1 [simp]: "(- (1::hypreal)) * z = -z"
-by (simp)
+by simp
 
 lemma hypreal_mult_minus_1_right [simp]: "z * (- (1::hypreal)) = -z"
 by (subst hypreal_mult_commute, simp)
@@ -613,12 +607,10 @@
 by (auto dest: hypreal_eq_minus_iff [THEN iffD2])
 
 lemma hypreal_mult_left_cancel: "(c::hypreal) \<noteq> 0 ==> (c*a=c*b) = (a=b)"
-apply auto
-done
+by auto
     
 lemma hypreal_mult_right_cancel: "(c::hypreal) \<noteq> 0 ==> (a*c=b*c) = (a=b)"
-apply auto
-done
+by auto
 
 lemma hypreal_mult_not_0: "[| x \<noteq> 0; y \<noteq> 0 |] ==> x * y \<noteq> (0::hypreal)"
 by simp
@@ -725,8 +717,7 @@
 lemma hypreal_less: 
       "(Abs_hypreal(hyprel``{%n. X n}) < Abs_hypreal(hyprel``{%n. Y n})) =  
        ({n. X n < Y n} \<in> FreeUltrafilterNat)"
-apply (auto simp add: hypreal_le linorder_not_le [symmetric]) 
-apply ultra+
+apply (auto simp add: hypreal_le linorder_not_le [symmetric], ultra+)
 done
 
 lemma hypreal_zero_num: "0 = Abs_hypreal (hyprel `` {%n. 0})"
@@ -773,7 +764,7 @@
 
 lemma lemma_omega_empty_singleton_disj: "{n::nat. x = real n} = {} |  
       (\<exists>y. {n::nat. x = real n} = {y})"
-by (force)
+by force
 
 lemma lemma_finite_omega_set: "finite {n::nat. x = real n}"
 by (cut_tac x = x in lemma_omega_empty_singleton_disj, auto)
@@ -794,7 +785,7 @@
 lemma lemma_epsilon_empty_singleton_disj:
      "{n::nat. x = inverse(real(Suc n))} = {} |  
       (\<exists>y. {n::nat. x = inverse(real(Suc n))} = {y})"
-by (auto)
+by auto
 
 lemma lemma_finite_epsilon_set: "finite {n. x = inverse(real(Suc n))}"
 by (cut_tac x = x in lemma_epsilon_empty_singleton_disj, auto)
--- a/src/HOL/Hyperreal/HyperPow.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Hyperreal/HyperPow.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -2,6 +2,7 @@
     Author      : Jacques D. Fleuriot  
     Copyright   : 1998  University of Cambridge
     Description : Powers theory for hyperreals
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
 *)
 
 header{*Exponentials on the Hyperreals*}
@@ -123,7 +124,7 @@
 
 lemma power_hypreal_of_real_number_of:
      "(number_of v :: hypreal) ^ n = hypreal_of_real ((number_of v) ^ n)"
-by (simp only: hypreal_number_of_def hypreal_of_real_power)
+by (simp only: hypreal_number_of [symmetric] hypreal_of_real_power)
 
 declare power_hypreal_of_real_number_of [of _ "number_of w", standard, simp]
 
--- a/src/HOL/Hyperreal/IntFloor.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Hyperreal/IntFloor.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -4,6 +4,8 @@
     Description: Floor and ceiling operations over reals
 *)
 
+val real_number_of = thm"real_number_of";
+
 Goal "((number_of n) < real (m::int)) = (number_of n < m)";
 by Auto_tac;
 by (rtac (real_of_int_less_iff RS iffD1) 1);
@@ -187,8 +189,8 @@
 by (auto_tac (claset() addIs [floor_eq3],simpset()));
 qed "floor_eq4";
 
-Goalw [real_number_of_def] 
-       "floor(number_of n :: real) = (number_of n :: int)";
+Goal "floor(number_of n :: real) = (number_of n :: int)";
+by (stac (real_number_of RS sym) 1);
 by (rtac floor_eq2 1);
 by (rtac (CLAIM "x < x + (1::real)") 2);
 by (rtac real_le_refl 1);
@@ -279,8 +281,9 @@
 by (auto_tac (claset() addIs [floor_eq2],simpset()));
 qed "ceiling_eq3";
 
-Goalw [ceiling_def,real_number_of_def] 
+Goalw [ceiling_def] 
   "ceiling (number_of n :: real) = (number_of n :: int)";
+by (stac (real_number_of RS sym) 1);
 by (rtac (CLAIM "x = - y ==> - (x::int) = y") 1);
 by (rtac (floor_minus_real_of_int RS ssubst) 1);
 by Auto_tac;
--- a/src/HOL/Hyperreal/Integration.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Hyperreal/Integration.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -4,6 +4,9 @@
     Description : Theory of integration (c.f. Harison's HOL development)
 *)
 
+val mult_2 = thm"mult_2";
+val mult_2_right = thm"mult_2_right";
+
 Goalw [psize_def] "a = b ==> psize (%n. if n = 0 then a else b) = 0";
 by Auto_tac;
 qed "partition_zero";
@@ -358,7 +361,7 @@
 by (dtac add_strict_mono 1 THEN assume_tac 1);
 by (auto_tac (claset(),
     HOL_ss addsimps [left_distrib RS sym,
-                     real_mult_2_right RS sym, mult_less_cancel_right]));
+                     mult_2_right RS sym, mult_less_cancel_right]));
 by (ALLGOALS(arith_tac));
 qed "Integral_unique";
 
@@ -956,7 +959,7 @@
     ("c","abs (rsum (D, p) g - k2) * 2")] 
     add_strict_mono 1 THEN assume_tac 1);
 by (auto_tac (claset(),HOL_ss addsimps [rsum_add,left_distrib RS sym,
-    real_mult_2_right RS sym,real_mult_less_iff1,CLAIM "(0::real) < 2"]));
+    mult_2_right RS sym,real_mult_less_iff1,CLAIM "(0::real) < 2"]));
 by (arith_tac 1);
 qed "Integral_add_fun";
 
@@ -1015,7 +1018,7 @@
 by (arith_tac 1);
 by (dtac add_strict_mono 1 THEN assume_tac 1);
 by (auto_tac (claset(),HOL_ss addsimps [left_distrib RS sym,
-    real_mult_2_right RS sym,real_mult_less_iff1,CLAIM "(0::real) < 2"]));
+    mult_2_right RS sym,real_mult_less_iff1,CLAIM "(0::real) < 2"]));
 by (arith_tac 1);
 qed "Integral_le";
 
@@ -1037,7 +1040,7 @@
 by (thin_tac "0 < e" 1);
 by (dtac add_strict_mono 1 THEN assume_tac 1);
 by (auto_tac (claset(),HOL_ss addsimps [left_distrib RS sym,
-    real_mult_2_right RS sym,real_mult_less_iff1,CLAIM "(0::real) < 2"]));
+    mult_2_right RS sym,real_mult_less_iff1,CLAIM "(0::real) < 2"]));
 by (arith_tac 1);
 qed "Integral_imp_Cauchy";
 
--- a/src/HOL/Hyperreal/Lim.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Hyperreal/Lim.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -1405,7 +1405,7 @@
 Goal "(\\<forall>z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l \
 \     ==> NSDERIV f x :> l";
 by (auto_tac (claset(), 
-              simpset() delsimprocs real_cancel_factor
+              simpset() delsimprocs field_cancel_factor
                         addsimps [NSDERIV_iff2]));
 by (auto_tac (claset(),
               simpset() addsimps [hypreal_mult_assoc]));
@@ -1877,9 +1877,25 @@
               addsplits [split_if_asm]) 1); 
 qed "DERIV_left_inc";
 
-Goalw [deriv_def,LIM_def] 
+val prems = goalw (the_context()) [deriv_def,LIM_def]
     "[| DERIV f x :> l;  l < 0 |] ==> \
 \      \\<exists>d. 0 < d & (\\<forall>h. 0 < h & h < d --> f(x) < f(x - h))";
+by (cut_facts_tac prems 1);  (*needed because arith removes the assumption l<0*)
+by (dres_inst_tac [("x","-l")] spec 1 THEN Auto_tac);
+by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
+by (dres_inst_tac [("x","-h")] spec 1);
+by (asm_full_simp_tac
+    (simpset() addsimps [real_abs_def, inverse_eq_divide, 
+                         pos_less_divide_eq,
+                         symmetric real_diff_def]
+               addsplits [split_if_asm]) 1);
+by (subgoal_tac "0 < (f (x - h) - f x)/h" 1);
+by (asm_full_simp_tac (simpset() addsimps [pos_less_divide_eq]) 1); 
+by (cut_facts_tac prems 1);
+by (arith_tac 1);
+qed "DERIV_left_dec";
+
+(*????previous proof, revealing arith problem:
 by (dres_inst_tac [("x","-l")] spec 1 THEN Auto_tac);
 by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
 by (subgoal_tac "l*h < 0" 1);
@@ -1896,6 +1912,7 @@
 by (asm_full_simp_tac
     (simpset() addsimps [pos_less_divide_eq]) 1); 
 qed "DERIV_left_dec";
+*)
 
 
 Goal "[| DERIV f x :> l; \
--- a/src/HOL/Hyperreal/Lim.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Hyperreal/Lim.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -5,7 +5,7 @@
                   differentiation of real=>real functions
 *)
 
-Lim = SEQ + RealArith + 
+Lim = SEQ + RealDef + 
 
 (*-----------------------------------------------------------------------
     Limits, continuity and differentiation: standard and NS definitions
--- a/src/HOL/Hyperreal/NSA.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Hyperreal/NSA.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -95,7 +95,7 @@
 declare SReal_hypreal_of_real [simp]
 
 lemma SReal_number_of: "(number_of w ::hypreal) \<in> Reals"
-apply (unfold hypreal_number_of_def)
+apply (simp only: hypreal_number_of [symmetric])
 apply (rule SReal_hypreal_of_real)
 done
 declare SReal_number_of [simp]
@@ -103,13 +103,13 @@
 (** As always with numerals, 0 and 1 are special cases **)
 
 lemma Reals_0: "(0::hypreal) \<in> Reals"
-apply (subst hypreal_numeral_0_eq_0 [symmetric])
+apply (subst numeral_0_eq_0 [symmetric])
 apply (rule SReal_number_of)
 done
 declare Reals_0 [simp]
 
 lemma Reals_1: "(1::hypreal) \<in> Reals"
-apply (subst hypreal_numeral_1_eq_1 [symmetric])
+apply (subst numeral_1_eq_1 [symmetric])
 apply (rule SReal_number_of)
 done
 declare Reals_1 [simp]
@@ -267,13 +267,13 @@
 (** As always with numerals, 0 and 1 are special cases **)
 
 lemma HFinite_0: "0 \<in> HFinite"
-apply (subst hypreal_numeral_0_eq_0 [symmetric])
+apply (subst numeral_0_eq_0 [symmetric])
 apply (rule HFinite_number_of)
 done
 declare HFinite_0 [simp]
 
 lemma HFinite_1: "1 \<in> HFinite"
-apply (subst hypreal_numeral_1_eq_1 [symmetric])
+apply (subst numeral_1_eq_1 [symmetric])
 apply (rule HFinite_number_of)
 done
 declare HFinite_1 [simp]
@@ -859,7 +859,7 @@
 
 (*again: 1 is a special case, but not 0 this time*)
 lemma one_not_Infinitesimal: "1 \<notin> Infinitesimal"
-apply (subst hypreal_numeral_1_eq_1 [symmetric])
+apply (subst numeral_1_eq_1 [symmetric])
 apply (rule number_of_not_Infinitesimal)
 apply (simp (no_asm))
 done
@@ -1424,12 +1424,14 @@
      "[| x < y;  u \<in> Infinitesimal |]
       ==> hypreal_of_real x + u < hypreal_of_real y"
 apply (simp add: Infinitesimal_def)
-apply (drule hypreal_of_real_less_iff [THEN iffD2])
 apply (drule_tac x = "hypreal_of_real y + -hypreal_of_real x" in bspec)
-apply (auto simp add: hypreal_add_commute abs_less_iff SReal_add SReal_minus)
+ apply (simp add: );  
+apply (auto simp add: add_commute abs_less_iff SReal_add SReal_minus)
+apply (simp add: compare_rls) 
 done
 
-lemma Infinitesimal_add_hrabs_hypreal_of_real_less: "[| x \<in> Infinitesimal; abs(hypreal_of_real r) < hypreal_of_real y |]
+lemma Infinitesimal_add_hrabs_hypreal_of_real_less:
+     "[| x \<in> Infinitesimal; abs(hypreal_of_real r) < hypreal_of_real y |]
       ==> abs (hypreal_of_real r + x) < hypreal_of_real y"
 apply (drule_tac x = "hypreal_of_real r" in approx_hrabs_add_Infinitesimal)
 apply (drule approx_sym [THEN bex_Infinitesimal_iff2 [THEN iffD2]])
@@ -1692,7 +1694,7 @@
 done
 
 lemma st_Infinitesimal: "x \<in> Infinitesimal ==> st x = 0"
-apply (subst hypreal_numeral_0_eq_0 [symmetric])
+apply (subst numeral_0_eq_0 [symmetric])
 apply (rule st_number_of [THEN subst])
 apply (rule approx_st_eq)
 apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mem_infmal_iff [symmetric])
@@ -1743,13 +1745,13 @@
 done
 
 lemma st_zero_le: "[| 0 <= x;  x \<in> HFinite |] ==> 0 <= st x"
-apply (subst hypreal_numeral_0_eq_0 [symmetric])
+apply (subst numeral_0_eq_0 [symmetric])
 apply (rule st_number_of [THEN subst])
 apply (rule st_le, auto)
 done
 
 lemma st_zero_ge: "[| x <= 0;  x \<in> HFinite |] ==> st x <= 0"
-apply (subst hypreal_numeral_0_eq_0 [symmetric])
+apply (subst numeral_0_eq_0 [symmetric])
 apply (rule st_number_of [THEN subst])
 apply (rule st_le, auto)
 done
--- a/src/HOL/Hyperreal/SEQ.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Hyperreal/SEQ.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -398,6 +398,7 @@
 Goal "(EX K. 0 < K & (ALL n. abs(X n) <= K)) = \
 \     (EX N. ALL n. abs(X n) <= real(Suc N))";
 by Auto_tac;
+by (Force_tac 2);
 by (cut_inst_tac [("x","K")] reals_Archimedean2 1);
 by (Clarify_tac 1); 
 by (res_inst_tac [("x","n")] exI 1); 
--- a/src/HOL/Hyperreal/Transcendental.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Hyperreal/Transcendental.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -75,8 +75,8 @@
 
 (*lcp: needed now because 2 is a binary numeral!*)
 Goal "root 2 = root (Suc (Suc 0))";
-by (simp_tac (simpset() delsimps [numeral_0_eq_0, numeral_1_eq_1]
-	                addsimps [numeral_0_eq_0 RS sym]) 1);  
+by (simp_tac (simpset() delsimps [nat_numeral_0_eq_0, nat_numeral_1_eq_1]
+	                addsimps [nat_numeral_0_eq_0 RS sym]) 1);  
 qed "root_2_eq";
 Addsimps [root_2_eq];
 
@@ -2077,7 +2077,7 @@
 Goal "[| cos x ~= 0; cos (2 * x) ~= 0 |] \
 \     ==> tan (2 * x) = (2 * tan x)/(1 - (tan(x) ^ 2))";
 by (auto_tac (claset(),simpset() addsimps [asm_full_simplify 
-    (simpset() addsimps [real_mult_2 RS sym] delsimps [lemma_tan_add1]) 
+    (simpset() addsimps [thm"mult_2" RS sym] delsimps [lemma_tan_add1]) 
     (read_instantiate [("x","x"),("y","x")] tan_add),numeral_2_eq_2]
     delsimps [lemma_tan_add1]));
 qed "tan_double";
--- a/src/HOL/Hyperreal/hypreal_arith.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Hyperreal/hypreal_arith.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -15,506 +15,11 @@
     read_instantiate_sg(sign_of (the_context())) [("a","?a::hypreal")] mult_left_mono;
 
 
-val hypreal_number_of = thm "hypreal_number_of";
-val hypreal_numeral_0_eq_0 = thm "hypreal_numeral_0_eq_0";
-val hypreal_numeral_1_eq_1 = thm "hypreal_numeral_1_eq_1";
-val hypreal_number_of_def = thm "hypreal_number_of_def";
-val add_hypreal_number_of = thm "add_hypreal_number_of";
-val minus_hypreal_number_of = thm "minus_hypreal_number_of";
-val diff_hypreal_number_of = thm "diff_hypreal_number_of";
-val mult_hypreal_number_of = thm "mult_hypreal_number_of";
-val hypreal_mult_2 = thm "hypreal_mult_2";
-val hypreal_mult_2_right = thm "hypreal_mult_2_right";
-val eq_hypreal_number_of = thm "eq_hypreal_number_of";
-val less_hypreal_number_of = thm "less_hypreal_number_of";
-val hypreal_minus_1_eq_m1 = thm "hypreal_minus_1_eq_m1";
-val hypreal_mult_minus1 = thm "hypreal_mult_minus1";
-val hypreal_mult_minus1_right = thm "hypreal_mult_minus1_right";
-val hypreal_add_number_of_left = thm "hypreal_add_number_of_left";
-val hypreal_mult_number_of_left = thm "hypreal_mult_number_of_left";
-val hypreal_add_number_of_diff1 = thm "hypreal_add_number_of_diff1";
-val hypreal_add_number_of_diff2 = thm "hypreal_add_number_of_diff2";
-
-(*Maps 0 to Numeral0 and 1 to Numeral1 and -(Numeral1) to -1*)
-val hypreal_numeral_ss =
-    real_numeral_ss addsimps [hypreal_numeral_0_eq_0 RS sym,
-                              hypreal_numeral_1_eq_1 RS sym,
-                              hypreal_minus_1_eq_m1]
-
-fun rename_numerals th =
-    asm_full_simplify hypreal_numeral_ss (Thm.transfer (the_context ()) th)
-
-
-structure Hyperreal_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 = [hypreal_numeral_0_eq_0 RS sym,
-                    hypreal_numeral_1_eq_1 RS sym]
-
-(*Utilities*)
-
-val hyprealT = Type("HyperDef.hypreal",[])
-
-fun mk_numeral n = HOLogic.number_of_const hyprealT $ HOLogic.mk_bin n
-
-val dest_numeral = Real_Numeral_Simprocs.dest_numeral
-
-val find_first_numeral = Real_Numeral_Simprocs.find_first_numeral
-
-val zero = mk_numeral 0
-val mk_plus = Real_Numeral_Simprocs.mk_plus
-
-val uminus_const = Const ("uminus", hyprealT --> hyprealT)
-
-(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
-fun mk_sum []        = zero
-  | mk_sum [t,u]     = mk_plus (t, u)
-  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts)
-
-(*this version ALWAYS includes a trailing zero*)
-fun long_mk_sum []        = zero
-  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts)
-
-val dest_plus = HOLogic.dest_bin "op +" hyprealT
-
-(*decompose additions AND subtractions as a sum*)
-fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
-        dest_summing (pos, t, dest_summing (pos, u, ts))
-  | dest_summing (pos, Const ("op -", _) $ t $ u, ts) =
-        dest_summing (pos, t, dest_summing (not pos, u, ts))
-  | dest_summing (pos, t, ts) =
-        if pos then t::ts else uminus_const$t :: ts
-
-fun dest_sum t = dest_summing (true, t, [])
-
-val mk_diff = HOLogic.mk_binop "op -"
-val dest_diff = HOLogic.dest_bin "op -" hyprealT
-
-val one = mk_numeral 1
-val mk_times = HOLogic.mk_binop "op *"
-
-fun mk_prod [] = one
-  | mk_prod [t] = t
-  | mk_prod (t :: ts) = if t = one then mk_prod ts
-                        else mk_times (t, mk_prod ts)
-
-val dest_times = HOLogic.dest_bin "op *" hyprealT
-
-fun dest_prod t =
-      let val (t,u) = dest_times t
-      in  dest_prod t @ dest_prod u  end
-      handle TERM _ => [t]
-
-(*DON'T do the obvious simplifications; that would create special cases*)
-fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts)
-
-(*Express t as a product of (possibly) a numeral with other sorted terms*)
-fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t
-  | dest_coeff sign t =
-    let val ts = sort Term.term_ord (dest_prod t)
-        val (n, ts') = find_first_numeral [] ts
-                          handle TERM _ => (1, ts)
-    in (sign*n, mk_prod ts') end
-
-(*Find first coefficient-term THAT MATCHES u*)
-fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
-  | find_first_coeff past u (t::terms) =
-        let val (n,u') = dest_coeff 1 t
-        in  if u aconv u' then (n, rev past @ terms)
-                          else find_first_coeff (t::past) u terms
-        end
-        handle TERM _ => find_first_coeff (t::past) u terms
-
-
-(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
-val add_0s = map rename_numerals
-                 [hypreal_add_zero_left, hypreal_add_zero_right]
-val mult_1s = map rename_numerals [hypreal_mult_1, hypreal_mult_1_right] @
-              [hypreal_mult_minus1, hypreal_mult_minus1_right]
-
-(*To perform binary arithmetic*)
-val bin_simps =
-    [hypreal_numeral_0_eq_0 RS sym, hypreal_numeral_1_eq_1 RS sym,
-     add_hypreal_number_of, hypreal_add_number_of_left,
-     minus_hypreal_number_of,
-     diff_hypreal_number_of, mult_hypreal_number_of,
-     hypreal_mult_number_of_left] @ bin_arith_simps @ bin_rel_simps
-
-(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
-  during re-arrangement*)
-val non_add_bin_simps = 
-    bin_simps \\ [hypreal_add_number_of_left, add_hypreal_number_of]
-
-(*To evaluate binary negations of coefficients*)
-val hypreal_minus_simps = NCons_simps @
-                   [hypreal_minus_1_eq_m1, minus_hypreal_number_of,
-                    bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min,
-                    bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min]
-
-(*To let us treat subtraction as addition*)
-val diff_simps = [hypreal_diff_def, minus_add_distrib, minus_minus]
-
-(*push the unary minus down: - x * y = x * - y *)
-val hypreal_minus_mult_eq_1_to_2 =
-    [minus_mult_left RS sym, minus_mult_right] MRS trans
-    |> standard
-
-(*to extract again any uncancelled minuses*)
-val hypreal_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 hypreal_mult_minus_simps =
-    [hypreal_mult_assoc, minus_mult_left, hypreal_minus_mult_eq_1_to_2]
-
-(*Final simplification: cancel + and *  *)
-val simplify_meta_eq =
-    Int_Numeral_Simprocs.simplify_meta_eq
-         [hypreal_add_zero_left, hypreal_add_zero_right,
-          mult_zero_left, mult_zero_right, mult_1, mult_1_right]
-
-val prep_simproc = Real_Numeral_Simprocs.prep_simproc
-
-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         = Real_Numeral_Simprocs.trans_tac
-  val norm_tac =
-     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
-                                         hypreal_minus_simps@add_ac))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@hypreal_mult_minus_simps))
-     THEN ALLGOALS
-              (simp_tac (HOL_ss addsimps hypreal_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 =" hyprealT
-  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 <" hyprealT
-  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 <=" hyprealT
-  val bal_add1 = le_add_iff1 RS trans
-  val bal_add2 = le_add_iff2 RS trans
-)
-
-val cancel_numerals =
-  map prep_simproc
-   [("hyprealeq_cancel_numerals",
-     ["(l::hypreal) + m = n", "(l::hypreal) = m + n",
-      "(l::hypreal) - m = n", "(l::hypreal) = m - n",
-      "(l::hypreal) * m = n", "(l::hypreal) = m * n"],
-     EqCancelNumerals.proc),
-    ("hyprealless_cancel_numerals",
-     ["(l::hypreal) + m < n", "(l::hypreal) < m + n",
-      "(l::hypreal) - m < n", "(l::hypreal) < m - n",
-      "(l::hypreal) * m < n", "(l::hypreal) < m * n"],
-     LessCancelNumerals.proc),
-    ("hyprealle_cancel_numerals",
-     ["(l::hypreal) + m <= n", "(l::hypreal) <= m + n",
-      "(l::hypreal) - m <= n", "(l::hypreal) <= m - n",
-      "(l::hypreal) * m <= n", "(l::hypreal) <= m * n"],
-     LeCancelNumerals.proc)]
-
-
-structure CombineNumeralsData =
-  struct
-  val add               = op + : int*int -> int
-  val mk_sum            = long_mk_sum    (*to work for e.g. 2*x + 3*x *)
-  val dest_sum          = dest_sum
-  val mk_coeff          = mk_coeff
-  val dest_coeff        = dest_coeff 1
-  val left_distrib      = combine_common_factor RS trans
-  val prove_conv        = Bin_Simprocs.prove_conv_nohyps
-  val trans_tac         = Real_Numeral_Simprocs.trans_tac
-  val norm_tac =
-     ALLGOALS (simp_tac (HOL_ss addsimps numeral_syms@add_0s@mult_1s@
-                                   diff_simps@hypreal_minus_simps@add_ac))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@hypreal_mult_minus_simps))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps hypreal_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 CombineNumerals = CombineNumeralsFun(CombineNumeralsData)
-
-val combine_numerals =
-  prep_simproc
-    ("hypreal_combine_numerals", ["(i::hypreal) + j", "(i::hypreal) - j"], CombineNumerals.proc)
-
-
-(** Declarations for ExtractCommonTerm **)
-
-(*this version ALWAYS includes a trailing one*)
-fun long_mk_prod []        = one
-  | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts)
-
-(*Find first term that matches u*)
-fun find_first past u []         = raise TERM("find_first", [])
-  | find_first past u (t::terms) =
-        if u aconv t then (rev past @ terms)
-        else find_first (t::past) u terms
-        handle TERM _ => find_first (t::past) u terms
-
-(*Final simplification: cancel + and *  *)
-fun cancel_simplify_meta_eq cancel_th th =
-    Int_Numeral_Simprocs.simplify_meta_eq
-        [mult_1, mult_1_right]
-        (([th, cancel_th]) MRS trans)
-
-(*** Making constant folding work for 0 and 1 too ***)
-
-structure HyperrealAbstractNumeralsData =
-  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  = hypreal_numeral_0_eq_0
-  val numeral_1_eq_1  = hypreal_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 HyperrealAbstractNumerals =
-  AbstractNumeralsFun (HyperrealAbstractNumeralsData)
-
-(*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
-   [("hypreal_add_eval_numerals",
-     ["(m::hypreal) + 1", "(m::hypreal) + number_of v"],
-     HyperrealAbstractNumerals.proc add_hypreal_number_of),
-    ("hypreal_diff_eval_numerals",
-     ["(m::hypreal) - 1", "(m::hypreal) - number_of v"],
-     HyperrealAbstractNumerals.proc diff_hypreal_number_of),
-    ("hypreal_eq_eval_numerals",
-     ["(m::hypreal) = 0", "(m::hypreal) = 1", "(m::hypreal) = number_of v"],
-     HyperrealAbstractNumerals.proc eq_hypreal_number_of),
-    ("hypreal_less_eval_numerals",
-     ["(m::hypreal) < 0", "(m::hypreal) < 1", "(m::hypreal) < number_of v"],
-     HyperrealAbstractNumerals.proc less_hypreal_number_of),
-    ("hypreal_le_eval_numerals",
-     ["(m::hypreal) <= 0", "(m::hypreal) <= 1", "(m::hypreal) <= number_of v"],
-     HyperrealAbstractNumerals.proc le_number_of_eq_not_less)]
-
-end;
-
-Addsimprocs Hyperreal_Numeral_Simprocs.eval_numerals;
-Addsimprocs Hyperreal_Numeral_Simprocs.cancel_numerals;
-Addsimprocs [Hyperreal_Numeral_Simprocs.combine_numerals];
-
-
-
-
-(**** Constant folding for hypreal plus and times ****)
-
-(*We do not need
-    structure Hyperreal_Plus_Assoc = Assoc_Fold (Hyperreal_Plus_Assoc_Data)
-  because combine_numerals does the same thing*)
-
-structure Hyperreal_Times_Assoc_Data : ASSOC_FOLD_DATA =
-struct
-  val ss                = HOL_ss
-  val eq_reflection     = eq_reflection
-  val sg_ref    = Sign.self_ref (Theory.sign_of (the_context ()))
-  val T      = Hyperreal_Numeral_Simprocs.hyprealT
-  val plus   = Const ("op *", [T,T] ---> T)
-  val add_ac = mult_ac
-end;
-
-structure Hyperreal_Times_Assoc = Assoc_Fold (Hyperreal_Times_Assoc_Data);
-
-Addsimprocs [Hyperreal_Times_Assoc.conv];
-
-
-
-(**** Simprocs for numeric literals ****)
-
-
-(****Common factor cancellation****)
-
-local
-  open Hyperreal_Numeral_Simprocs
-in
-
-val rel_hypreal_number_of = [eq_hypreal_number_of, less_hypreal_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         = Real_Numeral_Simprocs.trans_tac
-  val norm_tac =
-     ALLGOALS (simp_tac (HOL_ss addsimps hypreal_minus_from_mult_simps @ mult_1s))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@hypreal_mult_minus_simps))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps mult_ac))
-  val numeral_simp_tac  =
-         ALLGOALS (simp_tac (HOL_ss addsimps rel_hypreal_number_of@bin_simps))
-  val simplify_meta_eq  = simplify_meta_eq
-  end
-
-structure DivCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Bin_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binop "HOL.divide"
-  val dest_bal = HOLogic.dest_bin "HOL.divide" hyprealT
-  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 =" hyprealT
-  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 <" hyprealT
-  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 <=" hyprealT
-  val cancel = mult_le_cancel_left RS trans
-  val neg_exchanges = true
-)
-
-val hypreal_cancel_numeral_factors_relations =
-  map prep_simproc
-   [("hyprealeq_cancel_numeral_factor",
-     ["(l::hypreal) * m = n", "(l::hypreal) = m * n"],
-     EqCancelNumeralFactor.proc),
-    ("hyprealless_cancel_numeral_factor",
-     ["(l::hypreal) * m < n", "(l::hypreal) < m * n"],
-     LessCancelNumeralFactor.proc),
-    ("hyprealle_cancel_numeral_factor",
-     ["(l::hypreal) * m <= n", "(l::hypreal) <= m * n"],
-     LeCancelNumeralFactor.proc)];
-
-val hypreal_cancel_numeral_factors_divide = prep_simproc
-        ("hyprealdiv_cancel_numeral_factor",
-         ["((l::hypreal) * m) / n", "(l::hypreal) / (m * n)",
-          "((number_of v)::hypreal) / (number_of w)"],
-         DivCancelNumeralFactor.proc);
-
-val hypreal_cancel_numeral_factors =
-    hypreal_cancel_numeral_factors_relations @
-    [hypreal_cancel_numeral_factors_divide];
-
-end;
-
-Addsimprocs hypreal_cancel_numeral_factors;
-
-
-
-(** Declarations for ExtractCommonTerm **)
-
-local
-  open Hyperreal_Numeral_Simprocs
-in
-
-structure CancelFactorCommon =
-  struct
-  val mk_sum            = long_mk_prod
-  val dest_sum          = dest_prod
-  val mk_coeff          = mk_coeff
-  val dest_coeff        = dest_coeff
-  val find_first        = find_first []
-  val trans_tac         = Real_Numeral_Simprocs.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 =" hyprealT
-  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" hyprealT
-  val simplify_meta_eq  = cancel_simplify_meta_eq mult_divide_cancel_eq_if
-);
-
-val hypreal_cancel_factor =
-  map prep_simproc
-   [("hypreal_eq_cancel_factor", ["(l::hypreal) * m = n", "(l::hypreal) = m * n"],
-     EqCancelFactor.proc),
-    ("hypreal_divide_cancel_factor", ["((l::hypreal) * m) / n", "(l::hypreal) / (m * n)"],
-     DivideCancelFactor.proc)];
-
-end;
-
-Addsimprocs hypreal_cancel_factor;
-
-
-
-
 (****Instantiation of the generic linear arithmetic package****)
 
 
 local
 
-(* reduce contradictory <= to False *)
-val add_rules =
-    [hypreal_numeral_0_eq_0, hypreal_numeral_1_eq_1,
-     add_hypreal_number_of, minus_hypreal_number_of, diff_hypreal_number_of,
-     mult_hypreal_number_of, eq_hypreal_number_of, less_hypreal_number_of];
-
-val simprocs = [Hyperreal_Times_Assoc.conv, 
-                Hyperreal_Numeral_Simprocs.combine_numerals,
-                hypreal_cancel_numeral_factors_divide]@
-               Hyperreal_Numeral_Simprocs.cancel_numerals @
-               Hyperreal_Numeral_Simprocs.eval_numerals;
-
 fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
 
 val hypreal_mult_mono_thms =
@@ -529,6 +34,8 @@
 
 in
 
+val hyprealT = Type("Rational.hypreal", []);
+
 val fast_hypreal_arith_simproc =
     Simplifier.simproc (Theory.sign_of (the_context ()))
       "fast_hypreal_arith" 
@@ -541,10 +48,8 @@
     mult_mono_thms = mult_mono_thms @ hypreal_mult_mono_thms,
     inj_thms = inj_thms @ real_inj_thms, 
     lessD = lessD,  (*Can't change LA_Data_Ref.lessD: the hypreals are dense!*)
-    simpset = simpset addsimps add_rules
-                      addsimprocs simprocs}),
-  arith_inj_const ("HyperDef.hypreal_of_real", 
-                   HOLogic.realT --> Hyperreal_Numeral_Simprocs.hyprealT),
+    simpset = simpset}),
+  arith_inj_const ("HyperDef.hypreal_of_real", HOLogic.realT --> hyprealT),
   arith_discrete ("HyperDef.hypreal",false),
   Simplifier.change_simpset_of (op addsimprocs) [fast_hypreal_arith_simproc]];
 
--- a/src/HOL/Integ/Bin.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Integ/Bin.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -10,332 +10,266 @@
 
 theory Bin = IntDef + Numeral:
 
-text{*The sign @{term Pls} stands for an infinite string of leading Falses.*}
-text{*The sign @{term Min} stands for an infinite string of leading Trues.*}
-
-text{*A number can have multiple representations, namely leading Falses with
-sign @{term Pls} and leading Trues with sign @{term Min}.
-See @{text "ZF/Integ/twos-compl.ML"}, function @{text int_of_binary},
-for the numerical interpretation.
-
-The representation expects that @{term "(m mod 2)"} is 0 or 1,
-even if m is negative;
-For instance, @{term "-5 div 2 = -3"} and @{term "-5 mod 2 = 1"}; thus
-@{term "-5 = (-3)*2 + 1"}.
-*}
-
-consts
-  NCons     :: "[bin,bool]=>bin"
-  bin_succ  :: "bin=>bin"
-  bin_pred  :: "bin=>bin"
-  bin_minus :: "bin=>bin"
-  bin_add   :: "[bin,bin]=>bin"
-  bin_mult  :: "[bin,bin]=>bin"
-
-(*NCons inserts a bit, suppressing leading 0s and 1s*)
-primrec
-  NCons_Pls:  "NCons bin.Pls b = (if b then (bin.Pls BIT b) else bin.Pls)"
-  NCons_Min:  "NCons bin.Min b = (if b then bin.Min else (bin.Min BIT b))"
-  NCons_BIT:  "NCons (w BIT x) b = (w BIT x) BIT b"
-
-instance
-  int :: number ..
-
-primrec (*the type constraint is essential!*)
+axclass number_ring \<subseteq> number, ring
   number_of_Pls: "number_of bin.Pls = 0"
-  number_of_Min: "number_of bin.Min = - (1::int)"
+  number_of_Min: "number_of bin.Min = - 1"
   number_of_BIT: "number_of(w BIT x) = (if x then 1 else 0) +
 	                               (number_of w) + (number_of w)"
+subsection{*Converting Numerals to Rings: @{term number_of}*}
 
-primrec
-  bin_succ_Pls: "bin_succ bin.Pls = bin.Pls BIT True"
-  bin_succ_Min: "bin_succ bin.Min = bin.Pls"
-  bin_succ_BIT: "bin_succ(w BIT x) =
-  	            (if x then bin_succ w BIT False
-	                  else NCons w True)"
-
-primrec
-  bin_pred_Pls: "bin_pred bin.Pls = bin.Min"
-  bin_pred_Min: "bin_pred bin.Min = bin.Min BIT False"
-  bin_pred_BIT: "bin_pred(w BIT x) =
-	            (if x then NCons w False
-		          else (bin_pred w) BIT True)"
+lemmas number_of = number_of_Pls number_of_Min number_of_BIT
 
-primrec
-  bin_minus_Pls: "bin_minus bin.Pls = bin.Pls"
-  bin_minus_Min: "bin_minus bin.Min = bin.Pls BIT True"
-  bin_minus_BIT: "bin_minus(w BIT x) =
-	             (if x then bin_pred (NCons (bin_minus w) False)
-		           else bin_minus w BIT False)"
-
-primrec
-  bin_add_Pls: "bin_add bin.Pls w = w"
-  bin_add_Min: "bin_add bin.Min w = bin_pred w"
-  bin_add_BIT:
-    "bin_add (v BIT x) w =
-       (case w of Pls => v BIT x
-                | Min => bin_pred (v BIT x)
-                | (w BIT y) =>
-      	            NCons (bin_add v (if (x & y) then bin_succ w else w))
-	                  (x~=y))"
-
-primrec
-  bin_mult_Pls: "bin_mult bin.Pls w = bin.Pls"
-  bin_mult_Min: "bin_mult bin.Min w = bin_minus w"
-  bin_mult_BIT: "bin_mult (v BIT x) w =
-	            (if x then (bin_add (NCons (bin_mult v w) False) w)
-	                  else (NCons (bin_mult v w) False))"
+lemma number_of_NCons [simp]:
+     "number_of(NCons w b) = (number_of(w BIT b)::'a::number_ring)"
+by (induct_tac "w", simp_all add: number_of)
 
-
-(** extra rules for bin_succ, bin_pred, bin_add, bin_mult **)
-
-lemma NCons_Pls_0: "NCons bin.Pls False = bin.Pls"
-by simp
-
-lemma NCons_Pls_1: "NCons bin.Pls True = bin.Pls BIT True"
-by simp
-
-lemma NCons_Min_0: "NCons bin.Min False = bin.Min BIT False"
-by simp
-
-lemma NCons_Min_1: "NCons bin.Min True = bin.Min"
-by simp
-
-lemma bin_succ_1: "bin_succ(w BIT True) = (bin_succ w) BIT False"
-by simp
+lemma number_of_succ: "number_of(bin_succ w) = (1 + number_of w ::'a::number_ring)"
+apply (induct_tac "w")
+apply (simp_all add: number_of add_ac)
+done
 
-lemma bin_succ_0: "bin_succ(w BIT False) =  NCons w True"
-by simp
-
-lemma bin_pred_1: "bin_pred(w BIT True) = NCons w False"
-by simp
-
-lemma bin_pred_0: "bin_pred(w BIT False) = (bin_pred w) BIT True"
-by simp
+lemma number_of_pred: "number_of(bin_pred w) = (- 1 + number_of w ::'a::number_ring)"
+apply (induct_tac "w")
+apply (simp_all add: number_of add_assoc [symmetric]) 
+apply (simp add: add_ac)
+done
 
-lemma bin_minus_1: "bin_minus(w BIT True) = bin_pred (NCons (bin_minus w) False)"
-by simp
-
-lemma bin_minus_0: "bin_minus(w BIT False) = (bin_minus w) BIT False"
-by simp
-
-
-(*** bin_add: binary addition ***)
-
-lemma bin_add_BIT_11: "bin_add (v BIT True) (w BIT True) =
-     NCons (bin_add v (bin_succ w)) False"
-apply simp
+lemma number_of_minus: "number_of(bin_minus w) = (- (number_of w)::'a::number_ring)"
+apply (induct_tac "w")
+apply (simp_all del: bin_pred_Pls bin_pred_Min bin_pred_BIT 
+            add: number_of number_of_succ number_of_pred add_assoc)
 done
 
-lemma bin_add_BIT_10: "bin_add (v BIT True) (w BIT False) = NCons (bin_add v w) True"
-by simp
-
-lemma bin_add_BIT_0: "bin_add (v BIT False) (w BIT y) = NCons (bin_add v w) y"
-by auto
+text{*This proof is complicated by the mutual recursion*}
+lemma number_of_add [rule_format]:
+     "\<forall>w. number_of(bin_add v w) = (number_of v + number_of w::'a::number_ring)"
+apply (induct_tac "v")
+apply (simp add: number_of)
+apply (simp add: number_of number_of_pred)
+apply (rule allI)
+apply (induct_tac "w")
+apply (simp_all add: number_of bin_add_BIT_BIT number_of_succ number_of_pred add_ac)
+apply (simp add: add_left_commute [of "1::'a::number_ring"]) 
+done
 
-lemma bin_add_Pls_right: "bin_add w bin.Pls = w"
-by (induct_tac "w", auto)
+lemma number_of_mult:
+     "number_of(bin_mult v w) = (number_of v * number_of w::'a::number_ring)"
+apply (induct_tac "v", simp add: number_of) 
+apply (simp add: number_of number_of_minus) 
+apply (simp add: number_of number_of_add left_distrib add_ac)
+done
 
-lemma bin_add_Min_right: "bin_add w bin.Min = bin_pred w"
-by (induct_tac "w", auto)
-
-lemma bin_add_BIT_BIT: "bin_add (v BIT x) (w BIT y) =
-     NCons(bin_add v (if x & y then (bin_succ w) else w)) (x~= y)"
-apply simp
+text{*The correctness of shifting.  But it doesn't seem to give a measurable
+  speed-up.*}
+lemma double_number_of_BIT:
+     "(1+1) * number_of w = (number_of (w BIT False) ::'a::number_ring)"
+apply (induct_tac "w")
+apply (simp_all add: number_of number_of_add left_distrib add_ac)
 done
 
 
-(*** bin_mult: binary multiplication ***)
+text{*Converting numerals 0 and 1 to their abstract versions*}
+lemma numeral_0_eq_0 [simp]: "Numeral0 = (0::'a::number_ring)"
+by (simp add: number_of) 
+
+lemma numeral_1_eq_1 [simp]: "Numeral1 = (1::'a::number_ring)"
+by (simp add: number_of) 
 
-lemma bin_mult_1: "bin_mult (v BIT True) w = bin_add (NCons (bin_mult v w) False) w"
-by simp
+text{*Special-case simplification for small constants*}
+
+text{*Unary minus for the abstract constant 1. Cannot be inserted
+  as a simprule until later: it is @{text number_of_Min} re-oriented!*}
+lemma numeral_m1_eq_minus_1: "(-1::'a::number_ring) = - 1"
+by (simp add: number_of)
 
-lemma bin_mult_0: "bin_mult (v BIT False) w = NCons (bin_mult v w) False"
-by simp
+lemma mult_minus1 [simp]: "-1 * z = -(z::'a::number_ring)"
+by (simp add: numeral_m1_eq_minus_1)
+
+lemma mult_minus1_right [simp]: "z * -1 = -(z::'a::number_ring)"
+by (simp add: numeral_m1_eq_minus_1)
 
+(*Negation of a coefficient*)
+lemma minus_number_of_mult [simp]:
+     "- (number_of w) * z = number_of(bin_minus w) * (z::'a::number_ring)"
+by (simp add: number_of_minus)
 
-(**** The carry/borrow functions, bin_succ and bin_pred ****)
+text{*Subtraction*}
+lemma diff_number_of_eq:
+     "number_of v - number_of w =
+      (number_of(bin_add v (bin_minus w))::'a::number_ring)"
+by (simp add: diff_minus number_of_add number_of_minus)
 
 
-(** number_of **)
-
-lemma number_of_NCons [simp]:
-     "number_of(NCons w b) = (number_of(w BIT b)::int)"
-apply (induct_tac "w")
-apply (simp_all)
-done
+subsection{*Equality of Binary Numbers*}
 
-lemma number_of_succ: "number_of(bin_succ w) = (1 + number_of w :: int)"
-apply (induct_tac "w")
-apply (simp_all add: zadd_ac)
-done
-
-lemma number_of_pred: "number_of(bin_pred w) = (- 1 + number_of w :: int)"
-apply (induct_tac "w")
-apply (simp_all add: add_assoc [symmetric]) 
-apply (simp add: zadd_ac)
-done
-
-lemma number_of_minus: "number_of(bin_minus w) = (- (number_of w)::int)"
-apply (induct_tac "w", simp, simp)
-apply (simp del: bin_pred_Pls bin_pred_Min bin_pred_BIT add: number_of_succ number_of_pred zadd_assoc)
-done
+text{*First version by Norbert Voelker*}
 
-(*This proof is complicated by the mutual recursion*)
-lemma number_of_add [rule_format (no_asm)]: "! w. number_of(bin_add v w) = (number_of v + number_of w::int)"
-apply (induct_tac "v", simp)
-apply (simp add: number_of_pred)
-apply (rule allI)
-apply (induct_tac "w")
-apply (simp_all add: bin_add_BIT_BIT number_of_succ number_of_pred add_ac)
-apply (simp add: add_left_commute [of "1::int"]) 
-done
-
+lemma eq_number_of_eq:
+  "((number_of x::'a::number_ring) = number_of y) =
+   iszero (number_of (bin_add x (bin_minus y)) :: 'a)"
+by (simp add: iszero_def compare_rls number_of_add number_of_minus)
 
-(*Subtraction*)
-lemma diff_number_of_eq:
-     "number_of v - number_of w = (number_of(bin_add v (bin_minus w))::int)"
-apply (unfold zdiff_def)
-apply (simp add: number_of_add number_of_minus)
-done
+lemma iszero_number_of_Pls: "iszero ((number_of bin.Pls)::'a::number_ring)"
+by (simp add: iszero_def numeral_0_eq_0)
 
-lemmas bin_mult_simps = 
-       int_Suc0_eq_1 zmult_zminus number_of_minus number_of_add
-
-lemma number_of_mult: "number_of(bin_mult v w) = (number_of v * number_of w::int)"
-apply (induct_tac "v")
-apply (simp add: bin_mult_simps)
-apply (simp add: bin_mult_simps)
-apply (simp add: bin_mult_simps zadd_zmult_distrib zadd_ac)
-done
+lemma nonzero_number_of_Min: "~ iszero ((number_of bin.Min)::'a::number_ring)"
+by (simp add: iszero_def numeral_m1_eq_minus_1 eq_commute)
 
 
-(*The correctness of shifting.  But it doesn't seem to give a measurable
-  speed-up.*)
-lemma double_number_of_BIT: "(2::int) * number_of w = number_of (w BIT False)"
-apply (induct_tac "w")
-apply (simp_all add: bin_mult_simps zadd_zmult_distrib zadd_ac)
-done
+subsection{*Comparisons, for Ordered Rings*}
+
+lemma double_eq_0_iff: "(a + a = 0) = (a = (0::'a::ordered_ring))"
+proof -
+  have "a + a = (1+1)*a" by (simp add: left_distrib)
+  with zero_less_two [where 'a = 'a]
+  show ?thesis by force
+qed
 
-
-(** Converting numerals 0 and 1 to their abstract versions **)
-
-lemma int_numeral_0_eq_0: "Numeral0 = (0::int)"
-by simp
+lemma le_imp_0_less: 
+  assumes le: "0 \<le> z" shows "(0::int) < 1 + z"
+proof -
+  have "0 \<le> z" .
+  also have "... < z + 1" by (rule less_add_one) 
+  also have "... = 1 + z" by (simp add: add_ac)
+  finally show "0 < 1 + z" .
+qed
 
-lemma int_numeral_1_eq_1: "Numeral1 = (1::int)"
-by (simp add: int_1 int_Suc0_eq_1)
-
-(*Moving negation out of products: so far for type "int" only*)
-declare zmult_zminus [simp] zmult_zminus_right [simp]
-
-
-(** Special-case simplification for small constants **)
-
-lemma zmult_minus1 [simp]: "-1 * z = -(z::int)"
-by (simp add: compare_rls int_Suc0_eq_1 zmult_zminus)
-
-lemma zmult_minus1_right [simp]: "z * -1 = -(z::int)"
-by (subst zmult_commute, rule zmult_minus1)
+lemma odd_nonzero: "1 + z + z \<noteq> (0::int)";
+proof (cases z rule: int_cases)
+  case (nonneg n)
+  have le: "0 \<le> z+z" by (simp add: prems add_increasing) 
+  thus ?thesis using  le_imp_0_less [OF le]
+    by (auto simp add: add_assoc) 
+next
+  case (neg n)
+  show ?thesis
+  proof
+    assume eq: "1 + z + z = 0"
+    have "0 < 1 + (int n + int n)"
+      by (simp add: le_imp_0_less add_increasing) 
+    also have "... = - (1 + z + z)" by (simp add: prems int_Suc add_ac) 
+    also have "... = 0" by (simp add: eq) 
+    finally have "0<0" ..
+    thus False by blast
+  qed
+qed
 
 
-(*Negation of a coefficient*)
-lemma zminus_number_of_zmult [simp]: "- (number_of w) * z = number_of(bin_minus w) * (z::int)"
-by (simp add: number_of_minus zmult_zminus)
-
-(*Integer unary minus for the abstract constant 1. Cannot be inserted
-  as a simprule until later: it is number_of_Min re-oriented!*)
-lemma zminus_1_eq_m1: "- 1 = (-1::int)"
-by simp
-
-lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)"
-by (cut_tac w = 0 in zless_nat_conj, auto)
-
-
-(** Simplification rules for comparison of binary numbers (Norbert Voelker) **)
+text{*The premise involving @{term Ints} prevents @{term "a = 1/2"}.*}
+lemma Ints_odd_nonzero: "a \<in> Ints ==> 1 + a + a \<noteq> (0::'a::ordered_ring)"
+proof (unfold Ints_def) 
+  assume "a \<in> range of_int"
+  from this obtain z where a: "a = of_int z" ..
+  show ?thesis
+  proof
+    assume eq: "1 + a + a = 0"
+    hence "of_int (1 + z + z) = (of_int 0 :: 'a)" by (simp add: a)
+    hence "1 + z + z = 0" by (simp only: of_int_eq_iff)
+    with odd_nonzero show False by blast
+  qed
+qed 
 
-(** Equals (=) **)
-
-lemma eq_number_of_eq:
-  "((number_of x::int) = number_of y) =
-   iszero (number_of (bin_add x (bin_minus y)) :: int)"
-apply (unfold iszero_def)
-apply (simp add: compare_rls number_of_add number_of_minus)
-done
-
-lemma iszero_number_of_Pls: "iszero ((number_of bin.Pls)::int)"
-by (unfold iszero_def, simp)
-
-lemma nonzero_number_of_Min: "~ iszero ((number_of bin.Min)::int)"
-apply (unfold iszero_def)
-apply (simp add: eq_commute)
-done
+lemma Ints_number_of: "(number_of w :: 'a::number_ring) \<in> Ints"
+by (induct_tac "w", simp_all add: number_of)
 
 lemma iszero_number_of_BIT:
-     "iszero (number_of (w BIT x)::int) = (~x & iszero (number_of w::int))"
-apply (unfold iszero_def)
-apply (cases "(number_of w)::int" rule: int_cases) 
-apply (simp_all (no_asm_simp) add: compare_rls zero_reorient
-         zminus_zadd_distrib [symmetric] int_Suc0_eq_1 [symmetric] zadd_int)
-done
+     "iszero (number_of (w BIT x)::'a) = 
+      (~x & iszero (number_of w::'a::{ordered_ring,number_ring}))"
+by (simp add: iszero_def compare_rls zero_reorient double_eq_0_iff 
+              number_of Ints_odd_nonzero [OF Ints_number_of])
 
 lemma iszero_number_of_0:
-     "iszero (number_of (w BIT False)::int) = iszero (number_of w::int)"
+     "iszero (number_of (w BIT False) :: 'a::{ordered_ring,number_ring}) = 
+      iszero (number_of w :: 'a)"
 by (simp only: iszero_number_of_BIT simp_thms)
 
-lemma iszero_number_of_1: "~ iszero (number_of (w BIT True)::int)"
+lemma iszero_number_of_1:
+     "~ iszero (number_of (w BIT True)::'a::{ordered_ring,number_ring})"
 by (simp only: iszero_number_of_BIT simp_thms)
 
 
 
-(** Less-than (<) **)
+subsection{*The Less-Than Relation*}
 
 lemma less_number_of_eq_neg:
-    "((number_of x::int) < number_of y)
-     = neg (number_of (bin_add x (bin_minus y)) ::int )"
-by (simp add: neg_def number_of_add number_of_minus compare_rls) 
+    "((number_of x::'a::{ordered_ring,number_ring}) < number_of y)
+     = neg (number_of (bin_add x (bin_minus y)) :: 'a)"
+apply (subst less_iff_diff_less_0) 
+apply (simp add: neg_def diff_minus number_of_add number_of_minus)
+done
+
+text{*If @{term Numeral0} is rewritten to 0 then this rule can't be applied:
+  @{term Numeral0} IS @{term "number_of Pls"} *}
+lemma not_neg_number_of_Pls:
+     "~ neg (number_of bin.Pls ::'a::{ordered_ring,number_ring})"
+by (simp add: neg_def numeral_0_eq_0)
+
+lemma neg_number_of_Min:
+     "neg (number_of bin.Min ::'a::{ordered_ring,number_ring})"
+by (simp add: neg_def zero_less_one numeral_m1_eq_minus_1)
+
+lemma double_less_0_iff: "(a + a < 0) = (a < (0::'a::ordered_ring))"
+proof -
+  have "(a + a < 0) = ((1+1)*a < 0)" by (simp add: left_distrib)
+  also have "... = (a < 0)"
+    by (simp add: mult_less_0_iff zero_less_two 
+                  order_less_not_sym [OF zero_less_two]) 
+  finally show ?thesis .
+qed
+
+lemma odd_less_0: "(1 + z + z < 0) = (z < (0::int))";
+proof (cases z rule: int_cases)
+  case (nonneg n)
+  thus ?thesis by (simp add: linorder_not_less add_assoc add_increasing
+                             le_imp_0_less [THEN order_less_imp_le])  
+next
+  case (neg n)
+  thus ?thesis by (simp del: int_Suc
+			add: int_Suc0_eq_1 [symmetric] zadd_int compare_rls)
+qed
+
+text{*The premise involving @{term Ints} prevents @{term "a = 1/2"}.*}
+lemma Ints_odd_less_0: 
+     "a \<in> Ints ==> (1 + a + a < 0) = (a < (0::'a::ordered_ring))";
+proof (unfold Ints_def) 
+  assume "a \<in> range of_int"
+  from this obtain z where a: "a = of_int z" ..
+  hence "((1::'a) + a + a < 0) = (of_int (1 + z + z) < (of_int 0 :: 'a))"
+    by (simp add: prems)
+  also have "... = (z < 0)" by (simp only: of_int_less_iff odd_less_0)
+  also have "... = (a < 0)" by (simp add: prems)
+  finally show ?thesis .
+qed
+
+lemma neg_number_of_BIT:
+     "neg (number_of (w BIT x)::'a) = 
+      neg (number_of w :: 'a::{ordered_ring,number_ring})"
+by (simp add: number_of neg_def double_less_0_iff
+              Ints_odd_less_0 [OF Ints_number_of])
 
 
-(*But if Numeral0 is rewritten to 0 then this rule can't be applied:
-  Numeral0 IS (number_of Pls) *)
-lemma not_neg_number_of_Pls: "~ neg (number_of bin.Pls ::int)"
-by (simp add: neg_def)
-
-lemma neg_number_of_Min: "neg (number_of bin.Min ::int)"
-by (simp add: neg_def int_0_less_1)
-
-lemma neg_number_of_BIT:
-     "neg (number_of (w BIT x)::int) = neg (number_of w ::int)"
-apply simp
-apply (cases "(number_of w)::int" rule: int_cases) 
-apply (simp_all (no_asm_simp) add: int_Suc0_eq_1 [symmetric] zadd_int neg_def zdiff_def [symmetric] compare_rls)
-done
-
-
-(** Less-than-or-equals (\<le>) **)
+text{*Less-Than or Equals*}
 
 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]
+       linorder_not_less [of "number_of w" "number_of v", symmetric, 
+                          standard]
 
-declare le_number_of_eq_not_less [simp]
+lemma le_number_of_eq:
+    "((number_of x::'a::{ordered_ring,number_ring}) \<le> number_of y)
+     = (~ (neg (number_of (bin_add y (bin_minus x)) :: 'a)))"
+by (simp add: le_number_of_eq_not_less less_number_of_eq_neg)
 
 
-(** 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)"
-by (simp add: zabs_def)
+text{*Absolute value (@{term abs})*}
 
-(*0 and 1 require special rewrites because they aren't numerals*)
-lemma zabs_0: "abs (0::int) = 0"
-by (simp add: zabs_def)
+lemma abs_number_of:
+     "abs(number_of x::'a::{ordered_ring,number_ring}) =
+      (if number_of x < (0::'a) then -number_of x else number_of x)"
+by (simp add: abs_if)
 
-lemma zabs_1: "abs (1::int) = 1"
-by (simp del: int_0 int_1 add: int_0 [symmetric] int_1 [symmetric] zabs_def)
 
-(*Re-orientation of the equation nnn=x*)
+text{*Re-orientation of the equation nnn=x*}
 lemma number_of_reorient: "(number_of w = x) = (x = number_of w)"
 by auto
 
@@ -360,14 +294,14 @@
 
 lemmas bin_arith_extra_simps = 
        number_of_add [symmetric]
-       number_of_minus [symmetric] zminus_1_eq_m1
+       number_of_minus [symmetric] numeral_m1_eq_minus_1 [symmetric]
        number_of_mult [symmetric]
        bin_succ_1 bin_succ_0
        bin_pred_1 bin_pred_0
        bin_minus_1 bin_minus_0
        bin_add_Pls_right bin_add_Min_right
        bin_add_BIT_0 bin_add_BIT_10 bin_add_BIT_11
-       diff_number_of_eq zabs_number_of zabs_0 zabs_1
+       diff_number_of_eq abs_number_of abs_zero abs_one
        bin_mult_1 bin_mult_0 NCons_simps
 
 (*For making a minimal simpset, one must include these default simprules
@@ -386,47 +320,33 @@
        less_number_of_eq_neg
        not_neg_number_of_Pls not_neg_0 not_neg_1 not_iszero_1
        neg_number_of_Min neg_number_of_BIT
-       le_number_of_eq_not_less
+       le_number_of_eq
 
 declare bin_arith_extra_simps [simp]
 declare bin_rel_simps [simp]
 
 
-(** Simplification of arithmetic when nested to the right **)
+subsection{*Simplification of arithmetic when nested to the right*}
 
-lemma add_number_of_left [simp]: "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::int)"
-by (simp add: zadd_assoc [symmetric])
+lemma add_number_of_left [simp]:
+     "number_of v + (number_of w + z) =
+      (number_of(bin_add v w) + z::'a::number_ring)"
+by (simp add: add_assoc [symmetric])
 
-lemma mult_number_of_left [simp]: "number_of v * (number_of w * z) = (number_of(bin_mult v w) * z::int)"
-by (simp add: zmult_assoc [symmetric])
+lemma mult_number_of_left [simp]:
+    "number_of v * (number_of w * z) =
+     (number_of(bin_mult v w) * z::'a::number_ring)"
+by (simp add: mult_assoc [symmetric])
 
 lemma add_number_of_diff1:
-    "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::int)"
-apply (unfold zdiff_def)
-apply (rule add_number_of_left)
-done
+    "number_of v + (number_of w - c) = 
+     number_of(bin_add v w) - (c::'a::number_ring)"
+by (simp add: diff_minus add_number_of_left)
 
 lemma add_number_of_diff2 [simp]: "number_of v + (c - number_of w) =
-     number_of (bin_add v (bin_minus w)) + (c::int)"
+     number_of (bin_add v (bin_minus w)) + (c::'a::number_ring)"
 apply (subst diff_number_of_eq [symmetric])
 apply (simp only: compare_rls)
 done
 
-
-
-(** Inserting these natural simprules earlier would break many proofs! **)
-
-(* int (Suc n) = 1 + int n *)
-declare int_Suc [simp]
-
-(* 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/IntArith.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Integ/IntArith.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -8,9 +8,41 @@
 theory IntArith = Bin
 files ("int_arith1.ML"):
 
+text{*Duplicate: can't understand why it's necessary*}
+declare numeral_0_eq_0 [simp]
+
+subsection{*Instantiating Binary Arithmetic for the Integers*}
+
+instance
+  int :: number ..
+
+primrec (*the type constraint is essential!*)
+  number_of_Pls: "number_of bin.Pls = 0"
+  number_of_Min: "number_of bin.Min = - (1::int)"
+  number_of_BIT: "number_of(w BIT x) = (if x then 1 else 0) +
+	                               (number_of w) + (number_of w)"
+
+declare number_of_Pls [simp del]
+        number_of_Min [simp del]
+        number_of_BIT [simp del]
+
+instance int :: number_ring
+proof
+  show "Numeral0 = (0::int)" by (rule number_of_Pls)
+  show "-1 = - (1::int)" by (rule number_of_Min)
+  fix w :: bin and x :: bool
+  show "(number_of (w BIT x) :: int) =
+        (if x then 1 else 0) + number_of w + number_of w"
+    by (rule number_of_BIT)
+qed
+
+
 
 subsection{*Inequality Reasoning for the Arithmetic Simproc*}
 
+lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)"
+by (cut_tac w = 0 in zless_nat_conj, auto)
+
 lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z"
 apply (rule eq_Abs_Integ [of z])
 apply (rule eq_Abs_Integ [of w])
@@ -18,11 +50,115 @@
 done
 
 
+
+lemma add_numeral_0: "Numeral0 + a = (a::'a::number_ring)"
+by simp 
+
+lemma add_numeral_0_right: "a + Numeral0 = (a::'a::number_ring)"
+by simp
+
+lemma mult_numeral_1: "Numeral1 * a = (a::'a::number_ring)"
+by simp 
+
+lemma mult_numeral_1_right: "a * Numeral1 = (a::'a::number_ring)"
+by simp
+
+text{*Theorem lists for the cancellation simprocs. The use of binary numerals
+for 0 and 1 reduces the number of special cases.*}
+
+lemmas add_0s = add_numeral_0 add_numeral_0_right
+lemmas mult_1s = mult_numeral_1 mult_numeral_1_right 
+                 mult_minus1 mult_minus1_right
+
+
+subsection{*Special Arithmetic Rules for Abstract 0 and 1*}
+
+text{*Arithmetic computations are defined for binary literals, which leaves 0
+and 1 as special cases. Addition already has rules for 0, but not 1.
+Multiplication and unary minus already have rules for both 0 and 1.*}
+
+
+lemma binop_eq: "[|f x y = g x y; x = x'; y = y'|] ==> f x' y' = g x' y'"
+by simp
+
+
+lemmas add_number_of_eq = number_of_add [symmetric]
+
+text{*Allow 1 on either or both sides*}
+lemma one_add_one_is_two: "1 + 1 = (2::'a::number_ring)"
+by (simp del: numeral_1_eq_1 add: numeral_1_eq_1 [symmetric] add_number_of_eq)
+
+lemmas add_special =
+    one_add_one_is_two
+    binop_eq [of "op +", OF add_number_of_eq numeral_1_eq_1 refl, standard]
+    binop_eq [of "op +", OF add_number_of_eq refl numeral_1_eq_1, standard]
+
+text{*Allow 1 on either or both sides (1-1 already simplifies to 0)*}
+lemmas diff_special =
+    binop_eq [of "op -", OF diff_number_of_eq numeral_1_eq_1 refl, standard]
+    binop_eq [of "op -", OF diff_number_of_eq refl numeral_1_eq_1, standard]
+
+text{*Allow 0 or 1 on either side with a binary numeral on the other*}
+lemmas eq_special =
+    binop_eq [of "op =", OF eq_number_of_eq numeral_0_eq_0 refl, standard]
+    binop_eq [of "op =", OF eq_number_of_eq numeral_1_eq_1 refl, standard]
+    binop_eq [of "op =", OF eq_number_of_eq refl numeral_0_eq_0, standard]
+    binop_eq [of "op =", OF eq_number_of_eq refl numeral_1_eq_1, standard]
+
+text{*Allow 0 or 1 on either side with a binary numeral on the other*}
+lemmas less_special =
+  binop_eq [of "op <", OF less_number_of_eq_neg numeral_0_eq_0 refl, standard]
+  binop_eq [of "op <", OF less_number_of_eq_neg numeral_1_eq_1 refl, standard]
+  binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_0_eq_0, standard]
+  binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_1_eq_1, standard]
+
+text{*Allow 0 or 1 on either side with a binary numeral on the other*}
+lemmas le_special =
+    binop_eq [of "op \<le>", OF le_number_of_eq numeral_0_eq_0 refl, standard]
+    binop_eq [of "op \<le>", OF le_number_of_eq numeral_1_eq_1 refl, standard]
+    binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_0_eq_0, standard]
+    binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_1_eq_1, standard]
+
+lemmas arith_special = 
+       add_special diff_special eq_special less_special le_special
+
+
 use "int_arith1.ML"
 setup int_arith_setup
 
 
-subsection{*More inequality reasoning*}
+subsection{*Lemmas About Small Numerals*}
+
+lemma of_int_m1 [simp]: "of_int -1 = (-1 :: 'a :: number_ring)"
+proof -
+  have "(of_int -1 :: 'a) = of_int (- 1)" by simp
+  also have "... = - of_int 1" by (simp only: of_int_minus)
+  also have "... = -1" by simp
+  finally show ?thesis .
+qed
+
+lemma abs_minus_one [simp]: "abs (-1) = (1::'a::{ordered_ring,number_ring})"
+by (simp add: abs_if)
+
+lemma of_int_number_of_eq:
+     "of_int (number_of v) = (number_of v :: 'a :: number_ring)"
+apply (induct v)
+apply (simp_all only: number_of of_int_add, simp_all) 
+done
+
+text{*Lemmas for specialist use, NOT as default simprules*}
+lemma mult_2: "2 * z = (z+z::'a::number_ring)"
+proof -
+  have "2*z = (1 + 1)*z" by simp
+  also have "... = z+z" by (simp add: left_distrib)
+  finally show ?thesis .
+qed
+
+lemma mult_2_right: "z * 2 = (z+z::'a::number_ring)"
+by (subst mult_commute, rule mult_2)
+
+
+subsection{*More Inequality Reasoning*}
 
 lemma zless_add1_eq: "(w < z + (1::int)) = (w<z | w=z)"
 by arith
@@ -36,9 +172,6 @@
 lemma zle_add1_eq_le [simp]: "(w < z + 1) = (w\<le>(z::int))"
 by arith
 
-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
 
@@ -86,7 +219,6 @@
 apply (auto simp add: nat_less_iff)
 done
 
-
 lemma nat_le_eq_zle: "0 < w | 0 \<le> z ==> (nat w \<le> nat z) = (w\<le>z)"
 by (auto simp add: linorder_not_less [symmetric] zless_nat_conj)
 
@@ -114,6 +246,7 @@
   with False show ?thesis by simp
 qed
 
+
 subsubsection "Induction principles for int"
 
                      (* `set:int': dummy construction *)
@@ -177,7 +310,7 @@
   from this le show ?thesis by fast
 qed
 
-theorem int_less_induct[consumes 1,case_names base step]:
+theorem int_less_induct [consumes 1,case_names base step]:
   assumes less: "(i::int) < k" and
         base: "P(k - 1)" and
         step: "\<And>i. \<lbrakk>i < k; P i\<rbrakk> \<Longrightarrow> P(i - 1)"
@@ -228,12 +361,7 @@
 apply (drule mult_cancel_left [THEN iffD1], auto)
 done
 
-lemma zless_1_zmult: "[| 1 < m; 1 < n |] ==> 1 < m*(n::int)"
-apply (rule_tac y = "1*n" in order_less_trans)
-apply (rule_tac [2] mult_strict_right_mono)
-apply (simp_all (no_asm_simp))
-done
-
+text{*FIXME: tidy*}
 lemma pos_zmult_eq_1_iff: "0 < (m::int) ==> (m * n = 1) = (m = 1 & n = 1)"
 apply auto
 apply (case_tac "m=1")
@@ -242,7 +370,7 @@
 apply (case_tac [5] "n=1", auto)
 apply (tactic"distinct_subgoals_tac")
 apply (subgoal_tac "1<m*n", simp)
-apply (rule zless_1_zmult, arith)
+apply (rule less_1_mult, arith)
 apply (subgoal_tac "0<n", arith)
 apply (subgoal_tac "0<m*n")
 apply (drule zero_less_mult_iff [THEN iffD1], auto)
@@ -300,6 +428,7 @@
 val zle_diff1_eq = thm "zle_diff1_eq";
 val zle_add1_eq_le = thm "zle_add1_eq_le";
 val nonneg_eq_int = thm "nonneg_eq_int";
+val abs_minus_one = thm "abs_minus_one";
 val nat_eq_iff = thm "nat_eq_iff";
 val nat_eq_iff2 = thm "nat_eq_iff2";
 val nat_less_iff = thm "nat_less_iff";
@@ -312,7 +441,6 @@
 
 val nat_intermed_int_val = thm "nat_intermed_int_val";
 val zmult_eq_self_iff = thm "zmult_eq_self_iff";
-val zless_1_zmult = thm "zless_1_zmult";
 val pos_zmult_eq_1_iff = thm "pos_zmult_eq_1_iff";
 val zmult_eq_1_iff = thm "zmult_eq_1_iff";
 val nat_add_distrib = thm "nat_add_distrib";
--- a/src/HOL/Integ/IntDef.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Integ/IntDef.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -590,26 +590,6 @@
 apply (auto dest: order_less_trans simp add: order_less_imp_le)
 done
 
-
-
-subsection{*Monotonicity of Multiplication*}
-
-lemma zmult_zle_mono2: "[| i \<le> j;  (0::int) \<le> k |] ==> k*i \<le> k*j"
-  by (rule Ring_and_Field.mult_left_mono)
-
-lemma zmult_zless_cancel2: "(m*k < n*k) = (((0::int) < k & m<n) | (k<0 & n<m))"
-  by (rule Ring_and_Field.mult_less_cancel_right)
-
-lemma zmult_zless_cancel1:
-     "(k*m < k*n) = (((0::int) < k & m<n) | (k < 0 & n<m))"
-  by (rule Ring_and_Field.mult_less_cancel_left)
-
-lemma zmult_zle_cancel1:
-     "(k*m \<le> k*n) = (((0::int) < k --> m\<le>n) & (k < 0 --> n\<le>m))"
-  by (rule Ring_and_Field.mult_le_cancel_left)
-
-
-
 text{*A case theorem distinguishing non-negative and negative int*}
 
 lemma negD: "x<0 ==> \<exists>n. x = - (int (Suc n))"
@@ -728,6 +708,8 @@
 declare of_nat_le_iff [of 0, simplified, simp]
 declare of_nat_le_iff [of _ 0, simplified, simp]
 
+text{*The ordering on the semiring is necessary to exclude the possibility of
+a finite field, which indeed wraps back to zero.*}
 lemma of_nat_eq_iff [simp]:
      "(of_nat m = (of_nat n::'a::ordered_semiring)) = (m = n)"
 by (simp add: order_eq_iff) 
@@ -847,6 +829,7 @@
 declare of_int_less_iff [of 0, simplified, simp]
 declare of_int_less_iff [of _ 0, simplified, simp]
 
+text{*The ordering on the ring is necessary. See @{text of_nat_eq_iff} above.*}
 lemma of_int_eq_iff [simp]:
      "(of_int w = (of_int z::'a::ordered_ring)) = (w = z)"
 by (simp add: order_eq_iff) 
@@ -922,6 +905,14 @@
   by (rule Ints_cases) auto
 
 
+(* int (Suc n) = 1 + int n *)
+declare int_Suc [simp]
+
+text{*Simplification of @{term "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]
+
 
 (*Legacy ML bindings, but no longer the structure Int.*)
 ML
@@ -1048,6 +1039,7 @@
 val Nats_1 = thm "Nats_1";
 val Nats_add = thm "Nats_add";
 val Nats_mult = thm "Nats_mult";
+val int_eq_of_nat = thm"int_eq_of_nat";
 val of_int = thm "of_int";
 val of_int_0 = thm "of_int_0";
 val of_int_1 = thm "of_int_1";
--- a/src/HOL/Integ/IntDiv.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Integ/IntDiv.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -101,7 +101,7 @@
  prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2)
 apply (subgoal_tac "b * q' < b * (1 + q) ")
  prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2)
-apply (simp add: zmult_zless_cancel1)
+apply (simp add: mult_less_cancel_left)
 done
 
 lemma unique_quotient_lemma_neg:
@@ -526,7 +526,7 @@
       ==> q \<le> (q'::int)"
 apply (frule q_pos_lemma, assumption+) 
 apply (subgoal_tac "b*q < b* (q' + 1) ")
- apply (simp add: zmult_zless_cancel1)
+ apply (simp add: mult_less_cancel_left)
 apply (subgoal_tac "b*q = r' - r + b'*q'")
  prefer 2 apply simp
 apply (simp (no_asm_simp) add: zadd_zmult_distrib2)
@@ -558,7 +558,7 @@
       ==> q' \<le> (q::int)"
 apply (frule q_neg_lemma, assumption+) 
 apply (subgoal_tac "b*q' < b* (q + 1) ")
- apply (simp add: zmult_zless_cancel1)
+ apply (simp add: mult_less_cancel_left)
 apply (simp add: zadd_zmult_distrib2)
 apply (subgoal_tac "b*q' \<le> b'*q'")
  prefer 2 apply (simp add: mult_right_mono_neg)
@@ -725,7 +725,7 @@
 apply (simp add: zdiff_zmult_distrib2)
 apply (rule order_less_le_trans)
 apply (erule mult_strict_right_mono)
-apply (rule_tac [2] zmult_zle_mono2)
+apply (rule_tac [2] mult_left_mono)
 apply (auto simp add: compare_rls zadd_commute [of 1]
                       add1_zle_eq pos_mod_bound)
 done
@@ -904,11 +904,12 @@
           (if ~b | (0::int) \<le> number_of w                    
            then number_of v div (number_of w)     
            else (number_of v + (1::int)) div (number_of w))"
-apply (simp only: zadd_assoc number_of_BIT)
+apply (simp only: add_assoc number_of_BIT)
 (*create subgoal because the next step can't simplify numerals*)
-apply (subgoal_tac "2 ~= (0::int) ")
-apply (simp del: bin_arith_extra_simps 
-         add: zdiv_zmult_zmult1 pos_zdiv_mult_2 not_0_le_lemma neg_zdiv_mult_2, simp)
+apply (subgoal_tac "2 ~= (0::int) ") 
+apply (simp del: bin_arith_extra_simps arith_special
+         add: zdiv_zmult_zmult1 pos_zdiv_mult_2 not_0_le_lemma neg_zdiv_mult_2)
+apply simp
 done
 
 
@@ -922,7 +923,7 @@
 apply (subgoal_tac "1 < a * 2")
  prefer 2 apply arith
 apply (subgoal_tac "2* (1 + b mod a) \<le> 2*a")
- apply (rule_tac [2] zmult_zle_mono2)
+ apply (rule_tac [2] mult_left_mono)
 apply (auto simp add: zadd_commute [of 1] zmult_commute add1_zle_eq 
                       pos_mod_bound)
 apply (subst zmod_zadd1_eq)
@@ -953,8 +954,9 @@
                 else 2 * ((number_of v + (1::int)) mod number_of w) - 1   
            else 2 * (number_of v mod number_of w))"
 apply (simp only: zadd_assoc number_of_BIT)
-apply (simp del: bin_arith_extra_simps bin_rel_simps 
-         add: zmod_zmult_zmult1 pos_zmod_mult_2 not_0_le_lemma neg_zmod_mult_2, simp)
+apply (simp del: bin_arith_extra_simps bin_rel_simps arith_special
+         add: zmod_zmult_zmult1 pos_zmod_mult_2 not_0_le_lemma neg_zmod_mult_2
+ neg_def) 
 done
 
 
@@ -1114,7 +1116,7 @@
    apply (blast intro: order_less_trans)
   apply (simp add: zero_less_mult_iff)
   apply (subgoal_tac "n * k < n * 1")
-   apply (drule zmult_zless_cancel1 [THEN iffD1], auto)
+   apply (drule mult_less_cancel_left [THEN iffD1], auto)
   done
 
 lemma int_dvd_iff: "(int m dvd z) = (m dvd nat (abs z))"
--- a/src/HOL/Integ/IntDiv_setup.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Integ/IntDiv_setup.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -13,7 +13,7 @@
 val div_name = "Divides.op div";
 val mod_name = "Divides.op mod";
 val mk_binop = HOLogic.mk_binop;
-val mk_sum = Int_Numeral_Simprocs.mk_sum;
+val mk_sum = Int_Numeral_Simprocs.mk_sum HOLogic.intT;
 val dest_sum = Int_Numeral_Simprocs.dest_sum;
 
 (*logic*)
--- a/src/HOL/Integ/NatBin.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Integ/NatBin.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -26,14 +26,14 @@
 lemma nat_number_of [simp]: "nat (number_of w) = number_of w"
 by (simp add: nat_number_of_def)
 
-lemma numeral_0_eq_0: "Numeral0 = (0::nat)"
+lemma nat_numeral_0_eq_0 [simp]: "Numeral0 = (0::nat)"
 by (simp add: nat_number_of_def)
 
-lemma numeral_1_eq_1: "Numeral1 = (1::nat)"
+lemma nat_numeral_1_eq_1 [simp]: "Numeral1 = (1::nat)"
 by (simp add: nat_1 nat_number_of_def)
 
 lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0"
-by (simp add: numeral_1_eq_1)
+by (simp add: nat_numeral_1_eq_1)
 
 lemma numeral_2_eq_2: "2 = Suc (Suc 0)"
 apply (unfold nat_number_of_def)
@@ -52,11 +52,11 @@
 apply (auto elim!: nonneg_eq_int)
 apply (rename_tac m m')
 apply (subgoal_tac "0 <= int m div int m'")
- prefer 2 apply (simp add: numeral_0_eq_0 pos_imp_zdiv_nonneg_iff) 
+ prefer 2 apply (simp add: nat_numeral_0_eq_0 pos_imp_zdiv_nonneg_iff) 
 apply (rule inj_int [THEN injD], simp)
 apply (rule_tac r = "int (m mod m') " in quorem_div)
  prefer 2 apply force
-apply (simp add: nat_less_iff [symmetric] quorem_def numeral_0_eq_0 zadd_int 
+apply (simp add: nat_less_iff [symmetric] quorem_def nat_numeral_0_eq_0 zadd_int 
                  zmult_int)
 done
 
@@ -67,24 +67,23 @@
 apply (auto elim!: nonneg_eq_int)
 apply (rename_tac m m')
 apply (subgoal_tac "0 <= int m mod int m'")
- prefer 2 apply (simp add: nat_less_iff numeral_0_eq_0 pos_mod_sign) 
+ prefer 2 apply (simp add: nat_less_iff nat_numeral_0_eq_0 pos_mod_sign) 
 apply (rule inj_int [THEN injD], simp)
 apply (rule_tac q = "int (m div m') " in quorem_mod)
  prefer 2 apply force
-apply (simp add: nat_less_iff [symmetric] quorem_def numeral_0_eq_0 zadd_int zmult_int)
+apply (simp add: nat_less_iff [symmetric] quorem_def nat_numeral_0_eq_0 zadd_int zmult_int)
 done
 
 
 subsection{*Function @{term int}: Coercion from Type @{typ nat} to @{typ int}*}
 
 (*"neg" is used in rewrite rules for binary comparisons*)
-lemma int_nat_number_of:
+lemma int_nat_number_of [simp]:
      "int (number_of v :: nat) =  
          (if neg (number_of v :: int) then 0  
           else (number_of v :: int))"
 by (simp del: nat_number_of
 	 add: neg_nat nat_number_of_def not_neg_nat add_assoc)
-declare int_nat_number_of [simp]
 
 
 (** Successor **)
@@ -101,19 +100,18 @@
          add: nat_number_of_def neg_nat
               Suc_nat_eq_nat_zadd1 number_of_succ) 
 
-lemma Suc_nat_number_of:
+lemma Suc_nat_number_of [simp]:
      "Suc (number_of v) =  
         (if neg (number_of v :: int) then 1 else number_of (bin_succ v))"
 apply (cut_tac n = 0 in Suc_nat_number_of_add)
 apply (simp cong del: if_weak_cong)
 done
-declare Suc_nat_number_of [simp]
 
 
 (** Addition **)
 
 (*"neg" is used in rewrite rules for binary comparisons*)
-lemma add_nat_number_of:
+lemma add_nat_number_of [simp]:
      "(number_of v :: nat) + number_of v' =  
          (if neg (number_of v :: int) then number_of v'  
           else if neg (number_of v' :: int) then number_of v  
@@ -122,8 +120,6 @@
           simp del: nat_number_of
           simp add: nat_number_of_def nat_add_distrib [symmetric]) 
 
-declare add_nat_number_of [simp]
-
 
 (** Subtraction **)
 
@@ -136,31 +132,29 @@
 apply (simp add: diff_is_0_eq nat_le_eq_zle)
 done
 
-lemma diff_nat_number_of: 
+lemma diff_nat_number_of [simp]: 
      "(number_of v :: nat) - number_of v' =  
         (if neg (number_of v' :: int) then number_of v  
          else let d = number_of (bin_add v (bin_minus v')) in     
               if neg d then 0 else nat d)"
 by (simp del: nat_number_of add: diff_nat_eq_if nat_number_of_def) 
 
-declare diff_nat_number_of [simp]
 
 
 (** Multiplication **)
 
-lemma mult_nat_number_of:
+lemma mult_nat_number_of [simp]:
      "(number_of v :: nat) * number_of v' =  
        (if neg (number_of v :: int) then 0 else number_of (bin_mult v v'))"
 by (force dest!: neg_nat
           simp del: nat_number_of
           simp add: nat_number_of_def nat_mult_distrib [symmetric]) 
 
-declare mult_nat_number_of [simp]
 
 
 (** Quotient **)
 
-lemma div_nat_number_of:
+lemma div_nat_number_of [simp]:
      "(number_of v :: nat)  div  number_of v' =  
           (if neg (number_of v :: int) then 0  
            else nat (number_of v div number_of v'))"
@@ -168,12 +162,14 @@
           simp del: nat_number_of
           simp add: nat_number_of_def nat_div_distrib [symmetric]) 
 
-declare div_nat_number_of [simp]
+lemma one_div_nat_number_of [simp]:
+     "(Suc 0)  div  number_of v' = (nat (1 div number_of v'))" 
+by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) 
 
 
 (** Remainder **)
 
-lemma mod_nat_number_of:
+lemma mod_nat_number_of [simp]:
      "(number_of v :: nat)  mod  number_of v' =  
         (if neg (number_of v :: int) then 0  
          else if neg (number_of v' :: int) then number_of v  
@@ -182,15 +178,21 @@
           simp del: nat_number_of
           simp add: nat_number_of_def nat_mod_distrib [symmetric]) 
 
-declare mod_nat_number_of [simp]
+lemma one_mod_nat_number_of [simp]:
+     "(Suc 0)  mod  number_of v' =  
+        (if neg (number_of v' :: int) then Suc 0
+         else nat (1 mod number_of v'))"
+by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) 
+
+
 
 ML
 {*
 val nat_number_of_def = thm"nat_number_of_def";
 
 val nat_number_of = thm"nat_number_of";
-val numeral_0_eq_0 = thm"numeral_0_eq_0";
-val numeral_1_eq_1 = thm"numeral_1_eq_1";
+val nat_numeral_0_eq_0 = thm"nat_numeral_0_eq_0";
+val nat_numeral_1_eq_1 = thm"nat_numeral_1_eq_1";
 val numeral_1_eq_Suc_0 = thm"numeral_1_eq_Suc_0";
 val numeral_2_eq_2 = thm"numeral_2_eq_2";
 val nat_div_distrib = thm"nat_div_distrib";
@@ -208,29 +210,6 @@
 *}
 
 
-ML
-{*
-structure NatAbstractNumeralsData =
-  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    = numeral_0_eq_0
-  val numeral_1_eq_1    = numeral_1_eq_Suc_0
-  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 NatAbstractNumerals = AbstractNumeralsFun (NatAbstractNumeralsData);
-
-val nat_eval_numerals = 
-  map Bin_Simprocs.prep_simproc
-   [("nat_div_eval_numerals", ["(Suc 0) div m"], NatAbstractNumerals.proc div_nat_number_of),
-    ("nat_mod_eval_numerals", ["(Suc 0) mod m"], NatAbstractNumerals.proc mod_nat_number_of)];
-
-Addsimprocs nat_eval_numerals;
-*}
-
 (*** Comparisons ***)
 
 (** Equals (=) **)
@@ -270,7 +249,7 @@
 
 
 (*Maps #n to n for n = 0, 1, 2*)
-lemmas numerals = numeral_0_eq_0 numeral_1_eq_1 numeral_2_eq_2
+lemmas numerals = nat_numeral_0_eq_0 nat_numeral_1_eq_1 numeral_2_eq_2
 
 
 subsection{*General Theorems About Powers Involving Binary Numerals*}
@@ -398,7 +377,7 @@
 lemma eq_number_of_0:
      "(number_of v = (0::nat)) =  
       (if neg (number_of v :: int) then True else iszero (number_of v :: int))"
-apply (simp add: numeral_0_eq_0 [symmetric] iszero_0)
+apply (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric] iszero_0)
 done
 
 lemma eq_0_number_of:
@@ -409,13 +388,13 @@
 
 lemma less_0_number_of:
      "((0::nat) < number_of v) = neg (number_of (bin_minus v) :: int)"
-by (simp add: numeral_0_eq_0 [symmetric])
+by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric])
 
 (*Simplification already handles n<0, n<=0 and 0<=n.*)
 declare eq_number_of_0 [simp] eq_0_number_of [simp] less_0_number_of [simp]
 
 lemma neg_imp_number_of_eq_0: "neg (number_of v :: int) ==> number_of v = (0::nat)"
-by (simp add: numeral_0_eq_0 [symmetric] iszero_0)
+by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric] iszero_0)
 
 
 
@@ -563,7 +542,7 @@
            then (let w = z ^ (number_of w) in  z*w*w)    
            else 1)"
 apply (simp del: nat_number_of  add: nat_number_of_def number_of_BIT Let_def)
-apply (simp only: number_of_add int_numeral_1_eq_1 not_neg_eq_ge_0 neg_eq_less_0) 
+apply (simp only: number_of_add nat_numeral_1_eq_1 not_neg_eq_ge_0 neg_eq_less_0) 
 apply (rule_tac x = "number_of w" in spec, clarify)
 apply (auto simp add: nat_add_distrib nat_mult_distrib zpower_even power2_eq_square neg_nat)
 done
--- a/src/HOL/Integ/NatSimprocs.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Integ/NatSimprocs.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -31,8 +31,10 @@
 lemma Suc_diff_number_of:
      "neg (number_of (bin_minus v)::int) ==>  
       Suc m - (number_of v) = m - (number_of (bin_pred v))"
-apply (subst Suc_diff_eq_diff_pred, simp, simp)
-apply (force simp only: diff_nat_number_of less_0_number_of [symmetric] 
+apply (subst Suc_diff_eq_diff_pred)
+apply (simp add: ); 
+apply (simp del: nat_numeral_1_eq_1); 
+apply (auto simp only: diff_nat_number_of less_0_number_of [symmetric] 
                         neg_number_of_bin_pred_iff_0)
 done
 
@@ -54,7 +56,8 @@
          if neg pv then nat_case a f n else f (nat pv + n))"
 apply (subst add_eq_if)
 apply (simp split add: nat.split
-            add: numeral_1_eq_Suc_0 [symmetric] Let_def 
+            del: nat_numeral_1_eq_1
+	    add: numeral_1_eq_Suc_0 [symmetric] Let_def 
                  neg_imp_number_of_eq_0 neg_number_of_bin_pred_iff_0)
 done
 
@@ -74,6 +77,7 @@
                    else f (nat pv + n) (nat_rec a f (nat pv + n)))"
 apply (subst add_eq_if)
 apply (simp split add: nat.split
+            del: nat_numeral_1_eq_1
             add: numeral_1_eq_Suc_0 [symmetric] Let_def neg_imp_number_of_eq_0
                  neg_number_of_bin_pred_iff_0)
 done
@@ -106,8 +110,6 @@
 lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)"
 by simp
 
-declare numeral_0_eq_0 [simp] numeral_1_eq_1 [simp] 
-
 text{*Can be used to eliminate long strings of Sucs, but not by default*}
 lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n"
 by simp
@@ -194,4 +196,20 @@
 declare divide_eq_eq [of _ "number_of w", standard, simp]
 
 
+subsubsection{*Division By @{term "-1"}*}
+
+lemma divide_minus1 [simp]:
+     "x/-1 = -(x::'a::{field,division_by_zero,number_ring})" 
+by simp
+
+lemma minus1_divide [simp]:
+     "-1 / (x::'a::{field,division_by_zero,number_ring}) = - (1/x)"
+by (simp add: divide_inverse_zero inverse_minus_eq)
+
+ML
+{*
+val divide_minus1 = thm "divide_minus1";
+val minus1_divide = thm "minus1_divide";
+*}
+
 end
--- a/src/HOL/Integ/int_arith1.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Integ/int_arith1.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -30,11 +30,7 @@
 
 val neg_def = thm "neg_def";
 val iszero_def = thm "iszero_def";
-val not_neg_int = thm "not_neg_int";
-val neg_zminus_int = thm "neg_zminus_int";
 
-val zadd_ac = thms "Ring_and_Field.add_ac"
-val zmult_ac = thms "Ring_and_Field.mult_ac"
 val NCons_Pls_0 = thm"NCons_Pls_0";
 val NCons_Pls_1 = thm"NCons_Pls_1";
 val NCons_Min_0 = thm"NCons_Min_0";
@@ -61,12 +57,12 @@
 val diff_number_of_eq = thm"diff_number_of_eq";
 val number_of_mult = thm"number_of_mult";
 val double_number_of_BIT = thm"double_number_of_BIT";
-val int_numeral_0_eq_0 = thm"int_numeral_0_eq_0";
-val int_numeral_1_eq_1 = thm"int_numeral_1_eq_1";
-val zmult_minus1 = thm"zmult_minus1";
-val zmult_minus1_right = thm"zmult_minus1_right";
-val zminus_number_of_zmult = thm"zminus_number_of_zmult";
-val zminus_1_eq_m1 = thm"zminus_1_eq_m1";
+val numeral_0_eq_0 = thm"numeral_0_eq_0";
+val numeral_1_eq_1 = thm"numeral_1_eq_1";
+val numeral_m1_eq_minus_1 = thm"numeral_m1_eq_minus_1";
+val mult_minus1 = thm"mult_minus1";
+val mult_minus1_right = thm"mult_minus1_right";
+val minus_number_of_mult = thm"minus_number_of_mult";
 val zero_less_nat_eq = thm"zero_less_nat_eq";
 val eq_number_of_eq = thm"eq_number_of_eq";
 val iszero_number_of_Pls = thm"iszero_number_of_Pls";
@@ -75,13 +71,12 @@
 val iszero_number_of_0 = thm"iszero_number_of_0";
 val iszero_number_of_1 = thm"iszero_number_of_1";
 val less_number_of_eq_neg = thm"less_number_of_eq_neg";
+val le_number_of_eq = thm"le_number_of_eq";
 val not_neg_number_of_Pls = thm"not_neg_number_of_Pls";
 val neg_number_of_Min = thm"neg_number_of_Min";
 val neg_number_of_BIT = thm"neg_number_of_BIT";
 val le_number_of_eq_not_less = thm"le_number_of_eq_not_less";
-val zabs_number_of = thm"zabs_number_of";
-val zabs_0 = thm"zabs_0";
-val zabs_1 = thm"zabs_1";
+val abs_number_of = thm"abs_number_of";
 val number_of_reorient = thm"number_of_reorient";
 val add_number_of_left = thm"add_number_of_left";
 val mult_number_of_left = thm"mult_number_of_left";
@@ -91,7 +86,6 @@
 val eq_iff_diff_eq_0 = thm"eq_iff_diff_eq_0";
 val le_iff_diff_le_0 = thm"le_iff_diff_le_0";
 
-val bin_mult_simps = thms"bin_mult_simps";
 val NCons_simps = thms"NCons_simps";
 val bin_arith_extra_simps = thms"bin_arith_extra_simps";
 val bin_arith_simps = thms"bin_arith_simps";
@@ -107,6 +101,7 @@
 val le_add_iff1 = thm"le_add_iff1";
 val le_add_iff2 = thm"le_add_iff2";
 
+val arith_special = thms"arith_special";
 
 structure Bin_Simprocs =
   struct
@@ -128,43 +123,6 @@
   fun simplify_meta_eq f_number_of_eq f_eq =
       mk_meta_eq ([f_eq, f_number_of_eq] MRS trans)
 
-  structure IntAbstractNumeralsData =
-    struct
-    val dest_eq		= HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of
-    val is_numeral	= is_numeral
-    val numeral_0_eq_0    = int_numeral_0_eq_0
-    val numeral_1_eq_1    = int_numeral_1_eq_1
-    val prove_conv	= prove_conv_nohyps_novars
-    fun norm_tac simps	= ALLGOALS (simp_tac (HOL_ss addsimps simps))
-    val simplify_meta_eq  = simplify_meta_eq 
-    end
-
-  structure IntAbstractNumerals = AbstractNumeralsFun (IntAbstractNumeralsData)
-
-
-  (*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
-     [("int_add_eval_numerals",
-       ["(m::int) + 1", "(m::int) + number_of v"], 
-       IntAbstractNumerals.proc (number_of_add RS sym)),
-      ("int_diff_eval_numerals",
-       ["(m::int) - 1", "(m::int) - number_of v"], 
-       IntAbstractNumerals.proc diff_number_of_eq),
-      ("int_eq_eval_numerals",
-       ["(m::int) = 0", "(m::int) = 1", "(m::int) = number_of v"], 
-       IntAbstractNumerals.proc eq_number_of_eq),
-      ("int_less_eval_numerals",
-       ["(m::int) < 0", "(m::int) < 1", "(m::int) < number_of v"], 
-       IntAbstractNumerals.proc less_number_of_eq_neg),
-      ("int_le_eval_numerals",
-       ["(m::int) <= 0", "(m::int) <= 1", "(m::int) <= number_of v"],
-       IntAbstractNumerals.proc le_number_of_eq_not_less)]
-
   (*reorientation simprules using ==, for the following simproc*)
   val meta_zero_reorient = zero_reorient RS eq_reflection
   val meta_one_reorient = one_reorient RS eq_reflection
@@ -188,7 +146,7 @@
   end;
 
 
-Addsimprocs Bin_Simprocs.eval_numerals;
+Addsimps arith_special;
 Addsimprocs [Bin_Simprocs.reorient_simproc];
 
 
@@ -197,15 +155,11 @@
 
 (*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 = [int_numeral_0_eq_0 RS sym, int_numeral_1_eq_1 RS sym];
-val numeral_sym_ss = HOL_ss addsimps numeral_syms;
-
-fun rename_numerals th =
-    simplify numeral_sym_ss (Thm.transfer (the_context ()) th);
+val numeral_syms = [numeral_0_eq_0 RS sym, numeral_1_eq_1 RS sym];
 
 (*Utilities*)
 
-fun mk_numeral n = HOLogic.number_of_const HOLogic.intT $ HOLogic.mk_bin n;
+fun mk_numeral T n = HOLogic.number_of_const T $ HOLogic.mk_bin n;
 
 (*Decodes a binary INTEGER*)
 fun dest_numeral (Const("0", _)) = 0
@@ -220,21 +174,23 @@
          handle TERM _ => find_first_numeral (t::past) terms)
   | find_first_numeral past [] = raise TERM("find_first_numeral", []);
 
-val zero = mk_numeral 0;
 val mk_plus = HOLogic.mk_binop "op +";
 
-val uminus_const = Const ("uminus", HOLogic.intT --> HOLogic.intT);
+fun mk_minus t = 
+  let val T = Term.fastype_of t
+  in Const ("uminus", T --> T) $ t
+  end;
 
 (*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);
+fun mk_sum T []        = mk_numeral T 0
+  | mk_sum T [t,u]     = mk_plus (t, u)
+  | mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts);
 
 (*this version ALWAYS includes a trailing zero*)
-fun long_mk_sum []        = zero
-  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
+fun long_mk_sum T []        = mk_numeral T 0
+  | long_mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts);
 
-val dest_plus = HOLogic.dest_bin "op +" HOLogic.intT;
+val dest_plus = HOLogic.dest_bin "op +" Term.dummyT;
 
 (*decompose additions AND subtractions as a sum*)
 fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
@@ -242,22 +198,27 @@
   | 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;
+        if pos then t::ts else mk_minus 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.intT;
+val dest_diff = HOLogic.dest_bin "op -" Term.dummyT;
 
-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);
+fun mk_prod T = 
+  let val one = mk_numeral T 1
+  fun mk [] = one
+    | mk [t] = t
+    | mk (t :: ts) = if t = one then mk ts else mk_times (t, mk ts)
+  in mk end;
 
-val dest_times = HOLogic.dest_bin "op *" HOLogic.intT;
+(*This version ALWAYS includes a trailing one*)
+fun long_mk_prod T []        = mk_numeral T 1
+  | long_mk_prod T (t :: ts) = mk_times (t, mk_prod T ts);
+
+val dest_times = HOLogic.dest_bin "op *" Term.dummyT;
 
 fun dest_prod t =
       let val (t,u) = dest_times t
@@ -265,7 +226,7 @@
       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);
+fun mk_coeff (k, t) = mk_times (mk_numeral (Term.fastype_of t) k, t);
 
 (*Express t as a product of (possibly) a numeral with other sorted terms*)
 fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t
@@ -273,7 +234,7 @@
     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;
+    in (sign*n, mk_prod (Term.fastype_of t) ts') end;
 
 (*Find first coefficient-term THAT MATCHES u*)
 fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
@@ -286,13 +247,12 @@
 
 
 (*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
-val add_0s =  map rename_numerals [zadd_0, zadd_0_right];
-val mult_1s = map rename_numerals [zmult_1, zmult_1_right] @
-              [zmult_minus1, zmult_minus1_right];
+val add_0s =  thms "add_0s";
+val mult_1s = thms "mult_1s";
 
 (*To perform binary arithmetic.  The "left" rewriting handles patterns
   created by the simprocs, such as 3 * (5 * x). *)
-val bin_simps = [int_numeral_0_eq_0 RS sym, int_numeral_1_eq_1 RS sym,
+val bin_simps = [numeral_0_eq_0 RS sym, numeral_1_eq_1 RS sym,
                  add_number_of_left, mult_number_of_left] @
                 bin_arith_simps @ bin_rel_simps;
 
@@ -302,25 +262,25 @@
     bin_simps \\ [add_number_of_left, number_of_add RS sym];
 
 (*To evaluate binary negations of coefficients*)
-val zminus_simps = NCons_simps @
-                   [zminus_1_eq_m1, number_of_minus RS sym,
+val minus_simps = NCons_simps @
+                   [numeral_m1_eq_minus_1 RS sym, number_of_minus RS sym,
                     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 = [zdiff_def, zminus_zadd_distrib, zminus_zminus];
+val diff_simps = [diff_minus, minus_add_distrib, minus_minus];
 
 (*push the unary minus down: - x * y = x * - y *)
-val int_minus_mult_eq_1_to_2 =
-    [zmult_zminus, zmult_zminus_right RS sym] MRS trans |> standard;
+val minus_mult_eq_1_to_2 =
+    [minus_mult_left RS sym, minus_mult_right] MRS trans |> standard;
 
 (*to extract again any uncancelled minuses*)
-val int_minus_from_mult_simps =
-    [zminus_zminus, zmult_zminus, zmult_zminus_right];
+val 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 int_mult_minus_simps =
-    [zmult_assoc, zmult_zminus RS sym, int_minus_mult_eq_1_to_2];
+val mult_minus_simps =
+    [mult_assoc, minus_mult_left, minus_mult_eq_1_to_2];
 
 (*Apply the given rewrite (if present) just once*)
 fun trans_tac None      = all_tac
@@ -340,10 +300,10 @@
   val trans_tac         = trans_tac
   val norm_tac =
      ALLGOALS (simp_tac (HOL_ss addsimps numeral_syms@add_0s@mult_1s@
-                                         diff_simps@zminus_simps@zadd_ac))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@int_mult_minus_simps))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps int_minus_from_mult_simps@
-                                              zadd_ac@zmult_ac))
+                                         diff_simps@minus_simps@add_ac))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@mult_minus_simps))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps 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 (add_0s@mult_1s)
   end;
@@ -353,7 +313,7 @@
  (open CancelNumeralsCommon
   val prove_conv = Bin_Simprocs.prove_conv
   val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.intT
+  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
   val bal_add1 = eq_add_iff1 RS trans
   val bal_add2 = eq_add_iff2 RS trans
 );
@@ -362,7 +322,7 @@
  (open CancelNumeralsCommon
   val prove_conv = Bin_Simprocs.prove_conv
   val mk_bal   = HOLogic.mk_binrel "op <"
-  val dest_bal = HOLogic.dest_bin "op <" HOLogic.intT
+  val dest_bal = HOLogic.dest_bin "op <" Term.dummyT
   val bal_add1 = less_add_iff1 RS trans
   val bal_add2 = less_add_iff2 RS trans
 );
@@ -371,7 +331,7 @@
  (open CancelNumeralsCommon
   val prove_conv = Bin_Simprocs.prove_conv
   val mk_bal   = HOLogic.mk_binrel "op <="
-  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.intT
+  val dest_bal = HOLogic.dest_bin "op <=" Term.dummyT
   val bal_add1 = le_add_iff1 RS trans
   val bal_add2 = le_add_iff2 RS trans
 );
@@ -379,19 +339,28 @@
 val cancel_numerals =
   map Bin_Simprocs.prep_simproc
    [("inteq_cancel_numerals",
-     ["(l::int) + m = n", "(l::int) = m + n",
-      "(l::int) - m = n", "(l::int) = m - n",
-      "(l::int) * m = n", "(l::int) = m * n"],
+     ["(l::'a::number_ring) + m = n",
+      "(l::'a::number_ring) = m + n",
+      "(l::'a::number_ring) - m = n",
+      "(l::'a::number_ring) = m - n",
+      "(l::'a::number_ring) * m = n",
+      "(l::'a::number_ring) = m * n"],
      EqCancelNumerals.proc),
     ("intless_cancel_numerals",
-     ["(l::int) + m < n", "(l::int) < m + n",
-      "(l::int) - m < n", "(l::int) < m - n",
-      "(l::int) * m < n", "(l::int) < m * n"],
+     ["(l::'a::{ordered_ring,number_ring}) + m < n",
+      "(l::'a::{ordered_ring,number_ring}) < m + n",
+      "(l::'a::{ordered_ring,number_ring}) - m < n",
+      "(l::'a::{ordered_ring,number_ring}) < m - n",
+      "(l::'a::{ordered_ring,number_ring}) * m < n",
+      "(l::'a::{ordered_ring,number_ring}) < m * n"],
      LessCancelNumerals.proc),
     ("intle_cancel_numerals",
-     ["(l::int) + m <= n", "(l::int) <= m + n",
-      "(l::int) - m <= n", "(l::int) <= m - n",
-      "(l::int) * m <= n", "(l::int) <= m * n"],
+     ["(l::'a::{ordered_ring,number_ring}) + m <= n",
+      "(l::'a::{ordered_ring,number_ring}) <= m + n",
+      "(l::'a::{ordered_ring,number_ring}) - m <= n",
+      "(l::'a::{ordered_ring,number_ring}) <= m - n",
+      "(l::'a::{ordered_ring,number_ring}) * m <= n",
+      "(l::'a::{ordered_ring,number_ring}) <= m * n"],
      LeCancelNumerals.proc)];
 
 
@@ -407,10 +376,10 @@
   val trans_tac          = trans_tac
   val norm_tac =
      ALLGOALS (simp_tac (HOL_ss addsimps numeral_syms@add_0s@mult_1s@
-                                         diff_simps@zminus_simps@zadd_ac))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@int_mult_minus_simps))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps int_minus_from_mult_simps@
-                                              zadd_ac@zmult_ac))
+                                         diff_simps@minus_simps@add_ac))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@mult_minus_simps))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps 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 (add_0s@mult_1s)
@@ -420,7 +389,9 @@
 
 val combine_numerals =
   Bin_Simprocs.prep_simproc
-    ("int_combine_numerals", ["(i::int) + j", "(i::int) - j"], CombineNumerals.proc);
+    ("int_combine_numerals", 
+     ["(i::'a::number_ring) + j", "(i::'a::number_ring) - j"], 
+     CombineNumerals.proc);
 
 end;
 
@@ -465,43 +436,28 @@
 *)
 
 
-(** Constant folding for integer plus and times **)
+(** Constant folding for multiplication in semirings **)
 
-(*We do not need
-    structure Nat_Plus_Assoc = Assoc_Fold (Nat_Plus_Assoc_Data);
-    structure Int_Plus_Assoc = Assoc_Fold (Int_Plus_Assoc_Data);
-  because combine_numerals does the same thing*)
+(*We do not need folding for addition: combine_numerals does the same thing*)
 
-structure Int_Times_Assoc_Data : ASSOC_FOLD_DATA =
+structure Semiring_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.intT
-  val plus   = Const ("op *", [HOLogic.intT,HOLogic.intT] ---> HOLogic.intT);
-  val add_ac = zmult_ac
-end;
-
-structure Int_Times_Assoc = Assoc_Fold (Int_Times_Assoc_Data);
-
-Addsimprocs [Int_Times_Assoc.conv];
-
-
-(** The same for the naturals **)
-
-structure Nat_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.natT
-  val plus   = Const ("op *", [HOLogic.natT,HOLogic.natT] ---> HOLogic.natT);
   val add_ac = mult_ac
 end;
 
-structure Nat_Times_Assoc = Assoc_Fold (Nat_Times_Assoc_Data);
+structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data);
 
-Addsimprocs [Nat_Times_Assoc.conv];
+val assoc_fold_simproc =
+  Bin_Simprocs.prep_simproc
+   ("semiring_assoc_fold", ["(a::'a::semiring) * b"],
+    Semiring_Times_Assoc.proc);
+
+Addsimprocs [assoc_fold_simproc];
+
+
 
 
 (*** decision procedure for linear arithmetic ***)
@@ -519,18 +475,19 @@
 
 (* reduce contradictory <= to False *)
 val add_rules =
-    simp_thms @ bin_arith_simps @ bin_rel_simps @
-    [int_numeral_0_eq_0, int_numeral_1_eq_1,
+    simp_thms @ bin_arith_simps @ bin_rel_simps @ arith_special @
+    [numeral_0_eq_0, numeral_1_eq_1,
      minus_zero, diff_minus, 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,
      minus_add_distrib, minus_minus, mult_assoc,
-     int_0, int_1, int_Suc, zadd_int RS sym, zmult_int RS sym,
-     le_number_of_eq_not_less];
+     of_nat_0, of_nat_1, of_nat_Suc, of_nat_add, of_nat_mult,
+     of_int_0, of_int_1, of_int_add, of_int_mult, int_eq_of_nat,
+     zero_neq_one, zero_less_one, zero_le_one, 
+     zero_neq_one RS not_sym, not_one_le_zero, not_one_less_zero];
 
-val simprocs = [Int_Times_Assoc.conv, Int_Numeral_Simprocs.combine_numerals]@
-               Int_Numeral_Simprocs.cancel_numerals @
-               Bin_Simprocs.eval_numerals;
+val simprocs = [assoc_fold_simproc, Int_Numeral_Simprocs.combine_numerals]@
+               Int_Numeral_Simprocs.cancel_numerals;
 
 in
 
@@ -543,6 +500,7 @@
     simpset = simpset addsimps add_rules
                       addsimprocs simprocs
                       addcongs [if_weak_cong]}),
+  arith_inj_const ("IntDef.of_nat", HOLogic.natT --> HOLogic.intT),
   arith_inj_const ("IntDef.int", HOLogic.natT --> HOLogic.intT),
   arith_discrete ("IntDef.int", true)];
 
@@ -550,7 +508,10 @@
 
 val fast_int_arith_simproc =
   Simplifier.simproc (Theory.sign_of (the_context()))
-  "fast_int_arith" ["(m::int) < n","(m::int) <= n", "(m::int) = n"] Fast_Arith.lin_arith_prover;
+  "fast_int_arith" 
+     ["(m::'a::{ordered_ring,number_ring}) < n",
+      "(m::'a::{ordered_ring,number_ring}) <= n",
+      "(m::'a::{ordered_ring,number_ring}) = n"] Fast_Arith.lin_arith_prover;
 
 Addsimprocs [fast_int_arith_simproc]
 
--- a/src/HOL/Integ/int_factor_simprocs.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Integ/int_factor_simprocs.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -31,9 +31,9 @@
   val dest_coeff        = dest_coeff 1
   val trans_tac         = trans_tac
   val norm_tac =
-     ALLGOALS (simp_tac (HOL_ss addsimps int_minus_from_mult_simps@mult_1s))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@int_mult_minus_simps))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps zmult_ac))
+     ALLGOALS (simp_tac (HOL_ss addsimps minus_from_mult_simps@mult_1s))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@mult_minus_simps))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps mult_ac))
   val numeral_simp_tac  = ALLGOALS (simp_tac (HOL_ss addsimps bin_simps))
   val simplify_meta_eq  = simplify_meta_eq mult_1s
   end
@@ -132,11 +132,6 @@
   open Int_Numeral_Simprocs
 in
 
-
-(*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) =
@@ -147,7 +142,7 @@
 (*Final simplification: cancel + and *  *)
 fun cancel_simplify_meta_eq cancel_th th =
     Int_Numeral_Simprocs.simplify_meta_eq
-        [zmult_1, zmult_1_right]
+        [mult_1, mult_1_right]
         (([th, cancel_th]) MRS trans);
 
 structure CancelFactorCommon =
@@ -158,9 +153,11 @@
   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@zmult_ac))
+  val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@mult_ac))
   end;
 
+(*mult_cancel_left requires an ordered ring, such as int. The version in
+  rat_arith.ML works for all fields.*)
 structure EqCancelFactor = ExtractCommonTermFun
  (open CancelFactorCommon
   val prove_conv = Bin_Simprocs.prove_conv
@@ -169,6 +166,8 @@
   val simplify_meta_eq  = cancel_simplify_meta_eq mult_cancel_left
 );
 
+(*int_mult_div_cancel_disj is for integer division (div). The version in
+  rat_arith.ML works for all fields, using real division (/).*)
 structure DivideCancelFactor = ExtractCommonTermFun
  (open CancelFactorCommon
   val prove_conv = Bin_Simprocs.prove_conv
--- a/src/HOL/Integ/nat_simprocs.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Integ/nat_simprocs.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -11,7 +11,7 @@
 
 (*Maps n to #n for n = 0, 1, 2*)
 val numeral_syms =
-       [numeral_0_eq_0 RS sym, numeral_1_eq_1 RS sym, numeral_2_eq_2 RS sym];
+       [nat_numeral_0_eq_0 RS sym, nat_numeral_1_eq_1 RS sym, numeral_2_eq_2 RS sym];
 val numeral_sym_ss = HOL_ss addsimps numeral_syms;
 
 fun rename_numerals th =
@@ -65,7 +65,7 @@
 
 val trans_tac = Int_Numeral_Simprocs.trans_tac;
 
-val bin_simps = [numeral_0_eq_0 RS sym, numeral_1_eq_1 RS sym,
+val bin_simps = [nat_numeral_0_eq_0 RS sym, nat_numeral_1_eq_1 RS sym,
                  add_nat_number_of, nat_number_of_add_left,
                  diff_nat_number_of, le_number_of_eq_not_less,
                  less_nat_number_of, mult_nat_number_of,
@@ -126,7 +126,7 @@
 
 val simplify_meta_eq =
     Int_Numeral_Simprocs.simplify_meta_eq
-        ([numeral_0_eq_0, numeral_1_eq_Suc_0, add_0, add_0_right,
+        ([nat_numeral_0_eq_0, numeral_1_eq_Suc_0, add_0, add_0_right,
           mult_0, mult_0_right, mult_1, mult_1_right] @ contra_rules);
 
 
@@ -153,7 +153,7 @@
 
 structure CancelNumeralsCommon =
   struct
-  val mk_sum            = mk_sum
+  val mk_sum            = (fn T:typ => mk_sum)
   val dest_sum          = dest_Sucs_sum
   val mk_coeff          = mk_coeff
   val dest_coeff        = dest_coeff
@@ -236,7 +236,7 @@
 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 mk_sum            = (fn T:typ => long_mk_sum)  (*to work for 2*x + 3*x *)
   val dest_sum          = restricted_dest_Sucs_sum
   val mk_coeff          = mk_coeff
   val dest_coeff        = dest_coeff
@@ -346,7 +346,7 @@
 
 structure CancelFactorCommon =
   struct
-  val mk_sum            = long_mk_prod
+  val mk_sum            = (fn T:typ => long_mk_prod)
   val dest_sum          = dest_prod
   val mk_coeff          = mk_coeff
   val dest_coeff        = dest_coeff
@@ -514,8 +514,7 @@
    eq_number_of_0, eq_0_number_of, less_0_number_of,
    nat_number_of, if_True, if_False];
 
-val simprocs = [Nat_Times_Assoc.conv,
-                Nat_Numeral_Simprocs.combine_numerals]@
+val simprocs = [Nat_Numeral_Simprocs.combine_numerals]@
                 Nat_Numeral_Simprocs.cancel_numerals;
 
 in
--- a/src/HOL/IsaMakefile	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/IsaMakefile	Sun Feb 15 10:46:37 2004 +0100
@@ -138,10 +138,9 @@
 
 $(OUT)/HOL-Complex: $(OUT)/HOL Complex/ROOT.ML\
   Library/Zorn.thy\
-  Real/Lubs.thy Real/rat_arith.ML Real/RatArith.thy\
+  Real/Lubs.thy Real/rat_arith.ML\
   Real/Rational.thy Real/PReal.thy Real/RComplete.thy \
-  Real/ROOT.ML Real/Real.thy \
-  Real/RealArith.thy Real/real_arith.ML Real/RealDef.thy \
+  Real/ROOT.ML Real/Real.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\
@@ -161,13 +160,9 @@
   Complex/Complex_Main.thy\
   Complex/CLim.ML Complex/CLim.thy\
   Complex/CSeries.ML Complex/CSeries.thy\
-  Complex/CStar.ML Complex/CStar.thy Complex/Complex.thy\
-  Complex/ComplexArith0.ML Complex/ComplexArith0.thy\
-  Complex/ComplexBin.ML Complex/ComplexBin.thy\
+  Complex/CStar.ML Complex/CStar.thy Complex/Complex.thy Complex/ComplexBin.thy\
   Complex/NSCA.ML Complex/NSCA.thy\
-  Complex/NSComplex.thy\
-  Complex/hcomplex_arith.ML Complex/NSComplexArith.thy\
-  Complex/NSComplexBin.ML Complex/NSComplexBin.thy
+  Complex/NSComplex.thy
 	@cd Complex; $(ISATOOL) usedir -b $(OUT)/HOL HOL-Complex
 
 
--- a/src/HOL/NumberTheory/Int2.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/NumberTheory/Int2.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -62,7 +62,7 @@
     by (auto simp add: zmod_zdiv_equality [THEN sym] zmult_ac)
   also assume  "x < y * z";
   finally show ?thesis;
-    by (auto simp add: prems zmult_zless_cancel2, insert prems, arith)
+    by (auto simp add: prems mult_less_cancel_right, insert prems, arith)
 qed;
 
 lemma div_prop2: "[| 0 < z; (x::int) < (y * z) + z |] ==> x div z \<le> y";
--- a/src/HOL/NumberTheory/IntPrimes.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/NumberTheory/IntPrimes.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -124,8 +124,8 @@
   -- {* addition is an AC-operator *}
 
 lemma zgcd_zmult_distrib2: "0 \<le> k ==> k * zgcd (m, n) = zgcd (k * m, k * n)"
-  by (simp del: zmult_zminus_right
-      add: zmult_zminus_right [symmetric] nat_mult_distrib zgcd_def zabs_def
+  by (simp del: minus_mult_right [symmetric]
+      add: minus_mult_right nat_mult_distrib zgcd_def zabs_def
           mult_less_0_iff gcd_mult_distrib2 [symmetric] zmult_int [symmetric])
 
 lemma zgcd_zmult_distrib2_abs: "zgcd (k * m, k * n) = abs k * zgcd (m, n)"
@@ -365,7 +365,7 @@
   apply (subgoal_tac "0 < m")
    apply (simp add: zero_le_mult_iff)
    apply (subgoal_tac "m * k < m * 1")
-    apply (drule zmult_zless_cancel1 [THEN iffD1])
+    apply (drule mult_less_cancel_left [THEN iffD1])
     apply (auto simp add: linorder_neq_iff)
   done
 
--- a/src/HOL/NumberTheory/Quadratic_Reciprocity.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/NumberTheory/Quadratic_Reciprocity.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -320,7 +320,7 @@
   proof -;
     assume "b \<le> q * a div p";
     then have "p * b \<le> p * ((q * a) div p)";
-      by (insert p_g_2, auto simp add: zmult_zle_cancel1)
+      by (insert p_g_2, auto simp add: mult_le_cancel_left)
     also have "... \<le> q * a";
       by (rule zdiv_leq_prop, insert p_g_2, auto)
     finally have "p * b \<le> q * a" .;
@@ -353,7 +353,7 @@
   proof -;
     assume "a \<le> p * b div q";
     then have "q * a \<le> q * ((p * b) div q)";
-      by (insert q_g_2, auto simp add: zmult_zle_cancel1)
+      by (insert q_g_2, auto simp add: mult_le_cancel_left)
     also have "... \<le> p * b";
       by (rule zdiv_leq_prop, insert q_g_2, auto)
     finally have "q * a \<le> p * b" .;
@@ -425,7 +425,7 @@
         assume "0 < x" and "x \<le> q * j div p";
         with j_fact P_set_def  have "j \<le> (p - 1) div 2"; by auto
         with q_g_2; have "q * j \<le> q * ((p - 1) div 2)";
-          by (auto simp add: zmult_zle_cancel1)
+          by (auto simp add: mult_le_cancel_left)
         with p_g_2 have "q * j div p \<le> q * ((p - 1) div 2) div p";
           by (auto simp add: zdiv_mono1)
         also from prems have "... \<le> (q - 1) div 2";
@@ -437,7 +437,7 @@
   also have "... = (q * j) div p";
   proof -;
     from j_fact P_set_def have "0 \<le> j" by auto
-    with q_g_2 have "q * 0 \<le> q * j" by (auto simp only: zmult_zle_mono2)
+    with q_g_2 have "q * 0 \<le> q * j" by (auto simp only: mult_left_mono)
     then have "0 \<le> q * j" by auto
     then have "0 div p \<le> (q * j) div p";
       apply (rule_tac a = 0 in zdiv_mono1)
@@ -478,7 +478,7 @@
         assume "0 < x" and "x \<le> p * j div q";
         with j_fact Q_set_def  have "j \<le> (q - 1) div 2"; by auto
         with p_g_2; have "p * j \<le> p * ((q - 1) div 2)";
-          by (auto simp add: zmult_zle_cancel1)
+          by (auto simp add: mult_le_cancel_left)
         with q_g_2 have "p * j div q \<le> p * ((q - 1) div 2) div q";
           by (auto simp add: zdiv_mono1)
         also from prems have "... \<le> (p - 1) div 2";
@@ -490,7 +490,7 @@
   also have "... = (p * j) div q";
   proof -;
     from j_fact Q_set_def have "0 \<le> j" by auto
-    with p_g_2 have "p * 0 \<le> p * j" by (auto simp only: zmult_zle_mono2)
+    with p_g_2 have "p * 0 \<le> p * j" by (auto simp only: mult_left_mono)
     then have "0 \<le> p * j" by auto
     then have "0 div q \<le> (p * j) div q";
       apply (rule_tac a = 0 in zdiv_mono1)
--- a/src/HOL/Numeral.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Numeral.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -8,14 +8,24 @@
 theory Numeral = Datatype
 files "Tools/numeral_syntax.ML":
 
-(* The constructors Pls/Min are hidden in numeral_syntax.ML.
-   Only qualified access bin.Pls/Min is allowed.
-   Should also hide Bit, but that means one cannot use BIT anymore.
-*)
+text{* The file @{text numeral_syntax.ML} hides the constructors Pls and Min.
+   Only qualified access bin.Pls and bin.Min is allowed.
+   We do not hide Bit because we need the BIT infix syntax.*}
+
+text{*A number can have multiple representations, namely leading Falses with
+sign @{term Pls} and leading Trues with sign @{term Min}.
+See @{text "ZF/Integ/twos-compl.ML"}, function @{text int_of_binary},
+for the numerical interpretation.
+
+The representation expects that @{text "(m mod 2)"} is 0 or 1,
+even if m is negative;
+For instance, @{text "-5 div 2 = -3"} and @{text "-5 mod 2 = 1"}; thus
+@{text "-5 = (-3)*2 + 1"}.
+*}
 
 datatype
-  bin = Pls
-      | Min
+  bin = Pls  --{*Plus: Stands for an infinite string of leading Falses*}
+      | Min --{*Minus: Stands for an infinite string of leading Trues*}
       | Bit bin bool    (infixl "BIT" 90)
 
 axclass
@@ -58,4 +68,136 @@
   by (simp add: Let_def)
 
 
+consts
+  ring_of :: "bin => 'a::ring"
+
+  NCons     :: "[bin,bool]=>bin"
+  bin_succ  :: "bin=>bin"
+  bin_pred  :: "bin=>bin"
+  bin_minus :: "bin=>bin"
+  bin_add   :: "[bin,bin]=>bin"
+  bin_mult  :: "[bin,bin]=>bin"
+
+text{*@{term NCons} inserts a bit, suppressing leading 0s and 1s*}
+primrec
+  NCons_Pls:  "NCons bin.Pls b = (if b then (bin.Pls BIT b) else bin.Pls)"
+  NCons_Min:  "NCons bin.Min b = (if b then bin.Min else (bin.Min BIT b))"
+  NCons_BIT:  "NCons (w BIT x) b = (w BIT x) BIT b"
+
+
+primrec 
+  ring_of_Pls: "ring_of bin.Pls = 0"
+  ring_of_Min: "ring_of bin.Min = - (1::'a::ring)"
+  ring_of_BIT: "ring_of(w BIT x) = (if x then 1 else 0) +
+	                               (ring_of w) + (ring_of w)"
+
+primrec
+  bin_succ_Pls: "bin_succ bin.Pls = bin.Pls BIT True"
+  bin_succ_Min: "bin_succ bin.Min = bin.Pls"
+  bin_succ_BIT: "bin_succ(w BIT x) =
+  	            (if x then bin_succ w BIT False
+	                  else NCons w True)"
+
+primrec
+  bin_pred_Pls: "bin_pred bin.Pls = bin.Min"
+  bin_pred_Min: "bin_pred bin.Min = bin.Min BIT False"
+  bin_pred_BIT: "bin_pred(w BIT x) =
+	            (if x then NCons w False
+		          else (bin_pred w) BIT True)"
+
+primrec
+  bin_minus_Pls: "bin_minus bin.Pls = bin.Pls"
+  bin_minus_Min: "bin_minus bin.Min = bin.Pls BIT True"
+  bin_minus_BIT: "bin_minus(w BIT x) =
+	             (if x then bin_pred (NCons (bin_minus w) False)
+		           else bin_minus w BIT False)"
+
+primrec
+  bin_add_Pls: "bin_add bin.Pls w = w"
+  bin_add_Min: "bin_add bin.Min w = bin_pred w"
+  bin_add_BIT:
+    "bin_add (v BIT x) w =
+       (case w of Pls => v BIT x
+                | Min => bin_pred (v BIT x)
+                | (w BIT y) =>
+      	            NCons (bin_add v (if (x & y) then bin_succ w else w))
+	                  (x~=y))"
+
+primrec
+  bin_mult_Pls: "bin_mult bin.Pls w = bin.Pls"
+  bin_mult_Min: "bin_mult bin.Min w = bin_minus w"
+  bin_mult_BIT: "bin_mult (v BIT x) w =
+	            (if x then (bin_add (NCons (bin_mult v w) False) w)
+	                  else (NCons (bin_mult v w) False))"
+
+
+subsection{*Extra rules for @{term bin_succ}, @{term bin_pred}, 
+  @{term bin_add} and @{term bin_mult}*}
+
+lemma NCons_Pls_0: "NCons bin.Pls False = bin.Pls"
+by simp
+
+lemma NCons_Pls_1: "NCons bin.Pls True = bin.Pls BIT True"
+by simp
+
+lemma NCons_Min_0: "NCons bin.Min False = bin.Min BIT False"
+by simp
+
+lemma NCons_Min_1: "NCons bin.Min True = bin.Min"
+by simp
+
+lemma bin_succ_1: "bin_succ(w BIT True) = (bin_succ w) BIT False"
+by simp
+
+lemma bin_succ_0: "bin_succ(w BIT False) =  NCons w True"
+by simp
+
+lemma bin_pred_1: "bin_pred(w BIT True) = NCons w False"
+by simp
+
+lemma bin_pred_0: "bin_pred(w BIT False) = (bin_pred w) BIT True"
+by simp
+
+lemma bin_minus_1: "bin_minus(w BIT True) = bin_pred (NCons (bin_minus w) False)"
+by simp
+
+lemma bin_minus_0: "bin_minus(w BIT False) = (bin_minus w) BIT False"
+by simp
+
+
+subsection{*Binary Addition and Multiplication:
+         @{term bin_add} and @{term bin_mult}*}
+
+lemma bin_add_BIT_11:
+     "bin_add (v BIT True) (w BIT True) =
+     NCons (bin_add v (bin_succ w)) False"
+by simp
+
+lemma bin_add_BIT_10:
+     "bin_add (v BIT True) (w BIT False) = NCons (bin_add v w) True"
+by simp
+
+lemma bin_add_BIT_0:
+     "bin_add (v BIT False) (w BIT y) = NCons (bin_add v w) y"
+by auto
+
+lemma bin_add_Pls_right: "bin_add w bin.Pls = w"
+by (induct_tac "w", auto)
+
+lemma bin_add_Min_right: "bin_add w bin.Min = bin_pred w"
+by (induct_tac "w", auto)
+
+lemma bin_add_BIT_BIT:
+     "bin_add (v BIT x) (w BIT y) =
+     NCons(bin_add v (if x & y then (bin_succ w) else w)) (x~= y)"
+by simp
+
+lemma bin_mult_1:
+     "bin_mult (v BIT True) w = bin_add (NCons (bin_mult v w) False) w"
+by simp
+
+lemma bin_mult_0: "bin_mult (v BIT False) w = NCons (bin_mult v w) False"
+by simp
+
+
 end
--- a/src/HOL/ROOT.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/ROOT.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -29,7 +29,6 @@
 use "~~/src/Provers/Arith/abel_cancel.ML";
 use "~~/src/Provers/Arith/assoc_fold.ML";
 use "~~/src/Provers/quantifier1.ML";
-use "~~/src/Provers/Arith/abstract_numerals.ML";
 use "~~/src/Provers/Arith/cancel_numerals.ML";
 use "~~/src/Provers/Arith/combine_numerals.ML";
 use "~~/src/Provers/Arith/cancel_numeral_factor.ML";
--- a/src/HOL/Real/PReal.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Real/PReal.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -7,7 +7,7 @@
                   provides some of the definitions.
 *)
 
-theory PReal = RatArith:
+theory PReal = Rational:
 
 text{*Could be generalized and moved to @{text Ring_and_Field}*}
 lemma add_eq_exists: "\<exists>x. a+x = (b::rat)"
--- a/src/HOL/Real/RComplete.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Real/RComplete.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -8,11 +8,10 @@
 
 header{*Completeness Theorems for Positive Reals and Reals.*}
 
-theory RComplete = Lubs + RealArith:
+theory RComplete = Lubs + RealDef:
 
 lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
-apply (simp)
-done
+by simp
 
 
 subsection{*Completeness of Reals by Supremum Property of type @{typ preal}*} 
@@ -32,8 +31,7 @@
 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) 
+apply (drule real_gt_zero_preal_Ex [THEN iffD1], force) 
 done
 
 (*-------------------------------------------------------------
@@ -55,12 +53,10 @@
 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 (drule_tac [3] real_less_all_real2, auto)
 apply (rule preal_complete [THEN iffD1])
 apply (auto intro: order_less_imp_le)
-apply (frule real_gt_preal_preal_Ex)
-apply (force)
+apply (frule real_gt_preal_preal_Ex, 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])
@@ -131,8 +127,7 @@
 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 (frule posreals_complete [OF _ _ exI], blast, blast, safe)
 apply (rule_tac x = "t + X + (- 1) " in exI)
 apply (rule isLubI2)
 apply (rule_tac [2] setgeI, safe)
--- a/src/HOL/Real/RatArith.thy	Sat Feb 14 02:06:12 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,161 +0,0 @@
-(*  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 (overloaded)
-  rat_number_of_def:
-    "(number_of v :: rat) == of_int (number_of v)"
-     (*::bin=>rat         ::bin=>int*)
-
-
-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 only: rat_number_of_def of_int_add number_of_add)
-
-lemma minus_rat_number_of [simp]:
-     "- (number_of w :: rat) = number_of (bin_minus w)"
-by (simp only: rat_number_of_def of_int_minus number_of_minus)
-
-lemma diff_rat_number_of [simp]: 
-   "(number_of v :: rat) - number_of w = number_of (bin_add v (bin_minus w))"
-by (simp only: add_rat_number_of minus_rat_number_of diff_minus)
-
-lemma mult_rat_number_of [simp]:
-     "(number_of v :: rat) * number_of v' = number_of (bin_mult v v')"
-by (simp only: rat_number_of_def of_int_mult number_of_mult)
-
-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 del: rat_number_of_def 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')) :: int)"
-by (simp add: rat_number_of_def)
-
-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')) :: int)"
-by (simp add: rat_number_of_def)
-
-
-text{*New versions of existing theorems involving 0, 1*}
-
-lemma rat_minus_1_eq_m1 [simp]: "- 1 = (-1::rat)"
-by (simp del: rat_number_of_def 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 :: int)  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)
-
-declare rat_number_of_def [simp]
-
-
-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
--- a/src/HOL/Real/Rational.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Real/Rational.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -9,7 +9,8 @@
   \author{Markus Wenzel}
 *}
 
-theory Rational = Quotient + Ring_and_Field:
+theory Rational = Quotient + Main
+files ("rat_arith.ML"):
 
 subsection {* Fractions *}
 
@@ -693,4 +694,35 @@
 qed 
 
 
+
+subsection{*Numerals and Arithmetic*}
+
+instance rat :: number ..
+
+primrec (*the type constraint is essential!*)
+  number_of_Pls: "number_of bin.Pls = 0"
+  number_of_Min: "number_of bin.Min = - (1::rat)"
+  number_of_BIT: "number_of(w BIT x) = (if x then 1 else 0) +
+	                               (number_of w) + (number_of w)"
+
+declare number_of_Pls [simp del]
+        number_of_Min [simp del]
+        number_of_BIT [simp del]
+
+instance rat :: number_ring
+proof
+  show "Numeral0 = (0::rat)" by (rule number_of_Pls)
+  show "-1 = - (1::rat)" by (rule number_of_Min)
+  fix w :: bin and x :: bool
+  show "(number_of (w BIT x) :: rat) =
+        (if x then 1 else 0) + number_of w + number_of w"
+    by (rule number_of_BIT)
+qed
+
+declare diff_rat_def [symmetric]
+
+use "rat_arith.ML"
+
+setup rat_arith_setup
+
 end
--- a/src/HOL/Real/RealArith.thy	Sat Feb 14 02:06:12 2004 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,317 +0,0 @@
-(*  Title:      HOL/RealArith.thy
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1999  University of Cambridge
-*)
-
-header{*Binary arithmetic and Simplification for the Reals*}
-
-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')) :: int)"
-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')) :: int)"
-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 :: int) then 0  
-         else (number_of v :: real))"
-proof cases
-  assume "neg (number_of v :: int)" thus ?thesis by simp
-next
-  assume neg: "~ neg (number_of v :: int)"
-  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]
-
-
-use "real_arith.ML"
-
-setup real_arith_setup
-
-subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*}
-
-text{*Needed in this non-standard form by Hyperreal/Transcendental*}
-lemma real_0_le_divide_iff:
-     "((0::real) \<le> x/y) = ((x \<le> 0 | 0 \<le> y) & (0 \<le> x | y \<le> 0))"
-by (simp add: real_divide_def zero_le_mult_iff, auto)
-
-lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)" 
-by arith
-
-lemma real_add_eq_0_iff [iff]: "(x+y = (0::real)) = (y = -x)"
-by auto
-
-lemma real_add_less_0_iff [iff]: "(x+y < (0::real)) = (y < -x)"
-by auto
-
-lemma real_0_less_add_iff [iff]: "((0::real) < x+y) = (-x < y)"
-by auto
-
-lemma real_add_le_0_iff [iff]: "(x+y \<le> (0::real)) = (y \<le> -x)"
-by auto
-
-lemma real_0_le_add_iff [iff]: "((0::real) \<le> x+y) = (-x \<le> y)"
-by auto
-
-
-(** Simprules combining x-y and 0 (needed??) **)
-
-lemma real_0_less_diff_iff [iff]: "((0::real) < x-y) = (y < x)"
-by auto
-
-lemma real_0_le_diff_iff [iff]: "((0::real) \<le> x-y) = (y \<le> x)"
-by auto
-
-(*
-FIXME: we should have this, as for type int, but many proofs would break.
-It replaces x+-y by x-y.
-Addsimps [symmetric real_diff_def]
-*)
-
-subsubsection{*Division By @{term "-1"}*}
-
-lemma real_divide_minus1 [simp]: "x/-1 = -(x::real)"
-by simp
-
-lemma real_minus1_divide [simp]: "-1/(x::real) = - (1/x)"
-by (simp add: real_divide_def inverse_minus_eq)
-
-lemma real_lbound_gt_zero:
-     "[| (0::real) < d1; 0 < d2 |] ==> \<exists>e. 0 < e & e < d1 & e < d2"
-apply (rule_tac x = " (min d1 d2) /2" in exI)
-apply (simp add: min_def)
-done
-
-(*** Density of the Reals ***)
-
-text{*Similar results are proved in @{text Ring_and_Field}*}
-lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)"
-  by auto
-
-lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y"
-  by auto
-
-lemma real_dense: "x < y ==> \<exists>r::real. x < r & r < y"
-  by (rule Ring_and_Field.dense)
-
-
-subsection{*Absolute Value Function for the Reals*}
-
-lemma abs_nat_number_of [simp]: 
-     "abs (number_of v :: real) =  
-        (if neg (number_of v :: int) then number_of (bin_minus v)  
-         else number_of v)"
-by (simp add: real_abs_def bin_arith_simps minus_real_number_of
-       less_real_number_of real_of_int_le_iff)
-
-text{*FIXME: these should go!*}
-lemma abs_eqI1: "(0::real)\<le>x ==> abs x = x"
-by (unfold real_abs_def, simp)
-
-lemma abs_eqI2: "(0::real) < x ==> abs x = x"
-by (unfold real_abs_def, simp)
-
-lemma abs_minus_eqI2: "x < (0::real) ==> abs x = -x"
-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)
-
-lemma abs_minus_one [simp]: "abs (-1) = (1::real)"
-by (unfold real_abs_def, simp)
-
-lemma abs_interval_iff: "(abs x < r) = (-r < x & x < (r::real))"
-by (force simp add: Ring_and_Field.abs_less_iff)
-
-lemma abs_le_interval_iff: "(abs x \<le> r) = (-r\<le>x & x\<le>(r::real))"
-by (force simp add: Ring_and_Field.abs_le_iff)
-
-lemma abs_add_one_gt_zero [simp]: "(0::real) < 1 + abs(x)"
-by (unfold real_abs_def, auto)
-
-lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)"
-by (auto intro: abs_eqI1 simp add: real_of_nat_ge_zero)
-
-lemma abs_add_one_not_less_self [simp]: "~ abs(x) + (1::real) < x"
-apply (simp add: linorder_not_less)
-apply (auto intro: abs_ge_self [THEN order_trans])
-done
- 
-text{*Used only in Hyperreal/Lim.ML*}
-lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \<le> abs(x + -l) + abs(y + -m)"
-apply (simp add: real_add_assoc)
-apply (rule_tac a1 = y in add_left_commute [THEN ssubst])
-apply (rule real_add_assoc [THEN subst])
-apply (rule abs_triangle_ineq)
-done
-
-
-
-ML
-{*
-val real_0_le_divide_iff = thm"real_0_le_divide_iff";
-val real_add_minus_iff = thm"real_add_minus_iff";
-val real_add_eq_0_iff = thm"real_add_eq_0_iff";
-val real_add_less_0_iff = thm"real_add_less_0_iff";
-val real_0_less_add_iff = thm"real_0_less_add_iff";
-val real_add_le_0_iff = thm"real_add_le_0_iff";
-val real_0_le_add_iff = thm"real_0_le_add_iff";
-val real_0_less_diff_iff = thm"real_0_less_diff_iff";
-val real_0_le_diff_iff = thm"real_0_le_diff_iff";
-val real_divide_minus1 = thm"real_divide_minus1";
-val real_minus1_divide = thm"real_minus1_divide";
-val real_lbound_gt_zero = thm"real_lbound_gt_zero";
-val real_less_half_sum = thm"real_less_half_sum";
-val real_gt_half_sum = thm"real_gt_half_sum";
-val real_dense = thm"real_dense";
-
-val abs_nat_number_of = thm"abs_nat_number_of";
-val abs_eqI1 = thm"abs_eqI1";
-val abs_eqI2 = thm"abs_eqI2";
-val abs_minus_eqI2 = thm"abs_minus_eqI2";
-val abs_ge_zero = thm"abs_ge_zero";
-val abs_idempotent = thm"abs_idempotent";
-val abs_zero_iff = thm"abs_zero_iff";
-val abs_ge_self = thm"abs_ge_self";
-val abs_ge_minus_self = thm"abs_ge_minus_self";
-val abs_mult = thm"abs_mult";
-val abs_inverse = thm"abs_inverse";
-val abs_triangle_ineq = thm"abs_triangle_ineq";
-val abs_minus_cancel = thm"abs_minus_cancel";
-val abs_minus_add_cancel = thm"abs_minus_add_cancel";
-val abs_minus_one = thm"abs_minus_one";
-val abs_interval_iff = thm"abs_interval_iff";
-val abs_le_interval_iff = thm"abs_le_interval_iff";
-val abs_add_one_gt_zero = thm"abs_add_one_gt_zero";
-val abs_le_zero_iff = thm"abs_le_zero_iff";
-val abs_real_of_nat_cancel = thm"abs_real_of_nat_cancel";
-val abs_add_one_not_less_self = thm"abs_add_one_not_less_self";
-val abs_sum_triangle_ineq = thm"abs_sum_triangle_ineq";
-
-val abs_mult_less = thm"abs_mult_less";
-*}
-
-end
--- a/src/HOL/Real/RealDef.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Real/RealDef.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -2,10 +2,13 @@
     ID          : $Id$
     Author      : Jacques D. Fleuriot
     Copyright   : 1998  University of Cambridge
-    Description : The reals
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
 *)
 
-theory RealDef = PReal:
+header{*Defining the Reals from the Positive Reals*}
+
+theory RealDef = PReal
+files ("real_arith.ML"):
 
 constdefs
   realrel   ::  "((preal * preal) * (preal * preal)) set"
@@ -418,7 +421,7 @@
   "(Abs_REAL(realrel``{(x1,y1)}) \<le> Abs_REAL(realrel``{(x2,y2)})) =  
    (x1 + y2 \<le> x2 + y1)"
 apply (simp add: real_le_def) 
-apply (auto intro: real_le_lemma);
+apply (auto intro: real_le_lemma)
 done
 
 lemma real_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::real)"
@@ -464,8 +467,7 @@
 apply (rule eq_Abs_REAL [of z])
 apply (rule eq_Abs_REAL [of w]) 
 apply (auto simp add: real_le real_zero_def preal_add_ac preal_cancels)
-apply (cut_tac x="x+ya" and y="xa+y" in linorder_linear) 
-apply (auto ); 
+apply (cut_tac x="x+ya" and y="xa+y" in linorder_linear, auto) 
 done
 
 
@@ -514,9 +516,8 @@
 
 text{*lemma for proving @{term "0<(1::real)"}*}
 lemma real_zero_le_one: "0 \<le> (1::real)"
-apply (simp add: real_zero_def real_one_def real_le 
+by (simp add: real_zero_def real_one_def real_le 
                  preal_self_less_add_left order_less_imp_le)
-done
 
 
 subsection{*The Reals Form an Ordered Field*}
@@ -792,7 +793,7 @@
 lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)"
 by (simp add: real_of_nat_def)
 
-lemma real_of_nat_zero_iff: "(real (n::nat) = 0) = (n = 0)"
+lemma real_of_nat_zero_iff [iff]: "(real (n::nat) = 0) = (n = 0)"
 by (simp add: real_of_nat_def)
 
 lemma real_of_nat_diff: "n \<le> m ==> real (m - n) = real (m::nat) - real n"
@@ -814,8 +815,7 @@
 by (simp add: real_of_nat_def real_of_int_def int_eq_of_nat)
 
 
-
-text{*Still needed for binary arith*}
+text{*Still needed for binary arithmetic*}
 lemma real_of_nat_real_of_int: "~neg z ==> real (nat z) = real z"
 proof (simp add: not_neg_eq_ge_0 real_of_nat_def real_of_int_def)
   assume "0 \<le> z"
@@ -826,107 +826,194 @@
   finally show "of_nat (nat z) = of_int z" .
 qed
 
+
+
+subsection{*Numerals and Arithmetic*}
+
+instance real :: number ..
+
+primrec (*the type constraint is essential!*)
+  number_of_Pls: "number_of bin.Pls = 0"
+  number_of_Min: "number_of bin.Min = - (1::real)"
+  number_of_BIT: "number_of(w BIT x) = (if x then 1 else 0) +
+	                               (number_of w) + (number_of w)"
+
+declare number_of_Pls [simp del]
+        number_of_Min [simp del]
+        number_of_BIT [simp del]
+
+instance real :: number_ring
+proof
+  show "Numeral0 = (0::real)" by (rule number_of_Pls)
+  show "-1 = - (1::real)" by (rule number_of_Min)
+  fix w :: bin and x :: bool
+  show "(number_of (w BIT x) :: real) =
+        (if x then 1 else 0) + number_of w + number_of w"
+    by (rule number_of_BIT)
+qed
+
+
+text{*Collapse applications of @{term real} to @{term number_of}*}
+lemma real_number_of [simp]: "real (number_of v :: int) = number_of v"
+by (simp add:  real_of_int_def of_int_number_of_eq)
+
+lemma real_of_nat_number_of [simp]:
+     "real (number_of v :: nat) =  
+        (if neg (number_of v :: int) then 0  
+         else (number_of v :: real))"
+by (simp add: real_of_int_real_of_nat [symmetric] int_nat_number_of)
+ 
+
+use "real_arith.ML"
+
+setup real_arith_setup
+
+subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*}
+
+text{*Needed in this non-standard form by Hyperreal/Transcendental*}
+lemma real_0_le_divide_iff:
+     "((0::real) \<le> x/y) = ((x \<le> 0 | 0 \<le> y) & (0 \<le> x | y \<le> 0))"
+by (simp add: real_divide_def zero_le_mult_iff, auto)
+
+lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)" 
+by arith
+
+lemma real_add_eq_0_iff [iff]: "(x+y = (0::real)) = (y = -x)"
+by auto
+
+lemma real_add_less_0_iff [iff]: "(x+y < (0::real)) = (y < -x)"
+by auto
+
+lemma real_0_less_add_iff [iff]: "((0::real) < x+y) = (-x < y)"
+by auto
+
+lemma real_add_le_0_iff [iff]: "(x+y \<le> (0::real)) = (y \<le> -x)"
+by auto
+
+lemma real_0_le_add_iff [iff]: "((0::real) \<le> x+y) = (-x \<le> y)"
+by auto
+
+
+(** Simprules combining x-y and 0 (needed??) **)
+
+lemma real_0_less_diff_iff [iff]: "((0::real) < x-y) = (y < x)"
+by auto
+
+lemma real_0_le_diff_iff [iff]: "((0::real) \<le> x-y) = (y \<le> x)"
+by auto
+
+(*
+FIXME: we should have this, as for type int, but many proofs would break.
+It replaces x+-y by x-y.
+Addsimps [symmetric real_diff_def]
+*)
+
+
+subsubsection{*Density of the Reals*}
+
+lemma real_lbound_gt_zero:
+     "[| (0::real) < d1; 0 < d2 |] ==> \<exists>e. 0 < e & e < d1 & e < d2"
+apply (rule_tac x = " (min d1 d2) /2" in exI)
+apply (simp add: min_def)
+done
+
+
+text{*Similar results are proved in @{text Ring_and_Field}*}
+lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)"
+  by auto
+
+lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y"
+  by auto
+
+lemma real_dense: "x < y ==> \<exists>r::real. x < r & r < y"
+  by (rule Ring_and_Field.dense)
+
+
+subsection{*Absolute Value Function for the Reals*}
+
+text{*FIXME: these should go!*}
+lemma abs_eqI1: "(0::real)\<le>x ==> abs x = x"
+by (unfold real_abs_def, simp)
+
+lemma abs_eqI2: "(0::real) < x ==> abs x = x"
+by (unfold real_abs_def, simp)
+
+lemma abs_minus_eqI2: "x < (0::real) ==> abs x = -x"
+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)
+
+lemma abs_minus_one [simp]: "abs (-1) = (1::real)"
+by (unfold real_abs_def, simp)
+
+lemma abs_interval_iff: "(abs x < r) = (-r < x & x < (r::real))"
+by (force simp add: Ring_and_Field.abs_less_iff)
+
+lemma abs_le_interval_iff: "(abs x \<le> r) = (-r\<le>x & x\<le>(r::real))"
+by (force simp add: Ring_and_Field.abs_le_iff)
+
+lemma abs_add_one_gt_zero [simp]: "(0::real) < 1 + abs(x)"
+by (unfold real_abs_def, auto)
+
+lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)"
+by (auto intro: abs_eqI1 simp add: real_of_nat_ge_zero)
+
+lemma abs_add_one_not_less_self [simp]: "~ abs(x) + (1::real) < x"
+apply (simp add: linorder_not_less)
+apply (auto intro: abs_ge_self [THEN order_trans])
+done
+ 
+text{*Used only in Hyperreal/Lim.ML*}
+lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \<le> abs(x + -l) + abs(y + -m)"
+apply (simp add: real_add_assoc)
+apply (rule_tac a1 = y in add_left_commute [THEN ssubst])
+apply (rule real_add_assoc [THEN subst])
+apply (rule abs_triangle_ineq)
+done
+
+
+
 ML
 {*
-val real_abs_def = thm "real_abs_def";
-
-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 realrel_iff = thm"realrel_iff";
-val realrel_refl = thm"realrel_refl";
-val equiv_realrel = thm"equiv_realrel";
-val equiv_realrel_iff = thm"equiv_realrel_iff";
-val realrel_in_real = thm"realrel_in_real";
-val inj_on_Abs_REAL = thm"inj_on_Abs_REAL";
-val eq_realrelD = thm"eq_realrelD";
-val inj_Rep_REAL = thm"inj_Rep_REAL";
-val inj_real_of_preal = thm"inj_real_of_preal";
-val eq_Abs_REAL = thm"eq_Abs_REAL";
-val real_minus_congruent = thm"real_minus_congruent";
-val real_minus = thm"real_minus";
-val real_add = thm"real_add";
-val real_add_commute = thm"real_add_commute";
-val real_add_assoc = thm"real_add_assoc";
-val real_add_zero_left = thm"real_add_zero_left";
-val real_add_zero_right = thm"real_add_zero_right";
+val real_0_le_divide_iff = thm"real_0_le_divide_iff";
+val real_add_minus_iff = thm"real_add_minus_iff";
+val real_add_eq_0_iff = thm"real_add_eq_0_iff";
+val real_add_less_0_iff = thm"real_add_less_0_iff";
+val real_0_less_add_iff = thm"real_0_less_add_iff";
+val real_add_le_0_iff = thm"real_add_le_0_iff";
+val real_0_le_add_iff = thm"real_0_le_add_iff";
+val real_0_less_diff_iff = thm"real_0_less_diff_iff";
+val real_0_le_diff_iff = thm"real_0_le_diff_iff";
+val real_lbound_gt_zero = thm"real_lbound_gt_zero";
+val real_less_half_sum = thm"real_less_half_sum";
+val real_gt_half_sum = thm"real_gt_half_sum";
+val real_dense = thm"real_dense";
 
-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 preal_le_linear = thm"preal_le_linear";
-val real_mult_inverse_left = thm"real_mult_inverse_left";
-val real_not_refl2 = thm"real_not_refl2";
-val real_of_preal_add = thm"real_of_preal_add";
-val real_of_preal_mult = thm"real_of_preal_mult";
-val real_of_preal_trichotomy = thm"real_of_preal_trichotomy";
-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_le_imp_less_or_eq = thm"real_le_imp_less_or_eq";
-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_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";
-val real_less_all_preal = thm "real_less_all_preal";
-val real_less_all_real2 = thm "real_less_all_real2";
-val real_of_preal_le_iff = thm "real_of_preal_le_iff";
-val real_mult_order = thm "real_mult_order";
-val real_zero_less_one = thm "real_zero_less_one";
-val real_add_less_le_mono = thm "real_add_less_le_mono";
-val real_add_le_less_mono = thm "real_add_le_less_mono";
-val real_add_order = thm "real_add_order";
-val real_le_add_order = thm "real_le_add_order";
-val real_le_square = thm "real_le_square";
-val real_mult_less_mono2 = thm "real_mult_less_mono2";
+val abs_eqI1 = thm"abs_eqI1";
+val abs_eqI2 = thm"abs_eqI2";
+val abs_minus_eqI2 = thm"abs_minus_eqI2";
+val abs_ge_zero = thm"abs_ge_zero";
+val abs_idempotent = thm"abs_idempotent";
+val abs_zero_iff = thm"abs_zero_iff";
+val abs_ge_self = thm"abs_ge_self";
+val abs_ge_minus_self = thm"abs_ge_minus_self";
+val abs_mult = thm"abs_mult";
+val abs_inverse = thm"abs_inverse";
+val abs_triangle_ineq = thm"abs_triangle_ineq";
+val abs_minus_cancel = thm"abs_minus_cancel";
+val abs_minus_add_cancel = thm"abs_minus_add_cancel";
+val abs_minus_one = thm"abs_minus_one";
+val abs_interval_iff = thm"abs_interval_iff";
+val abs_le_interval_iff = thm"abs_le_interval_iff";
+val abs_add_one_gt_zero = thm"abs_add_one_gt_zero";
+val abs_le_zero_iff = thm"abs_le_zero_iff";
+val abs_real_of_nat_cancel = thm"abs_real_of_nat_cancel";
+val abs_add_one_not_less_self = thm"abs_add_one_not_less_self";
+val abs_sum_triangle_ineq = thm"abs_sum_triangle_ineq";
 
-val real_mult_less_iff1 = thm "real_mult_less_iff1";
-val real_mult_le_cancel_iff1 = thm "real_mult_le_cancel_iff1";
-val real_mult_le_cancel_iff2 = thm "real_mult_le_cancel_iff2";
-val real_mult_less_mono = thm "real_mult_less_mono";
-val real_mult_less_mono' = thm "real_mult_less_mono'";
-val real_sum_squares_cancel = thm "real_sum_squares_cancel";
-val real_sum_squares_cancel2 = thm "real_sum_squares_cancel2";
-
-val real_mult_left_cancel = thm"real_mult_left_cancel";
-val real_mult_right_cancel = thm"real_mult_right_cancel";
-val real_inverse_unique = thm "real_inverse_unique";
-val real_inverse_gt_one = thm "real_inverse_gt_one";
-
-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_real_of_nat = thm"real_of_int_real_of_nat";
-val real_of_int_inject = thm"real_of_int_inject";
-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";
-val real_of_nat_Suc = thm "real_of_nat_Suc";
-val real_of_nat_less_iff = thm "real_of_nat_less_iff";
-val real_of_nat_le_iff = thm "real_of_nat_le_iff";
-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_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";
-val real_of_nat_ge_zero_cancel_iff = thm "real_of_nat_ge_zero_cancel_iff";
+val abs_mult_less = thm"abs_mult_less";
 *}
 
+
 end
--- a/src/HOL/Real/RealPow.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Real/RealPow.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -6,7 +6,7 @@
 
 *)
 
-theory RealPow = RealArith:
+theory RealPow = RealDef:
 
 declare abs_mult_self [simp]
 
@@ -60,7 +60,7 @@
 lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n"
 apply (induct_tac "n")
 apply (auto simp add: real_of_nat_Suc)
-apply (subst real_mult_2)
+apply (subst mult_2)
 apply (rule real_add_less_le_mono)
 apply (auto simp add: two_realpow_ge_one)
 done
@@ -137,13 +137,13 @@
 
 lemma real_of_int_power: "real (x::int) ^ n = real (x ^ n)"
 apply (induct_tac "n")
-apply (simp_all (no_asm_simp) add: nat_mult_distrib)
+apply (simp_all add: nat_mult_distrib)
 done
 declare real_of_int_power [symmetric, simp]
 
 lemma power_real_number_of:
      "(number_of v :: real) ^ n = real ((number_of v :: int) ^ n)"
-by (simp only: real_number_of_def real_of_int_power)
+by (simp only: real_number_of [symmetric] real_of_int_power)
 
 declare power_real_number_of [of _ "number_of w", standard, simp]
 
@@ -254,6 +254,7 @@
 apply (auto simp add: realpow_num_eq_if)
 done
 
+(*???generalize the type!*)
 lemma zero_le_x_squared [simp]: "(0::real) \<le> x^2"
 by (simp add: power2_eq_square)
 
--- a/src/HOL/Real/rat_arith.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Real/rat_arith.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -13,352 +13,9 @@
     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_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;
+ read_instantiate_sg(sign_of (the_context())) [("a","?a::rat")] mult_left_mono;
 
-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];
+	val le_number_of_eq = thm"le_number_of_eq";
 
 
 (****Common factor cancellation****)
@@ -373,32 +30,33 @@
 and d = gcd(m,m') and n=m/d and n'=m'/d.
 *)
 
-local
-  open Rat_Numeral_Simprocs
+val rel_number_of = [eq_number_of_eq, less_number_of_eq_neg, le_number_of_eq]
+
+local open Int_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))
+     ALLGOALS (simp_tac (HOL_ss addsimps minus_from_mult_simps @ mult_1s))
+     THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@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
+         ALLGOALS (simp_tac (HOL_ss addsimps rel_number_of@bin_simps))
+  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];
   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 dest_bal = HOLogic.dest_bin "HOL.divide" Term.dummyT
   val cancel = mult_divide_cancel_left RS trans
   val neg_exchanges = false
 )
@@ -407,8 +65,8 @@
  (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 dest_bal = HOLogic.dest_bin "op =" Term.dummyT
+  val cancel = field_mult_cancel_left RS trans
   val neg_exchanges = false
 )
 
@@ -416,7 +74,7 @@
  (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 dest_bal = HOLogic.dest_bin "op <" Term.dummyT
   val cancel = mult_less_cancel_left RS trans
   val neg_exchanges = true
 )
@@ -425,36 +83,41 @@
  (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 dest_bal = HOLogic.dest_bin "op <=" Term.dummyT
   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"],
+val field_cancel_numeral_factors_relations =
+  map Bin_Simprocs.prep_simproc
+   [("field_eq_cancel_numeral_factor",
+     ["(l::'a::{field,number_ring}) * m = n",
+      "(l::'a::{field,number_ring}) = m * n"],
      EqCancelNumeralFactor.proc),
-    ("ratless_cancel_numeral_factor",
-     ["(l::rat) * m < n", "(l::rat) < m * n"],
+    ("field_less_cancel_numeral_factor",
+     ["(l::'a::{ordered_field,number_ring}) * m < n",
+      "(l::'a::{ordered_field,number_ring}) < m * n"],
      LessCancelNumeralFactor.proc),
-    ("ratle_cancel_numeral_factor",
-     ["(l::rat) * m <= n", "(l::rat) <= m * n"],
+    ("field_le_cancel_numeral_factor",
+     ["(l::'a::{ordered_field,number_ring}) * m <= n",
+      "(l::'a::{ordered_field,number_ring}) <= 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)"],
+val field_cancel_numeral_factors_divide = 
+    Bin_Simprocs.prep_simproc
+        ("field_cancel_numeral_factor",
+         ["((l::'a::{field,number_ring}) * m) / n",
+          "(l::'a::{field,number_ring}) / (m * n)",
+          "((number_of v)::'a::{field,number_ring}) / (number_of w)"],
          DivCancelNumeralFactor.proc)
 
-val rat_cancel_numeral_factors =
-    rat_cancel_numeral_factors_relations @
-    [rat_cancel_numeral_factors_divide]
+val field_cancel_numeral_factors =
+    field_cancel_numeral_factors_relations @
+    [field_cancel_numeral_factors_divide]
 
 end;
 
-Addsimprocs rat_cancel_numeral_factors;
+Addsimprocs field_cancel_numeral_factors;
 
 
 (*examples:
@@ -497,8 +160,7 @@
 
 (** Declarations for ExtractCommonTerm **)
 
-local
-  open Rat_Numeral_Simprocs
+local open Int_Numeral_Simprocs
 in
 
 structure CancelFactorCommon =
@@ -512,32 +174,39 @@
   val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@mult_ac))
   end;
 
+(*This version works for all fields, including unordered ones (complex).
+  The version declared in int_factor_simprocs.ML is for integers.*)
 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
+  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
+  val simplify_meta_eq  = cancel_simplify_meta_eq field_mult_cancel_left
 );
 
 
+(*This version works for fields, with the generic divides operator (/).
+  The version declared in int_factor_simprocs.ML for integers with div.*)
 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 dest_bal = HOLogic.dest_bin "HOL.divide" Term.dummyT
   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)"],
+val field_cancel_factor =
+  map Bin_Simprocs.prep_simproc
+   [("field_eq_cancel_factor",
+     ["(l::'a::field) * m = n", "(l::'a::field) = m * n"], 
+     EqCancelFactor.proc),
+    ("field_divide_cancel_factor",
+     ["((l::'a::field) * m) / n", "(l::'a::field) / (m * n)"],
      DivideCancelFactor.proc)];
 
 end;
 
-Addsimprocs rat_cancel_factor;
+Addsimprocs field_cancel_factor;
 
 
 (*examples:
@@ -563,22 +232,12 @@
 
 
 
-(****Instantiation of the generic linear arithmetic package****)
+(****Instantiation of the generic linear arithmetic package for fields****)
 
 
 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];
-
-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 simprocs = [field_cancel_numeral_factors_divide]
 
 val mono_ss = simpset() addsimps
                 [add_mono,add_strict_mono,add_less_le_mono,add_le_less_mono];
@@ -600,15 +259,17 @@
   (rat_mult_left_mono,
    cvar(rat_mult_left_mono, hd(tl(prems_of rat_mult_left_mono))))]
 
-val simps = [True_implies_equals,
+val simps = [order_less_irrefl, True_implies_equals,
              inst "a" "(number_of ?v)" right_distrib,
-             divide_1, times_divide_eq_right, times_divide_eq_left,
+             divide_1, divide_zero_left,
+             times_divide_eq_right, times_divide_eq_left,
 	     of_int_0, of_int_1, of_int_add, of_int_minus, of_int_diff,
-	     of_int_mult, of_int_of_nat_eq, rat_number_of_def];
+	     of_int_mult, of_int_of_nat_eq];
 
 in
 
-val fast_rat_arith_simproc = Simplifier.simproc (Theory.sign_of(the_context()))
+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;
 
@@ -618,21 +279,19 @@
 val int_inj_thms = [of_int_le_iff RS iffD2, of_int_less_iff RS iffD2,
                     of_int_eq_iff RS iffD2];
 
+val ratT = Type("Rational.rat", []);
+
 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_ordered_field,
     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
+    simpset = simpset addsimps simps
                       addsimprocs simprocs}),
-  arith_inj_const("IntDef.of_nat", HOLogic.natT --> Rat_Numeral_Simprocs.ratT),
-  arith_inj_const("IntDef.of_int", HOLogic.intT --> Rat_Numeral_Simprocs.ratT),
+  arith_inj_const("IntDef.of_nat", HOLogic.natT --> ratT),
+  arith_inj_const("IntDef.of_int", HOLogic.intT --> ratT),
   arith_discrete ("Rational.rat",false),
   Simplifier.change_simpset_of (op addsimprocs) [fast_rat_arith_simproc]];
 
-
 end;
-
-
--- a/src/HOL/Real/real_arith.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Real/real_arith.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -12,598 +12,111 @@
 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;
+val real_abs_def = thm "real_abs_def";
 
-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];
+val real_le_def = thm "real_le_def";
+val real_diff_def = thm "real_diff_def";
+val real_divide_def = thm "real_divide_def";
 
-(*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)
+val realrel_iff = thm"realrel_iff";
+val realrel_refl = thm"realrel_refl";
+val equiv_realrel = thm"equiv_realrel";
+val equiv_realrel_iff = thm"equiv_realrel_iff";
+val realrel_in_real = thm"realrel_in_real";
+val inj_on_Abs_REAL = thm"inj_on_Abs_REAL";
+val eq_realrelD = thm"eq_realrelD";
+val inj_Rep_REAL = thm"inj_Rep_REAL";
+val inj_real_of_preal = thm"inj_real_of_preal";
+val eq_Abs_REAL = thm"eq_Abs_REAL";
+val real_minus_congruent = thm"real_minus_congruent";
+val real_minus = thm"real_minus";
+val real_add = thm"real_add";
+val real_add_commute = thm"real_add_commute";
+val real_add_assoc = thm"real_add_assoc";
+val real_add_zero_left = thm"real_add_zero_left";
+val real_add_zero_right = thm"real_add_zero_right";
 
-(*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****)
-
-(*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 Real_Numeral_Simprocs
-in
-
-val rel_real_number_of = [eq_real_number_of, less_real_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 real_minus_from_mult_simps @ mult_1s))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@real_mult_minus_simps))
-     THEN ALLGOALS (simp_tac (HOL_ss addsimps mult_ac))
-  val numeral_simp_tac  =
-         ALLGOALS (simp_tac (HOL_ss addsimps rel_real_number_of@bin_simps))
-  val simplify_meta_eq  = simplify_meta_eq
-  end
-
-structure DivCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Bin_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binop "HOL.divide"
-  val dest_bal = HOLogic.dest_bin "HOL.divide" HOLogic.realT
-  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 =" HOLogic.realT
-  val cancel = mult_cancel_left RS trans
-  val neg_exchanges = false
-)
+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 preal_le_linear = thm"preal_le_linear";
+val real_mult_inverse_left = thm"real_mult_inverse_left";
+val real_not_refl2 = thm"real_not_refl2";
+val real_of_preal_add = thm"real_of_preal_add";
+val real_of_preal_mult = thm"real_of_preal_mult";
+val real_of_preal_trichotomy = thm"real_of_preal_trichotomy";
+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_le_imp_less_or_eq = thm"real_le_imp_less_or_eq";
+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_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";
+val real_less_all_preal = thm "real_less_all_preal";
+val real_less_all_real2 = thm "real_less_all_real2";
+val real_of_preal_le_iff = thm "real_of_preal_le_iff";
+val real_mult_order = thm "real_mult_order";
+val real_zero_less_one = thm "real_zero_less_one";
+val real_add_less_le_mono = thm "real_add_less_le_mono";
+val real_add_le_less_mono = thm "real_add_le_less_mono";
+val real_add_order = thm "real_add_order";
+val real_le_add_order = thm "real_le_add_order";
+val real_le_square = thm "real_le_square";
+val real_mult_less_mono2 = thm "real_mult_less_mono2";
 
-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 <" HOLogic.realT
-  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 <=" HOLogic.realT
-  val cancel = mult_le_cancel_left RS trans
-  val neg_exchanges = true
-)
-
-val real_cancel_numeral_factors_relations =
-  map prep_simproc
-   [("realeq_cancel_numeral_factor",
-     ["(l::real) * m = n", "(l::real) = m * n"],
-     EqCancelNumeralFactor.proc),
-    ("realless_cancel_numeral_factor",
-     ["(l::real) * m < n", "(l::real) < m * n"],
-     LessCancelNumeralFactor.proc),
-    ("realle_cancel_numeral_factor",
-     ["(l::real) * m <= n", "(l::real) <= m * n"],
-     LeCancelNumeralFactor.proc)]
-
-val real_cancel_numeral_factors_divide = prep_simproc
-        ("realdiv_cancel_numeral_factor",
-         ["((l::real) * m) / n", "(l::real) / (m * n)",
-          "((number_of v)::real) / (number_of w)"],
-         DivCancelNumeralFactor.proc)
+val real_mult_less_iff1 = thm "real_mult_less_iff1";
+val real_mult_le_cancel_iff1 = thm "real_mult_le_cancel_iff1";
+val real_mult_le_cancel_iff2 = thm "real_mult_le_cancel_iff2";
+val real_mult_less_mono = thm "real_mult_less_mono";
+val real_mult_less_mono' = thm "real_mult_less_mono'";
+val real_sum_squares_cancel = thm "real_sum_squares_cancel";
+val real_sum_squares_cancel2 = thm "real_sum_squares_cancel2";
 
-val real_cancel_numeral_factors =
-    real_cancel_numeral_factors_relations @
-    [real_cancel_numeral_factors_divide]
-
-end;
-
-Addsimprocs real_cancel_numeral_factors;
-
-
-(*examples:
-print_depth 22;
-set timing;
-set trace_simp;
-fun test s = (Goal s; by (Simp_tac 1));
-
-test "0 <= (y::real) * -2";
-test "9*x = 12 * (y::real)";
-test "(9*x) / (12 * (y::real)) = z";
-test "9*x < 12 * (y::real)";
-test "9*x <= 12 * (y::real)";
-
-test "-99*x = 132 * (y::real)";
-test "(-99*x) / (132 * (y::real)) = z";
-test "-99*x < 132 * (y::real)";
-test "-99*x <= 132 * (y::real)";
-
-test "999*x = -396 * (y::real)";
-test "(999*x) / (-396 * (y::real)) = z";
-test "999*x < -396 * (y::real)";
-test "999*x <= -396 * (y::real)";
-
-test  "(- ((2::real) * x) <= 2 * y)";
-test "-99*x = -81 * (y::real)";
-test "(-99*x) / (-81 * (y::real)) = z";
-test "-99*x <= -81 * (y::real)";
-test "-99*x < -81 * (y::real)";
+val real_mult_left_cancel = thm"real_mult_left_cancel";
+val real_mult_right_cancel = thm"real_mult_right_cancel";
+val real_inverse_unique = thm "real_inverse_unique";
+val real_inverse_gt_one = thm "real_inverse_gt_one";
 
-test "-2 * x = -1 * (y::real)";
-test "-2 * x = -(y::real)";
-test "(-2 * x) / (-1 * (y::real)) = z";
-test "-2 * x < -(y::real)";
-test "-2 * x <= -1 * (y::real)";
-test "-x < -23 * (y::real)";
-test "-x <= -23 * (y::real)";
-*)
-
-
-(** Declarations for ExtractCommonTerm **)
-
-local
-  open Real_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 =" HOLogic.realT
-  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" HOLogic.realT
-  val simplify_meta_eq  = cancel_simplify_meta_eq mult_divide_cancel_eq_if
-);
-
-val real_cancel_factor =
-  map prep_simproc
-   [("real_eq_cancel_factor", ["(l::real) * m = n", "(l::real) = m * n"], EqCancelFactor.proc),
-    ("real_divide_cancel_factor", ["((l::real) * m) / n", "(l::real) / (m * n)"],
-     DivideCancelFactor.proc)];
-
-end;
-
-Addsimprocs real_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::real)";
-test "k = k*(y::real)";
-test "a*(b*c) = (b::real)";
-test "a*(b*c) = d*(b::real)*(x*a)";
-
-
-test "(x*k) / (k*(y::real)) = (uu::real)";
-test "(k) / (k*(y::real)) = (uu::real)";
-test "(a*(b*c)) / ((b::real)) = (uu::real)";
-test "(a*(b*c)) / (d*(b::real)*(x*a)) = (uu::real)";
-
-(*FIXME: what do we do about this?*)
-test "a*(b*c)/(y*z) = d*(b::real)*(x*a)/z";
-*)
-
+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_real_of_nat = thm"real_of_int_real_of_nat";
+val real_of_int_inject = thm"real_of_int_inject";
+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";
+val real_of_nat_Suc = thm "real_of_nat_Suc";
+val real_of_nat_less_iff = thm "real_of_nat_less_iff";
+val real_of_nat_le_iff = thm "real_of_nat_le_iff";
+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_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";
+val real_of_nat_ge_zero_cancel_iff = thm "real_of_nat_ge_zero_cancel_iff";
 
 
 (****Instantiation of the generic linear arithmetic package****)
 
 local
 
-(* reduce contradictory <= to False *)
-val add_rules = 
-    [real_numeral_0_eq_0, real_numeral_1_eq_1,
-     add_real_number_of, minus_real_number_of, diff_real_number_of,
-     mult_real_number_of, eq_real_number_of, less_real_number_of];
-
-val simprocs = [Real_Times_Assoc.conv, Real_Numeral_Simprocs.combine_numerals,
-                real_cancel_numeral_factors_divide]@
-               Real_Numeral_Simprocs.cancel_numerals @
-               Real_Numeral_Simprocs.eval_numerals;
-
 fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
 
 val real_mult_mono_thms =
@@ -615,7 +128,7 @@
 val simps = [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, symmetric real_number_of_def];
+       real_of_int_mult RS sym];
 
 val int_inj_thms = [real_of_int_le_iff RS iffD2, real_of_int_less_iff RS iffD2,
                     real_of_int_inject RS iffD2];
@@ -625,7 +138,8 @@
 
 in
 
-val fast_real_arith_simproc = Simplifier.simproc (Theory.sign_of (the_context ()))
+val fast_real_arith_simproc =
+ Simplifier.simproc (Theory.sign_of (the_context ()))
   "fast_real_arith" ["(m::real) < n","(m::real) <= n", "(m::real) = n"]
   Fast_Arith.lin_arith_prover;
 
@@ -635,9 +149,7 @@
     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 LA_Data_Ref.lessD: the reals are dense!*)
-    simpset = simpset addsimps add_rules
-                      addsimps simps
-                      addsimprocs simprocs}),
+    simpset = simpset addsimps simps}),
   arith_inj_const ("RealDef.real", HOLogic.natT --> HOLogic.realT),
   arith_inj_const ("RealDef.real", HOLogic.intT --> HOLogic.realT),
   arith_discrete ("RealDef.real",false),
@@ -645,7 +157,6 @@
 
 (* some thms for injection nat => real:
 real_of_nat_zero
-?zero_eq_numeral_0
 real_of_nat_add
 *)
 
--- a/src/HOL/Ring_and_Field.thy	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/Ring_and_Field.thy	Sun Feb 15 10:46:37 2004 +0100
@@ -36,7 +36,7 @@
   diff_minus: "a - b = a + (-b)"
 
 axclass ordered_semiring \<subseteq> semiring, linorder
-  zero_less_one: "0 < 1" --{*This axiom too is needed for semirings only.*}
+  zero_less_one [simp]: "0 < 1" --{*This too is needed for semirings only.*}
   add_left_mono: "a \<le> b ==> c + a \<le> c + b"
   mult_strict_left_mono: "a < b ==> 0 < c ==> c * a < c * b"
 
@@ -301,6 +301,9 @@
       "a + c \<le> b + c ==> a \<le> (b::'a::ordered_semiring)"
 by simp
 
+lemma add_increasing: "[|0\<le>a; b\<le>c|] ==> b \<le> a + (c::'a::ordered_semiring)"
+by (insert add_mono [of 0 a b c], simp)
+
 
 subsection {* Ordering Rules for Unary Minus *}
 
@@ -571,9 +574,19 @@
 lemma zero_le_square: "(0::'a::ordered_ring) \<le> a*a"
 by (simp add: zero_le_mult_iff linorder_linear) 
 
-lemma zero_le_one: "(0::'a::ordered_semiring) \<le> 1"
+text{*All three types of comparision involving 0 and 1 are covered.*}
+
+declare zero_neq_one [THEN not_sym, simp]
+
+lemma zero_le_one [simp]: "(0::'a::ordered_semiring) \<le> 1"
   by (rule zero_less_one [THEN order_less_imp_le]) 
 
+lemma not_one_le_zero [simp]: "~ (1::'a::ordered_semiring) \<le> 0"
+by (simp add: linorder_not_le zero_less_one) 
+
+lemma not_one_less_zero [simp]: "~ (1::'a::ordered_semiring) < 0"
+by (simp add: linorder_not_less zero_le_one) 
+
 
 subsection{*More Monotonicity*}
 
@@ -611,6 +624,11 @@
 apply (erule mult_left_mono, assumption)
 done
 
+lemma less_1_mult: "[| 1 < m; 1 < n |] ==> 1 < m*(n::'a::ordered_semiring)"
+apply (insert mult_strict_mono [of 1 m 1 n]) 
+apply (simp add:  order_less_trans [OF zero_less_one]); 
+done
+
 
 subsection{*Cancellation Laws for Relationships With a Common Factor*}
 
@@ -976,12 +994,20 @@
 declare minus_divide_left  [symmetric, simp]
 declare minus_divide_right [symmetric, simp]
 
+text{*Also, extract signs from products*}
+declare minus_mult_left [symmetric, simp]
+declare minus_mult_right [symmetric, simp]
+
 lemma minus_divide_divide [simp]:
      "(-a)/(-b) = a / (b::'a::{field,division_by_zero})"
 apply (case_tac "b=0", simp) 
 apply (simp add: nonzero_minus_divide_divide) 
 done
 
+lemma diff_divide_distrib:
+     "(a-b)/(c::'a::{field,division_by_zero}) = a/c - b/c"
+by (simp add: diff_minus add_divide_distrib) 
+
 
 subsection {* Ordered Fields *}
 
@@ -1732,6 +1758,8 @@
 val zero_le_square = thm"zero_le_square";
 val zero_less_one = thm"zero_less_one";
 val zero_le_one = thm"zero_le_one";
+val not_one_less_zero = thm"not_one_less_zero";
+val not_one_le_zero = thm"not_one_le_zero";
 val mult_left_mono_neg = thm"mult_left_mono_neg";
 val mult_right_mono_neg = thm"mult_right_mono_neg";
 val mult_strict_mono = thm"mult_strict_mono";
--- a/src/HOL/hologic.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/HOL/hologic.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -174,8 +174,10 @@
     Const (c, [T, T] ---> boolT) $ t $ u
   end;
 
+(*destruct the application of a binary operator. The dummyT case is a crude
+  way of handling polymorphic operators.*)
 fun dest_bin c T (tm as Const (c', Type ("fun", [T', _])) $ t $ u) =
-      if c = c' andalso T = T' then (t, u)
+      if c = c' andalso (T=T' orelse T=dummyT) then (t, u)
       else raise TERM ("dest_bin " ^ c, [tm])
   | dest_bin c _ tm = raise TERM ("dest_bin " ^ c, [tm]);
 
--- a/src/Provers/Arith/assoc_fold.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/Provers/Arith/assoc_fold.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -14,8 +14,6 @@
   val ss                : simpset       (*basic simpset of object-logtic*)
   val eq_reflection     : thm           (*object-equality to meta-equality*)
   val sg_ref            : Sign.sg_ref   (*the operator's signature*)
-  val T                 : typ           (*the operator's numeric type*)
-  val plus              : term          (*the operator being folded*)
   val add_ac            : thm list      (*AC-rewrites for plus*)
 end;
 
@@ -27,17 +25,17 @@
 
  exception Assoc_fail;
 
- fun mk_sum []  = raise Assoc_fail
-   | mk_sum tms = foldr1 (fn (x,y) => Data.plus $ x $ y) tms;
+ fun mk_sum plus []  = raise Assoc_fail
+   | mk_sum plus tms = foldr1 (fn (x,y) => plus $ x $ y) tms;
 
  (*Separate the literals from the other terms being combined*)
- fun sift_terms (t, (lits,others)) =
+ fun sift_terms plus (t, (lits,others)) =
      case t of
           Const("Numeral.number_of", _) $ _ =>
               (t::lits, others)         (*new literal*)
         | (f as Const _) $ x $ y =>
-              if f = Data.plus
-              then sift_terms (x, sift_terms (y, (lits,others)))
+              if f = plus
+              then sift_terms plus (x, sift_terms plus (y, (lits,others)))
               else (lits, t::others)    (*arbitrary summand*)
         | _ => (lits, t::others);
 
@@ -48,11 +46,13 @@
    let fun show t = string_of_cterm (Thm.cterm_of sg t)
        val _ = if !trace then tracing ("assoc_fold simproc: LHS = " ^ show lhs)
                else ()
-       val (lits,others) = sift_terms (lhs, ([],[]))
+       val plus =
+         (case lhs of f $ _ $ _ => f | _ => error "Assoc_fold: bad pattern")
+       val (lits,others) = sift_terms plus (lhs, ([],[]))
        val _ = if length lits < 2
                then raise Assoc_fail (*we can't reduce the number of terms*)
                else ()
-       val rhs = Data.plus $ mk_sum lits $ mk_sum others
+       val rhs = plus $ mk_sum plus lits $ mk_sum plus others
        val _ = if !trace then tracing ("RHS = " ^ show rhs) else ()
        val th = Tactic.prove sg [] [] (Logic.mk_equals (lhs, rhs))
                    (fn _ => rtac Data.eq_reflection 1 THEN
@@ -60,10 +60,6 @@
    in Some th end
    handle Assoc_fail => None;
 
- val conv =
-     Simplifier.simproc_i (Sign.deref Data.sg_ref) "assoc_fold"
-       [Data.plus $ Free ("x", Data.T) $ Free ("y",Data.T)] proc;
-
 end;
 
 
--- a/src/Provers/Arith/cancel_numerals.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/Provers/Arith/cancel_numerals.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -25,7 +25,7 @@
 signature CANCEL_NUMERALS_DATA =
 sig
   (*abstract syntax*)
-  val mk_sum: term list -> term
+  val mk_sum: typ -> term list -> term
   val dest_sum: term -> term list
   val mk_bal: term * term -> term
   val dest_bal: term -> term * term
@@ -78,7 +78,8 @@
       val u = find_common (terms1,terms2)
       val (n1, terms1') = Data.find_first_coeff u terms1
       and (n2, terms2') = Data.find_first_coeff u terms2
-      fun newshape (i,terms) = Data.mk_sum (Data.mk_coeff(i,u)::terms)
+      and T = Term.fastype_of u
+      fun newshape (i,terms) = Data.mk_sum T (Data.mk_coeff(i,u)::terms)
       val reshape =  (*Move i*u to the front and put j*u into standard form
 		       i + #m + j + k == #m + i + (j + k) *)
 	    if n1=0 orelse n2=0 then   (*trivial, so do nothing*)
@@ -94,12 +95,12 @@
 	       [Data.trans_tac reshape, rtac Data.bal_add1 1,
 		Data.numeral_simp_tac] sg hyps xs
 	       (t', Data.mk_bal (newshape(n1-n2,terms1'), 
-				 Data.mk_sum terms2'))
+				 Data.mk_sum T terms2'))
 	else
 	    Data.prove_conv 
 	       [Data.trans_tac reshape, rtac Data.bal_add2 1,
 		Data.numeral_simp_tac] sg hyps xs
-	       (t', Data.mk_bal (Data.mk_sum terms1', 
+	       (t', Data.mk_bal (Data.mk_sum T terms1', 
 				 newshape(n2-n1,terms2'))))
   end
   handle TERM _ => None
--- a/src/Provers/Arith/combine_numerals.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/Provers/Arith/combine_numerals.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -20,7 +20,7 @@
 sig
   (*abstract syntax*)
   val add: int * int -> int          (*addition (or multiplication) *)
-  val mk_sum: term list -> term
+  val mk_sum: typ -> term list -> term
   val dest_sum: term -> term list
   val mk_coeff: int * term -> term
   val dest_coeff: term -> int * term
@@ -66,22 +66,23 @@
 (*the simplification procedure*)
 fun proc sg _ t =
   let (*first freeze any Vars in the term to prevent flex-flex problems*)
-      val (t', xs) = Term.adhoc_freeze_vars t;
+      val (t', xs) = Term.adhoc_freeze_vars t
       val (u,m,n,terms) = find_repeated (Termtab.empty, [], Data.dest_sum t')
+      val T = Term.fastype_of u
       val reshape =  (*Move i*u to the front and put j*u into standard form
 		       i + #m + j + k == #m + i + (j + k) *)
 	    if m=0 orelse n=0 then   (*trivial, so do nothing*)
 		raise TERM("combine_numerals", []) 
 	    else Data.prove_conv [Data.norm_tac] sg xs
 			(t', 
-			 Data.mk_sum ([Data.mk_coeff(m,u),
-				       Data.mk_coeff(n,u)] @ terms))
+			 Data.mk_sum T ([Data.mk_coeff(m,u),
+				         Data.mk_coeff(n,u)] @ terms))
   in
       apsome Data.simplify_meta_eq
 	 (Data.prove_conv 
 	    [Data.trans_tac reshape, rtac Data.left_distrib 1,
 	     Data.numeral_simp_tac] sg xs
-	    (t', Data.mk_sum (Data.mk_coeff(Data.add(m,n), u) :: terms)))
+	    (t', Data.mk_sum T (Data.mk_coeff(Data.add(m,n), u) :: terms)))
   end
   handle TERM _ => None
        | TYPE _ => None;   (*Typically (if thy doesn't include Numeral)
--- a/src/Provers/Arith/extract_common_term.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/Provers/Arith/extract_common_term.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -17,7 +17,7 @@
 signature EXTRACT_COMMON_TERM_DATA =
 sig
   (*abstract syntax*)
-  val mk_sum: term list -> term
+  val mk_sum: typ -> term list -> term
   val dest_sum: term -> term list
   val mk_bal: term * term -> term
   val dest_bal: term -> term * term
@@ -59,11 +59,12 @@
       val u = find_common (terms1,terms2)
       val terms1' = Data.find_first u terms1
       and terms2' = Data.find_first u terms2
+      and T = Term.fastype_of u
       val reshape = 
 	    Data.prove_conv [Data.norm_tac] sg hyps xs
 	        (t', 
-		 Data.mk_bal (Data.mk_sum (u::terms1'), 
-		              Data.mk_sum (u::terms2')))
+		 Data.mk_bal (Data.mk_sum T (u::terms1'), 
+		              Data.mk_sum T (u::terms2')))
   in
       apsome Data.simplify_meta_eq reshape
   end
--- a/src/Pure/drule.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/Pure/drule.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -371,6 +371,12 @@
 (*Standard form of object-rule: no hypotheses, Frees, or outer quantifiers;
     all generality expressed by Vars having index 0.*)
 
+fun flexflex_unique th =
+    case Seq.chop (2, flexflex_rule th) of
+      ([th],_) => th
+    | ([],_)   => raise THM("flexflex_unique: impossible constraints", 0, [th])
+    |      _   => raise THM("flexflex_unique: multiple unifiers", 0, [th]);
+
 fun close_derivation thm =
   if Thm.get_name_tags thm = ("", []) then Thm.name_thm ("", thm)
   else thm;
@@ -378,7 +384,7 @@
 fun standard' th =
   let val {maxidx,...} = rep_thm th in
     th
-    |> implies_intr_hyps
+    |> flexflex_unique |> implies_intr_hyps 
     |> forall_intr_frees |> forall_elim_vars (maxidx + 1)
     |> strip_shyps_warning
     |> zero_var_indexes |> Thm.varifyT |> Thm.compress
--- a/src/ZF/Integ/int_arith.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/ZF/Integ/int_arith.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -219,7 +219,7 @@
 
 structure CancelNumeralsCommon =
   struct
-  val mk_sum            = mk_sum
+  val mk_sum            = (fn T:typ => mk_sum)
   val dest_sum          = dest_sum
   val mk_coeff          = mk_coeff
   val dest_coeff        = dest_coeff 1
@@ -292,7 +292,7 @@
 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 mk_sum            = (fn T:typ => long_mk_sum) (*to work for #2*x $+ #3*x *)
   val dest_sum          = dest_sum
   val mk_coeff          = mk_coeff
   val dest_coeff        = dest_coeff 1
@@ -328,7 +328,7 @@
 structure CombineNumeralsProdData =
   struct
   val add               = op * : int*int -> int
-  val mk_sum            = mk_prod
+  val mk_sum            = (fn T:typ => mk_prod)
   val dest_sum          = dest_prod
   fun mk_coeff(k,t) = if t=one then mk_numeral k
                       else raise TERM("mk_coeff", [])
--- a/src/ZF/arith_data.ML	Sat Feb 14 02:06:12 2004 +0100
+++ b/src/ZF/arith_data.ML	Sun Feb 15 10:46:37 2004 +0100
@@ -138,7 +138,7 @@
 
 structure CancelNumeralsCommon =
   struct
-  val mk_sum            = mk_sum
+  val mk_sum            = (fn T:typ => mk_sum)
   val dest_sum          = dest_sum
   val mk_coeff          = mk_coeff
   val dest_coeff        = dest_coeff