Installation of target HOL-Real
authorpaulson
Thu, 25 Jun 1998 13:57:34 +0200
changeset 5078 7b5ea59c0275
parent 5077 71043526295f
child 5079 2a8ed71f791f
Installation of target HOL-Real
src/HOL/Arith.ML
src/HOL/Induct/ROOT.ML
src/HOL/Integ/Group.ML
src/HOL/Integ/Group.thy
src/HOL/Integ/IntRing.ML
src/HOL/Integ/IntRing.thy
src/HOL/Integ/IntRingDefs.ML
src/HOL/Integ/IntRingDefs.thy
src/HOL/Integ/Lagrange.ML
src/HOL/Integ/Lagrange.thy
src/HOL/Integ/ROOT.ML
src/HOL/Integ/Ring.ML
src/HOL/Integ/Ring.thy
src/HOL/IsaMakefile
src/HOL/ROOT.ML
src/HOL/Real/Lubs.ML
src/HOL/Real/Lubs.thy
src/HOL/Real/PNat.ML
src/HOL/Real/PNat.thy
src/HOL/Real/PRat.ML
src/HOL/Real/PRat.thy
src/HOL/Real/PReal.ML
src/HOL/Real/PReal.thy
src/HOL/Real/RComplete.ML
src/HOL/Real/RComplete.thy
src/HOL/Real/README.html
src/HOL/Real/ROOT.ML
src/HOL/Real/Real.ML
src/HOL/Real/Real.thy
src/HOL/Real/RealAbs.ML
src/HOL/Real/RealAbs.thy
src/HOL/ex/Group.ML
src/HOL/ex/Group.thy
src/HOL/ex/IntRing.ML
src/HOL/ex/IntRing.thy
src/HOL/ex/IntRingDefs.ML
src/HOL/ex/IntRingDefs.thy
src/HOL/ex/Lagrange.ML
src/HOL/ex/Lagrange.thy
src/HOL/ex/Primrec.ML
src/HOL/ex/ROOT.ML
src/HOL/ex/Ring.ML
src/HOL/ex/Ring.thy
--- a/src/HOL/Arith.ML	Wed Jun 24 13:59:45 1998 +0200
+++ b/src/HOL/Arith.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -118,19 +118,16 @@
 qed "add_pred";
 Addsimps [add_pred];
 
+Goal "!!m::nat. m + n = m ==> n = 0";
+by (dtac (add_0_right RS ssubst) 1);
+by (asm_full_simp_tac (simpset() addsimps [add_assoc]
+                                 delsimps [add_0_right]) 1);
+qed "add_eq_self_zero";
+
 
 (**** Additional theorems about "less than" ****)
 
-Goal "i<j --> (EX k. j = Suc(i+k))";
-by (induct_tac "j" 1);
-by (Simp_tac 1);
-by (blast_tac (claset() addSEs [less_SucE] 
-                       addSIs [add_0_right RS sym, add_Suc_right RS sym]) 1);
-val lemma = result();
-
-(* [| i<j;  !!x. j = Suc(i+x) ==> Q |] ==> Q *)
-bind_thm ("less_natE", lemma RS mp RS exE);
-
+(*Deleted less_natE; instead use less_eq_Suc_add RS exE*)
 Goal "!!m. m<n --> (? k. n=Suc(m+k))";
 by (induct_tac "n" 1);
 by (ALLGOALS (simp_tac (simpset() addsimps [less_Suc_eq])));
@@ -442,6 +439,12 @@
 by (ALLGOALS Asm_simp_tac);
 qed "less_imp_diff_positive";
 
+Goal "!! (i::nat). i < j  ==> ? k. 0<k & i+k = j";
+by (res_inst_tac [("x","j - i")] exI 1);
+by (fast_tac (claset() addDs [less_trans, less_irrefl] 
+   	               addIs [less_imp_diff_positive, add_diff_inverse]) 1);
+qed "less_imp_add_positive";
+
 Goal "Suc(m)-n = (if m<n then 0 else Suc(m-n))";
 by (simp_tac (simpset() addsimps [less_imp_diff_is_0, not_less_eq, Suc_diff_n]) 1);
 qed "if_Suc_diff_n";
@@ -527,7 +530,7 @@
 
 (*strict, in 1st argument; proof is by induction on k>0*)
 Goal "!!i::nat. [| i<j; 0<k |] ==> k*i < k*j";
-by (eres_inst_tac [("i","0")] less_natE 1);
+by (eres_inst_tac [("m1","0")] (less_eq_Suc_add RS exE) 1);
 by (Asm_simp_tac 1);
 by (induct_tac "x" 1);
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_less_mono])));
--- a/src/HOL/Induct/ROOT.ML	Wed Jun 24 13:59:45 1998 +0200
+++ b/src/HOL/Induct/ROOT.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -3,7 +3,7 @@
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1997  University of Cambridge
 
-Exampled of Inductive Definitions
+Examples of Inductive and Coinductive Definitions
 *)
 
 HOL_build_completed;    (*Make examples fail if HOL did*)
--- a/src/HOL/Integ/Group.ML	Wed Jun 24 13:59:45 1998 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,222 +0,0 @@
-(*  Title:      HOL/Integ/Group.ML
-    ID:         $Id$
-    Author:     Tobias Nipkow
-    Copyright   1997 TU Muenchen
-*)
-
-(*** Groups ***)
-
-(* Derives the well-known convergent set of equations for groups
-   based on the unary inverse zero-x.
-*)
-
-Goal "!!x::'a::add_group. (zero-x)+(x+y) = y";
-by (rtac trans 1);
-by (rtac (plus_assoc RS sym) 1);
-by (stac left_inv 1);
-by (rtac zeroL 1);
-qed "left_inv2";
-
-Goal "!!x::'a::add_group. (zero-(zero-x)) = x";
-by (rtac trans 1);
-by (res_inst_tac [("x","zero-x")] left_inv2 2);
-by (stac left_inv 1);
-by (rtac (zeroR RS sym) 1);
-qed "inv_inv";
-
-Goal "zero-zero = (zero::'a::add_group)";
-by (rtac trans 1);
-by (rtac (zeroR RS sym) 1);
-by (rtac trans 1);
-by (res_inst_tac [("x","zero")] left_inv2 2);
-by (simp_tac (simpset() addsimps [zeroR]) 1);
-qed "inv_zero";
-
-Goal "!!x::'a::add_group. x+(zero-x) = zero";
-by (rtac trans 1);
-by (res_inst_tac [("x","zero-x")] left_inv 2);
-by (stac inv_inv 1);
-by (rtac refl 1);
-qed "right_inv";
-
-Goal "!!x::'a::add_group. x+((zero-x)+y) = y";
-by (rtac trans 1);
-by (res_inst_tac [("x","zero-x")] left_inv2 2);
-by (stac inv_inv 1);
-by (rtac refl 1);
-qed "right_inv2";
-
-val plus_cong = read_instantiate [("f1","op +")] (arg_cong RS cong);
-
-Goal "!!x::'a::add_group. zero-(x+y) = (zero-y)+(zero-x)";
-by (rtac trans 1);
- by (rtac zeroR 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (rtac refl 2);
- by (res_inst_tac [("x","x+y")] right_inv 2);
-by (rtac trans 1);
- by (rtac plus_assoc 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (simp_tac (simpset() addsimps
-        [plus_assoc,left_inv,left_inv2,right_inv,right_inv2]) 2);
- by (rtac refl 2);
-by (rtac (zeroL RS sym) 1);
-qed "inv_plus";
-
-(*** convergent TRS for groups with unary inverse zero-x ***)
-val group1_simps =
-  [zeroL,zeroR,plus_assoc,left_inv,left_inv2,right_inv,right_inv2,inv_inv,
-   inv_zero,inv_plus];
-
-val group1_tac =
-  let val ss = HOL_basic_ss addsimps group1_simps
-  in simp_tac ss end;
-
-(* I believe there is no convergent TRS for groups with binary `-',
-   unless you have an extra unary `-' and simply define x-y = x+(-y).
-   This does not work with only a binary `-' because x-y = x+(zero-y) does
-   not terminate. Hence we have a special tactic for converting all
-   occurrences of x-y into x+(zero-y):
-*)
-
-local
-fun find(Const("op -",Type("fun",[T,_]))$s$t) = [(T,s,t)] @ find s @ find t
-  | find(s$t) = find s @ find t
-  | find _ = [];
-
-fun subst_tac sg (tacf,(T,s,t)) = 
-  let val typinst = [(("'a",0),ctyp_of sg T)];
-      val terminst = [(cterm_of sg (Var(("x",0),T)),cterm_of sg s),
-                      (cterm_of sg (Var(("y",0),T)),cterm_of sg t)];
-  in tacf THEN' rtac ((instantiate(typinst,terminst) minus_inv) RS ssubst) end;
-
-in
-val mk_group1_tac = SUBGOAL(fn (t,i) => fn st =>
-      let val sg = #sign(rep_thm st)
-      in foldl (subst_tac sg) (K all_tac,find t) i st
-      end)
-end;
-
-(* The following two equations are not used in any of the decision procedures,
-   but are still very useful. They also demonstrate mk_group1_tac.
-*)
-Goal "x-x = (zero::'a::add_group)";
-by (mk_group1_tac 1);
-by (group1_tac 1);
-qed "minus_self_zero";
-
-Goal "x-zero = (x::'a::add_group)";
-by (mk_group1_tac 1);
-by (group1_tac 1);
-qed "minus_zero";
-
-(*** Abelian Groups ***)
-
-Goal "x+(y+z)=y+(x+z::'a::add_agroup)";
-by (rtac trans 1);
-by (rtac plus_commute 1);
-by (rtac trans 1);
-by (rtac plus_assoc 1);
-by (simp_tac (simpset() addsimps [plus_commute]) 1);
-qed "plus_commuteL";
-
-(* Convergent TRS for Abelian groups with unary inverse zero-x.
-   Requires ordered rewriting
-*)
-
-val agroup1_simps = plus_commute::plus_commuteL::group1_simps;
-
-val agroup1_tac =
-  let val ss = HOL_basic_ss addsimps agroup1_simps
-  in simp_tac ss end;
-
-(* Again, I do not believe there is a convergent TRS for Abelian Groups with
-   binary `-'. However, we can still decide the word problem using additional
-   rules for 
-   1. floating `-' to the top:
-      "x + (y - z) = (x + y) - (z::'a::add_group)"
-      "(x - y) + z = (x + z) - (y::'a::add_agroup)"
-      "(x - y) - z = x - (y + (z::'a::add_agroup))"
-      "x - (y - z) = (x + z) - (y::'a::add_agroup)"
-   2. and for moving `-' over to the other side:
-      (x-y = z) = (x = z+y) and (x = y-z) = (x+z = y)
-*)
-Goal "x + (y - z) = (x + y) - (z::'a::add_group)";
-by (mk_group1_tac 1);
-by (group1_tac 1);
-qed "plus_minusR";
-
-Goal "(x - y) + z = (x + z) - (y::'a::add_agroup)";
-by (mk_group1_tac 1);
-by (agroup1_tac 1);
-qed "plus_minusL";
-
-Goal "(x - y) - z = x - (y + (z::'a::add_agroup))";
-by (mk_group1_tac 1);
-by (agroup1_tac 1);
-qed "minus_minusL";
-
-Goal "x - (y - z) = (x + z) - (y::'a::add_agroup)";
-by (mk_group1_tac 1);
-by (agroup1_tac 1);
-qed "minus_minusR";
-
-Goal "!!x::'a::add_group. (x-y = z) = (x = z+y)";
-by (stac minus_inv 1);
-by (fast_tac (claset() addss (HOL_basic_ss addsimps group1_simps)) 1);
-qed "minusL_iff";
-
-Goal "!!x::'a::add_group. (x = y-z) = (x+z = y)";
-by (stac minus_inv 1);
-by (fast_tac (claset() addss (HOL_basic_ss addsimps group1_simps)) 1);
-qed "minusR_iff";
-
-val agroup2_simps =
-  [zeroL,zeroR,plus_assoc,plus_commute,plus_commuteL,
-   plus_minusR,plus_minusL,minus_minusL,minus_minusR,minusL_iff,minusR_iff];
-
-(* This two-phase ordered rewriting tactic works, but agroup_simps are
-   simpler. However, agroup_simps is not confluent for arbitrary terms,
-   it merely decides equality.
-(* Phase 1 *)
-
-Goal "!!x::'a::add_agroup. (x+(zero-y))+z = (x+z)+(zero-y)";
-by (Simp_tac 1);
-val lemma = result();
-bind_thm("plus_minusL",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma);
-
-Goal "!!x::'a::add_agroup. x+(zero-(y+z)) = (x+(zero-y))+(zero-z)";
-by (Simp_tac 1);
-val lemma = result();
-bind_thm("minus_plusR",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma);
-
-Goal "!!x::'a::add_agroup. x+(zero-(y+(zero-z))) = (x+z)+(zero-y)";
-by (Simp_tac 1);
-val lemma = result();
-bind_thm("minus_minusR",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma);
-
-Goal "!!x::'a::add_agroup. x+(y+(zero-z)) = (x+y)+(zero-z)";
-by (Simp_tac 1);
-val lemma = result();
-bind_thm("plus_minusR",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma);
-
-(* Phase 2 *)
-
-Goal "!!x::'a::add_agroup. (x+y)+(zero-z) = x+(y+(zero-z))";
-by (Simp_tac 1);
-val lemma = result();
-bind_thm("minus_plusL2",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma);
-
-Goal "!!x::'a::add_agroup. (x+y)+(zero-x) = y";
-by (rtac (plus_assoc RS trans) 1);
-by (rtac trans 1);
- by (rtac plus_cong 1);
-  by (rtac refl 1);
- by (rtac right_inv2 2);
-by (rtac plus_commute 1);
-val lemma = result();
-bind_thm("minus_plusL3",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma);
-
-*)
--- a/src/HOL/Integ/Group.thy	Wed Jun 24 13:59:45 1998 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,44 +0,0 @@
-(*  Title:      HOL/Integ/Group.thy
-    ID:         $Id$
-    Author:     Tobias Nipkow
-    Copyright   1996 TU Muenchen
-
-A little bit of group theory leading up to rings. Hence groups are additive.
-*)
-
-Group = Set +
-
-(* 0 already used in Nat *)
-axclass  zero < term
-consts   zero :: "'a::zero"
-
-(* additive semigroups *)
-
-axclass  add_semigroup < plus
-  plus_assoc   "(x + y) + z = x + (y + z)"
-
-
-(* additive monoids *)
-
-axclass  add_monoid < add_semigroup, zero
-  zeroL    "zero + x = x"
-  zeroR    "x + zero = x"
-
-(* additive groups *)
-(*
-The inverse is the binary `-'. Groups with unary and binary inverse are
-interdefinable: x-y := x+(zero-y) and -x := zero-x. The law left_inv is
-simply the translation of (-x)+x = zero. This characterizes groups already,
-provided we only allow (zero-x). Law minus_inv `defines' the general x-y in
-terms of the specific zero-y.
-*)
-axclass  add_group < add_monoid, minus
-  left_inv  "(zero-x)+x = zero"
-  minus_inv "x-y = x + (zero-y)"
-
-(* additive abelian groups *)
-
-axclass  add_agroup < add_group
-  plus_commute  "x + y = y + x"
-
-end
--- a/src/HOL/Integ/IntRing.ML	Wed Jun 24 13:59:45 1998 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,18 +0,0 @@
-(*  Title:      HOL/Integ/IntRing.ML
-    ID:         $Id$
-    Author:     Tobias Nipkow and Markus Wenzel
-    Copyright   1996 TU Muenchen
-
-The instantiation of Lagrange's lemma for int.
-*)
-
-open IntRing;
-
-Goal "!!i1::int. \
-\  (sq i1 + sq i2 + sq i3 + sq i4) * (sq j1 + sq j2 + sq j3 + sq j4) = \
-\  sq(i1*j1 - i2*j2 - i3*j3 - i4*j4)  + \
-\  sq(i1*j2 + i2*j1 + i3*j4 - i4*j3)  + \
-\  sq(i1*j3 - i2*j4 + i3*j1 + i4*j2)  + \
-\  sq(i1*j4 + i2*j3 - i3*j2 + i4*j1)";
-by (rtac Lagrange_lemma 1);
-qed "Lagrange_lemma_int";
--- a/src/HOL/Integ/IntRing.thy	Wed Jun 24 13:59:45 1998 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,19 +0,0 @@
-(*  Title:      HOL/Integ/IntRing.thy
-    ID:         $Id$
-    Author:     Tobias Nipkow and Markus Wenzel
-    Copyright   1996 TU Muenchen
-
-The integers form a commutative ring.
-With an application of Lagrange's lemma.
-*)
-
-IntRing = IntRingDefs + Lagrange +
-
-instance int :: add_semigroup (zadd_assoc)
-instance int :: add_monoid (zero_int_def,zadd_0,zadd_0_right)
-instance int :: add_group (left_inv_int,minus_inv_int)
-instance int :: add_agroup (zadd_commute)
-instance int :: ring (zmult_assoc,zadd_zmult_distrib2,zadd_zmult_distrib)
-instance int :: cring (zmult_commute)
-
-end
--- a/src/HOL/Integ/IntRingDefs.ML	Wed Jun 24 13:59:45 1998 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,16 +0,0 @@
-(*  Title:      HOL/Integ/IntRingDefs.thy
-    ID:         $Id$
-    Author:     Tobias Nipkow and Markus Wenzel
-    Copyright   1996 TU Muenchen
-
-*)
-
-open IntRingDefs;
-
-Goalw [zero_int_def,zdiff_def] "(zero-x)+(x::int) = zero";
-by (Simp_tac 1);
-qed "left_inv_int";
-
-Goalw [zero_int_def,zdiff_def] "x-y = (x::int) + (zero-y)";
-by (Simp_tac 1);
-qed "minus_inv_int";
--- a/src/HOL/Integ/IntRingDefs.thy	Wed Jun 24 13:59:45 1998 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,15 +0,0 @@
-(*  Title:      HOL/Integ/IntRingDefs.thy
-    ID:         $Id$
-    Author:     Tobias Nipkow and Markus Wenzel
-    Copyright   1996 TU Muenchen
-
-Provides the basic defs and thms for showing that int is a commutative ring.
-Most of it has been defined and shown already.
-*)
-
-IntRingDefs = Integ + Ring +
-
-instance int :: zero
-defs zero_int_def "zero::int == $# 0"
-
-end
--- a/src/HOL/Integ/Lagrange.ML	Wed Jun 24 13:59:45 1998 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,37 +0,0 @@
-(*  Title:      HOL/Integ/Lagrange.ML
-    ID:         $Id$
-    Author:     Tobias Nipkow
-    Copyright   1996 TU Muenchen
-
-
-The following lemma essentially shows that all composite natural numbers are
-sums of fours squares, provided all prime numbers are. However, this is an
-abstract thm about commutative rings and has a priori nothing to do with nat.
-*)
-
-Goalw [Lagrange.sq_def] "!!x1::'a::cring. \
-\  (sq x1 + sq x2 + sq x3 + sq x4) * (sq y1 + sq y2 + sq y3 + sq y4) = \
-\  sq(x1*y1 - x2*y2 - x3*y3 - x4*y4)  + \
-\  sq(x1*y2 + x2*y1 + x3*y4 - x4*y3)  + \
-\  sq(x1*y3 - x2*y4 + x3*y1 + x4*y2)  + \
-\  sq(x1*y4 + x2*y3 - x3*y2 + x4*y1)";
-(*Takes up to three minutes...*)
-by (cring_tac 1);
-qed "Lagrange_lemma";
-
-(* A challenge by John Harrison.
-   Takes forever because of the naive bottom-up strategy of the rewriter.
-
-Goalw [Lagrange.sq_def] "!!p1::'a::cring.\
-\ (sq p1 + sq q1 + sq r1 + sq s1 + sq t1 + sq u1 + sq v1 + sq w1) * \
-\ (sq p2 + sq q2 + sq r2 + sq s2 + sq t2 + sq u2 + sq v2 + sq w2) \
-\  = sq (p1*p2 - q1*q2 - r1*r2 - s1*s2 - t1*t2 - u1*u2 - v1*v2 - w1*w2) + \
-\    sq (p1*q2 + q1*p2 + r1*s2 - s1*r2 + t1*u2 - u1*t2 - v1*w2 + w1*v2) +\
-\    sq (p1*r2 - q1*s2 + r1*p2 + s1*q2 + t1*v2 + u1*w2 - v1*t2 - w1*u2) +\
-\    sq (p1*s2 + q1*r2 - r1*q2 + s1*p2 + t1*w2 - u1*v2 + v1*u2 - w1*t2) +\
-\    sq (p1*t2 - q1*u2 - r1*v2 - s1*w2 + t1*p2 + u1*q2 + v1*r2 + w1*s2) +\
-\    sq (p1*u2 + q1*t2 - r1*w2 + s1*v2 - t1*q2 + u1*p2 - v1*s2 + w1*r2) +\
-\    sq (p1*v2 + q1*w2 + r1*t2 - s1*u2 - t1*r2 + u1*s2 + v1*p2 - w1*q2) +\
-\    sq (p1*w2 - q1*v2 + r1*u2 + s1*t2 - t1*s2 - u1*r2 + v1*q2 + w1*p2)";
-
-*)
--- a/src/HOL/Integ/Lagrange.thy	Wed Jun 24 13:59:45 1998 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,18 +0,0 @@
-(*  Title:      HOL/Integ/Lagrange.thy
-    ID:         $Id$
-    Author:     Tobias Nipkow
-    Copyright   1996 TU Muenchen
-
-
-This theory only contains a single thm, which is a lemma in Lagrange's proof
-that every natural number is the sum of 4 squares.  It's sole purpose is to
-demonstrate ordered rewriting for commutative rings.
-
-The enterprising reader might consider proving all of Lagrange's thm.
-*)
-Lagrange = Ring +
-
-constdefs sq :: 'a::times => 'a
-         "sq x == x*x"
-
-end
--- a/src/HOL/Integ/ROOT.ML	Wed Jun 24 13:59:45 1998 +0200
+++ b/src/HOL/Integ/ROOT.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -6,7 +6,4 @@
 The Integers in HOL (ported from ZF by Riccardo Mattolini)
 *)
 
-HOL_build_completed;    (*Cause examples to fail if HOL did*)
-
 time_use_thy "Bin";
-time_use_thy "IntRing";
--- a/src/HOL/Integ/Ring.ML	Wed Jun 24 13:59:45 1998 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,139 +0,0 @@
-(*  Title:      HOL/Integ/Ring.ML
-    ID:         $Id$
-    Author:     Tobias Nipkow
-    Copyright   1996 TU Muenchen
-
-Derives a few equational consequences about rings
-and defines cring_simpl, a simplification tactic for commutative rings.
-*)
-
-Goal "!!x::'a::cring. x*(y*z)=y*(x*z)";
-by (rtac trans 1);
-by (rtac times_commute 1);
-by (rtac trans 1);
-by (rtac times_assoc 1);
-by (simp_tac (HOL_basic_ss addsimps [times_commute]) 1);
-qed "times_commuteL";
-
-val times_cong = read_instantiate [("f1","op *")] (arg_cong RS cong);
-
-Goal "!!x::'a::ring. zero*x = zero";
-by (rtac trans 1);
- by (rtac right_inv 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (rtac refl 3);
- by (rtac trans 2);
-  by (rtac times_cong 3);
-   by (rtac zeroL 3);
-  by (rtac refl 3);
- by (rtac (distribR RS sym) 2);
-by (rtac trans 1);
- by (rtac (plus_assoc RS sym) 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (rtac refl 2);
- by (rtac (right_inv RS sym) 2);
-by (rtac (zeroR RS sym) 1);
-qed "mult_zeroL";
-
-Goal "!!x::'a::ring. x*zero = zero";
-by (rtac trans 1);
- by (rtac right_inv 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (rtac refl 3);
- by (rtac trans 2);
-  by (rtac times_cong 3);
-   by (rtac zeroL 4);
-  by (rtac refl 3);
- by (rtac (distribL RS sym) 2);
-by (rtac trans 1);
- by (rtac (plus_assoc RS sym) 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (rtac refl 2);
- by (rtac (right_inv RS sym) 2);
-by (rtac (zeroR RS sym) 1);
-qed "mult_zeroR";
-
-Goal "!!x::'a::ring. (zero-x)*y = zero-(x*y)";
-by (rtac trans 1);
- by (rtac zeroL 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (rtac refl 3);
- by (rtac mult_zeroL 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (rtac refl 3);
- by (rtac times_cong 2);
-  by (rtac left_inv 2);
- by (rtac refl 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (rtac refl 3);
- by (rtac (distribR RS sym) 2);
-by (rtac trans 1);
- by (rtac (plus_assoc RS sym) 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (rtac refl 2);
- by (rtac (right_inv RS sym) 2);
-by (rtac (zeroR RS sym) 1);
-qed "mult_invL";
-
-Goal "!!x::'a::ring. x*(zero-y) = zero-(x*y)";
-by (rtac trans 1);
- by (rtac zeroL 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (rtac refl 3);
- by (rtac mult_zeroR 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (rtac refl 3);
- by (rtac times_cong 2);
-  by (rtac refl 2);
- by (rtac left_inv 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (rtac refl 3);
- by (rtac (distribL RS sym) 2);
-by (rtac trans 1);
- by (rtac (plus_assoc RS sym) 2);
-by (rtac trans 1);
- by (rtac plus_cong 2);
-  by (rtac refl 2);
- by (rtac (right_inv RS sym) 2);
-by (rtac (zeroR RS sym) 1);
-qed "mult_invR";
-
-Goal "x*(y-z) = (x*y - x*z::'a::ring)";
-by (mk_group1_tac 1);
-by (simp_tac (HOL_basic_ss addsimps [distribL,mult_invR]) 1);
-qed "minus_distribL";
-
-Goal "(x-y)*z = (x*z - y*z::'a::ring)";
-by (mk_group1_tac 1);
-by (simp_tac (HOL_basic_ss addsimps [distribR,mult_invL]) 1);
-qed "minus_distribR";
-
-val cring_simps = [times_assoc,times_commute,times_commuteL,
-                   distribL,distribR,minus_distribL,minus_distribR]
-                  @ agroup2_simps;
-
-val cring_tac =
-  let val ss = HOL_basic_ss addsimps cring_simps
-  in simp_tac ss end;
-
-
-(*** The order [minus_plusL3,minus_plusL2] is important because minus_plusL3
-     MUST be tried first
-val cring_simp =
-  let val phase1 = simpset() addsimps
-                   [plus_minusL,minus_plusR,minus_minusR,plus_minusR]
-      val phase2 = HOL_ss addsimps [minus_plusL3,minus_plusL2,
-                                    zeroL,zeroR,mult_zeroL,mult_zeroR]
-  in simp_tac phase1 THEN' simp_tac phase2 end;
-***)
--- a/src/HOL/Integ/Ring.thy	Wed Jun 24 13:59:45 1998 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-(*  Title:      HOL/Integ/Ring.thy
-    ID:         $Id$
-    Author:     Tobias Nipkow
-    Copyright   1996 TU Muenchen
-
-Bits of rings.
-Main output: a simplification tactic for commutative rings.
-*)
-
-Ring = Group +
-
-(* Ring *)
-
-axclass  ring < add_agroup, times
-  times_assoc "(x*y)*z = x*(y*z)"
-  distribL    "x*(y+z) = x*y + x*z"
-  distribR    "(x+y)*z = x*z + y*z"
-
-(* Commutative ring *)
-
-axclass cring < ring
-  times_commute "x*y = y*x"
-
-end
--- a/src/HOL/IsaMakefile	Wed Jun 24 13:59:45 1998 +0200
+++ b/src/HOL/IsaMakefile	Thu Jun 25 13:57:34 1998 +0200
@@ -8,7 +8,7 @@
 
 default: HOL
 images: HOL TLA
-test: HOL-Subst HOL-Induct HOL-IMP HOL-Hoare HOL-Lex HOL-Integ \
+test: HOL-Subst HOL-Induct HOL-IMP HOL-Hoare HOL-Lex HOL-Real \
   HOL-Auth HOL-UNITY HOL-Modelcheck HOL-Lambda HOL-W0 HOL-MiniML HOL-IOA \
   HOL-AxClasses HOL-AxClasses-Group HOL-AxClasses-Lattice \
   HOL-AxClasses-Tutorial HOL-Quot HOL-ex TLA-Inc TLA-Buffer TLA-Memory
@@ -39,6 +39,8 @@
   $(SRC)/TFL/tfl.sml $(SRC)/TFL/thms.sig $(SRC)/TFL/thms.sml \
   $(SRC)/TFL/thry.sig $(SRC)/TFL/thry.sml $(SRC)/TFL/usyntax.sig \
   $(SRC)/TFL/usyntax.sml $(SRC)/TFL/utils.sig $(SRC)/TFL/utils.sml \
+  Integ/Bin.ML Integ/Bin.thy  Integ/Equiv.ML Integ/Equiv.thy \
+  Integ/Integ.ML Integ/Integ.thy Integ/ROOT.ML \
   Arith.ML Arith.thy Divides.ML Divides.thy Finite.ML Finite.thy \
   Fun.ML Fun.thy Gfp.ML Gfp.thy HOL.ML HOL.thy Inductive.ML \
   Inductive.thy Lfp.ML Lfp.thy List.ML List.thy Map.ML Map.thy Nat.ML \
@@ -117,16 +119,16 @@
 	@$(ISATOOL) usedir $(OUT)/HOL Lex
 
 
-## HOL-Integ
+## HOL-Real
 
-HOL-Integ: HOL $(LOG)/HOL-Integ.gz
+HOL-Real: HOL $(LOG)/HOL-Real.gz
 
-$(LOG)/HOL-Integ.gz: $(OUT)/HOL Integ/Bin.ML Integ/Bin.thy \
-  Integ/Equiv.ML Integ/Equiv.thy Integ/Group.ML Integ/Group.thy \
-  Integ/IntRing.ML Integ/IntRing.thy Integ/IntRingDefs.ML \
-  Integ/IntRingDefs.thy Integ/Integ.ML Integ/Integ.thy Integ/Lagrange.ML \
-  Integ/Lagrange.thy Integ/ROOT.ML Integ/Ring.ML Integ/Ring.thy
-	@$(ISATOOL) usedir $(OUT)/HOL Integ
+$(LOG)/HOL-Real.gz: $(OUT)/HOL \
+  Real/Lubs.ML Real/Lubs.thy Real/PNat.ML Real/PNat.thy \
+  Real/PRat.ML Real/PRat.thy Real/PReal.ML Real/PReal.thy \
+  Real/RComplete.ML Real/RComplete.thy Real/Real.ML Real/Real.thy \
+  Real/RealAbs.ML Real/RealAbs.thy Real/ROOT.ML
+	@$(ISATOOL) usedir $(OUT)/HOL Real
 
 
 ## HOL-Auth
@@ -282,7 +284,10 @@
   ex/NatSum.thy ex/Primes.ML ex/Primes.thy ex/Primrec.ML ex/Primrec.thy \
   ex/Puzzle.ML ex/Puzzle.thy ex/Qsort.ML ex/Qsort.thy ex/ROOT.ML \
   ex/Recdef.ML ex/Recdef.thy ex/String.ML ex/String.thy ex/cla.ML \
-  ex/meson.ML ex/mesontest.ML ex/set.ML
+  ex/meson.ML ex/mesontest.ML ex/set.ML \
+  ex/Group.ML ex/Group.thy  ex/IntRing.ML ex/IntRing.thy \
+  ex/IntRingDefs.ML ex/IntRingDefs.thy \
+  ex/Lagrange.ML ex/Lagrange.thy ex/Ring.ML ex/Ring.thy
 	@$(ISATOOL) usedir $(OUT)/HOL ex
 
 
@@ -336,8 +341,7 @@
 clean:
 	@rm -f $(OUT)/HOL $(LOG)/HOL.gz $(LOG)/HOL-Subst.gz \
 	  $(LOG)/HOL-Induct.gz $(LOG)/HOL-IMP.gz $(LOG)/HOL-Hoare.gz \
-	  $(LOG)/HOL-Lex.gz $(LOG)/HOL-Integ.gz \
-	  $(LOG)/HOL-Auth.gz $(LOG)/HOL-UNITY.gz \
+	  $(LOG)/HOL-Lex.gz $(LOG)/HOL-Auth.gz $(LOG)/HOL-UNITY.gz \
 	  $(LOG)/HOL-Modelcheck.gz $(LOG)/HOL-Lambda.gz $(LOG)/HOL-W0.gz \
 	  $(LOG)/HOL-MiniML.gz $(LOG)/HOL-IOA.gz $(LOG)/HOL-AxClasses.gz \
 	  $(LOG)/HOL-AxClasses-Group.gz $(LOG)/HOL-AxClasses-Lattice.gz \
--- a/src/HOL/ROOT.ML	Wed Jun 24 13:59:45 1998 +0200
+++ b/src/HOL/ROOT.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -59,9 +59,12 @@
 use_thy "Map";
 use_thy "Update";
 
+use_dir "Integ";
+
 (*TFL: recursive function definitions*)
 cd "$ISABELLE_HOME/src/TFL";
 use "sys.sml";
+cd "$ISABELLE_HOME/src/HOL";
 
 print_depth 8;
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/Lubs.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,101 @@
+(*  Title       : Lubs.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Completeness of the reals. A few of the 
+                  definitions suggested by James Margetson
+*) 
+
+open Lubs;
+
+(*------------------------------------------------------------------------
+                        Rules for *<= and <=*
+ ------------------------------------------------------------------------*)
+Goalw [setle_def] "!!x. ALL y: S. y <= x ==> S *<= x";
+by (assume_tac 1);
+qed "setleI";
+
+Goalw [setle_def] "!!x. [| S *<= x; y: S |] ==> y <= x";
+by (Fast_tac 1);
+qed "setleD";
+
+Goalw [setge_def] "!!x. ALL y: S. x<= y ==> x <=* S";
+by (assume_tac 1);
+qed "setgeI";
+
+Goalw [setge_def] "!!x. [| x <=* S; y: S |] ==> x <= y";
+by (Fast_tac 1);
+qed "setgeD";
+
+(*------------------------------------------------------------------------
+                        Rules about leastP, ub and lub
+ ------------------------------------------------------------------------*)
+Goalw [leastP_def] "!!x. leastP P x ==> P x";
+by (Step_tac 1);
+qed "leastPD1";
+
+Goalw [leastP_def] "!!x. leastP P x ==> x <=* Collect P";
+by (Step_tac 1);
+qed "leastPD2";
+
+Goal "!!x. [| leastP P x; y: Collect P |] ==> x <= y";
+by (blast_tac (claset() addSDs [leastPD2,setgeD]) 1);
+qed "leastPD3";
+
+Goalw [isLub_def,isUb_def,leastP_def] 
+      "!!x. isLub R S x ==> S *<= x";
+by (Step_tac 1);
+qed "isLubD1";
+
+Goalw [isLub_def,isUb_def,leastP_def] 
+      "!!x. isLub R S x ==> x: R";
+by (Step_tac 1);
+qed "isLubD1a";
+
+Goalw [isUb_def] "!!x. isLub R S x ==> isUb R S x";
+by (blast_tac (claset() addDs [isLubD1,isLubD1a]) 1);
+qed "isLub_isUb";
+
+Goal "!!x. [| isLub R S x; y : S |] ==> y <= x";
+by (blast_tac (claset() addSDs [isLubD1,setleD]) 1);
+qed "isLubD2";
+
+Goalw [isLub_def] "!!x. isLub R S x ==> leastP(isUb R S) x";
+by (assume_tac 1);
+qed "isLubD3";
+
+Goalw [isLub_def] "!!x. leastP(isUb R S) x ==> isLub R S x";
+by (assume_tac 1);
+qed "isLubI1";
+
+Goalw [isLub_def,leastP_def] 
+      "!!x. [| isUb R S x; x <=* Collect (isUb R S) |] ==> isLub R S x";
+by (Step_tac 1);
+qed "isLubI2";
+
+Goalw [isUb_def] "!!x. [| isUb R S x; y : S |] ==> y <= x";
+by (fast_tac (claset() addDs [setleD]) 1);
+qed "isUbD";
+
+Goalw [isUb_def] "!!x. isUb R S x ==> S *<= x";
+by (Step_tac 1);
+qed "isUbD2";
+
+Goalw [isUb_def] "!!x. isUb R S x ==> x: R";
+by (Step_tac 1);
+qed "isUbD2a";
+
+Goalw [isUb_def] "!!x. [| S *<= x; x: R |] ==> isUb R S x";
+by (Step_tac 1);
+qed "isUbI";
+
+Goalw [isLub_def] "!!x. [| isLub R S x; isUb R S y |] ==> x <= y";
+by (blast_tac (claset() addSIs [leastPD3]) 1);
+qed "isLub_le_isUb";
+
+Goalw [ubs_def,isLub_def] "!!x. isLub R S x ==> x <=* ubs R S";
+by (etac leastPD2 1);
+qed "isLub_ubs";
+
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/Lubs.thy	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,38 @@
+(*  Title       : Lubs.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Upper bounds, lubs definitions
+                  suggested by James Margetson
+*) 
+
+
+Lubs = Set +
+
+consts
+    
+    "*<=" :: ['a set, 'a] => bool     (infixl 70)
+    "<=*" :: ['a, 'a set] => bool     (infixl 70)
+
+constdefs
+    leastP      :: ['a =>bool,'a] => bool
+    "leastP P x == (P x & x <=* Collect P)"
+
+    isLub       :: ['a set, 'a set, 'a] => bool    
+    "isLub R S x  == leastP (isUb R S) x"
+
+    isUb        :: ['a set, 'a set, 'a] => bool     
+    "isUb R S x   == S *<= x & x: R"
+
+    ubs         :: ['a set, 'a set] => 'a set
+    "ubs R S      == Collect (isUb R S)"
+
+defs
+    setle_def
+    "S *<= x    == (ALL y: S. y <= x)"
+
+    setge_def
+    "x <=* S    == (ALL y: S. x <= y)"
+
+end                    
+
+    
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/PNat.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,723 @@
+(*  Title       : PNat.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : The positive naturals -- proofs 
+                : mainly as in Nat.thy
+*)
+
+open PNat;
+
+Goal "mono(%X. {1} Un (Suc``X))";
+by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1));
+qed "pnat_fun_mono";
+
+val pnat_unfold = pnat_fun_mono RS (pnat_def RS def_lfp_Tarski);
+
+Goal "1 : pnat";
+by (stac pnat_unfold 1);
+by (rtac (singletonI RS UnI1) 1);
+qed "one_RepI";
+
+Addsimps [one_RepI];
+
+Goal "i: pnat ==> Suc(i) : pnat";
+by (stac pnat_unfold 1);
+by (etac (imageI RS UnI2) 1);
+qed "pnat_Suc_RepI";
+
+Goal "2 : pnat";
+by (rtac (one_RepI RS pnat_Suc_RepI) 1);
+qed "two_RepI";
+
+(*** Induction ***)
+
+val major::prems = goal thy
+    "[| i: pnat;  P(1);   \
+\       !!j. [| j: pnat; P(j) |] ==> P(Suc(j)) |]  ==> P(i)";
+by (rtac ([pnat_def, pnat_fun_mono, major] MRS def_induct) 1);
+by (blast_tac (claset() addIs prems) 1);
+qed "PNat_induct";
+
+val prems = goalw thy [pnat_one_def,pnat_Suc_def]
+    "[| P(1p);   \
+\       !!n. P(n) ==> P(pSuc n) |]  ==> P(n)";
+by (rtac (Rep_pnat_inverse RS subst) 1);   
+by (rtac (Rep_pnat RS PNat_induct) 1);
+by (REPEAT (ares_tac prems 1
+     ORELSE eresolve_tac [Abs_pnat_inverse RS subst] 1));
+qed "pnat_induct";
+
+(*Perform induction on n. *)
+local fun raw_pnat_ind_tac a i = 
+    res_inst_tac [("n",a)] pnat_induct i  THEN  rename_last_tac a [""] (i+1)
+in
+val pnat_ind_tac = Datatype.occs_in_prems raw_pnat_ind_tac
+end;
+
+val prems = goal thy
+    "[| !!x. P x 1p;  \
+\       !!y. P 1p (pSuc y);  \
+\       !!x y. [| P x y |] ==> P (pSuc x) (pSuc y)  \
+\    |] ==> P m n";
+by (res_inst_tac [("x","m")] spec 1);
+by (pnat_ind_tac "n" 1);
+by (rtac allI 2);
+by (pnat_ind_tac "x" 2);
+by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1));
+qed "pnat_diff_induct";
+
+(*Case analysis on the natural numbers*)
+val prems = goal thy 
+    "[| n=1p ==> P;  !!x. n = pSuc(x) ==> P |] ==> P";
+by (subgoal_tac "n=1p | (EX x. n = pSuc(x))" 1);
+by (fast_tac (claset() addSEs prems) 1);
+by (pnat_ind_tac "n" 1);
+by (rtac (refl RS disjI1) 1);
+by (Blast_tac 1);
+qed "pnatE";
+
+(*** Isomorphisms: Abs_Nat and Rep_Nat ***)
+
+Goal "inj_on Abs_pnat pnat";
+by (rtac inj_on_inverseI 1);
+by (etac Abs_pnat_inverse 1);
+qed "inj_on_Abs_pnat";
+
+Addsimps [inj_on_Abs_pnat RS inj_on_iff];
+
+Goal "inj(Rep_pnat)";
+by (rtac inj_inverseI 1);
+by (rtac Rep_pnat_inverse 1);
+qed "inj_Rep_pnat";
+
+bind_thm ("Zero_not_Suc", Suc_not_Zero RS not_sym);
+
+Goal "0 ~: pnat";
+by (stac pnat_unfold 1);
+by Auto_tac;
+qed "zero_not_mem_pnat";
+
+(* 0 : pnat ==> P *)
+bind_thm ("zero_not_mem_pnatE", zero_not_mem_pnat RS notE);
+
+Addsimps [zero_not_mem_pnat];
+
+Goal "!!x. x : pnat ==> 0 < x";
+by (dtac (pnat_unfold RS subst) 1);
+by Auto_tac;
+qed "mem_pnat_gt_zero";
+
+Goal "!!x. 0 < x ==> x: pnat";
+by (stac pnat_unfold 1);
+by (dtac (gr_implies_not0 RS not0_implies_Suc) 1); 
+by (etac exE 1 THEN Asm_simp_tac 1);
+by (induct_tac "m" 1);
+by (auto_tac (claset(),simpset() 
+    addsimps [one_RepI]) THEN dtac pnat_Suc_RepI 1);
+by (Blast_tac 1);
+qed "gt_0_mem_pnat";
+
+Goal "(x: pnat) = (0 < x)";
+by (blast_tac (claset() addDs [mem_pnat_gt_zero,gt_0_mem_pnat]) 1);
+qed "mem_pnat_gt_0_iff";
+
+Goal "0 < Rep_pnat x";
+by (rtac (Rep_pnat RS mem_pnat_gt_zero) 1);
+qed "Rep_pnat_gt_zero";
+
+Goalw [pnat_add_def] "(x::pnat) + y = y + x";
+by (simp_tac (simpset() addsimps [add_commute]) 1);
+qed "pnat_add_commute";
+
+(** alternative definition for pnat **)
+(** order isomorphism **)
+Goal "pnat = {x::nat. 0 < x}";
+by (rtac set_ext 1);
+by (simp_tac (simpset() addsimps 
+    [mem_pnat_gt_0_iff]) 1);
+qed "Collect_pnat_gt_0";
+
+(*** Distinctness of constructors ***)
+
+Goalw [pnat_one_def,pnat_Suc_def] "pSuc(m) ~= 1p";
+by (rtac (inj_on_Abs_pnat RS inj_on_contraD) 1);
+by (rtac (Rep_pnat_gt_zero RS Suc_mono RS less_not_refl2) 1);
+by (REPEAT (resolve_tac [Rep_pnat RS  pnat_Suc_RepI, one_RepI] 1));
+qed "pSuc_not_one";
+
+bind_thm ("one_not_pSuc", pSuc_not_one RS not_sym);
+
+AddIffs [pSuc_not_one,one_not_pSuc];
+
+bind_thm ("pSuc_neq_one", (pSuc_not_one RS notE));
+val one_neq_pSuc = sym RS pSuc_neq_one;
+
+(** Injectiveness of pSuc **)
+
+Goalw [pnat_Suc_def] "inj(pSuc)";
+by (rtac injI 1);
+by (dtac (inj_on_Abs_pnat RS inj_onD) 1);
+by (REPEAT (resolve_tac [Rep_pnat, pnat_Suc_RepI] 1));
+by (dtac (inj_Suc RS injD) 1);
+by (etac (inj_Rep_pnat RS injD) 1);
+qed "inj_pSuc"; 
+
+val pSuc_inject = inj_pSuc RS injD;
+
+Goal "(pSuc(m)=pSuc(n)) = (m=n)";
+by (EVERY1 [rtac iffI, etac pSuc_inject, etac arg_cong]); 
+qed "pSuc_pSuc_eq";
+
+AddIffs [pSuc_pSuc_eq];
+
+Goal "n ~= pSuc(n)";
+by (pnat_ind_tac "n" 1);
+by (ALLGOALS Asm_simp_tac);
+qed "n_not_pSuc_n";
+
+bind_thm ("pSuc_n_not_n", n_not_pSuc_n RS not_sym);
+
+Goal "!!n. n ~= 1p ==> EX m. n = pSuc m";
+by (rtac pnatE 1);
+by (REPEAT (Blast_tac 1));
+qed "not1p_implies_pSuc";
+
+Goal "pSuc m = m + 1p";
+by (auto_tac (claset(),simpset() addsimps [pnat_Suc_def,
+    pnat_one_def,Abs_pnat_inverse,pnat_add_def]));
+qed "pSuc_is_plus_one";
+
+Goal
+      "(Rep_pnat x + Rep_pnat y): pnat";
+by (cut_facts_tac [[Rep_pnat_gt_zero,
+    Rep_pnat_gt_zero] MRS add_less_mono,Collect_pnat_gt_0] 1);
+by (etac ssubst 1);
+by Auto_tac;
+qed "sum_Rep_pnat";
+
+Goalw [pnat_add_def] 
+      "Rep_pnat x + Rep_pnat y = Rep_pnat (x + y)";
+by (simp_tac (simpset() addsimps [sum_Rep_pnat RS 
+                          Abs_pnat_inverse]) 1);
+qed "sum_Rep_pnat_sum";
+
+Goalw [pnat_add_def] 
+      "(x + y) + z = x + (y + (z::pnat))";
+by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
+by (simp_tac (simpset() addsimps [sum_Rep_pnat RS 
+                Abs_pnat_inverse,add_assoc]) 1);
+qed "pnat_add_assoc";
+
+Goalw [pnat_add_def] "x + (y + z) = y + (x + (z::pnat))";
+by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
+by (simp_tac (simpset() addsimps [sum_Rep_pnat RS 
+          Abs_pnat_inverse,add_left_commute]) 1);
+qed "pnat_add_left_commute";
+
+(*Addition is an AC-operator*)
+val pnat_add_ac = [pnat_add_assoc, pnat_add_commute, pnat_add_left_commute];
+
+Goalw [pnat_add_def] "((x::pnat) + y = x + z) = (y = z)";
+by (auto_tac (claset() addDs [(inj_on_Abs_pnat RS inj_onD),
+     inj_Rep_pnat RS injD],simpset() addsimps [sum_Rep_pnat]));
+qed "pnat_add_left_cancel";
+
+Goalw [pnat_add_def] "(y + (x::pnat) = z + x) = (y = z)";
+by (auto_tac (claset() addDs [(inj_on_Abs_pnat RS inj_onD),
+     inj_Rep_pnat RS injD],simpset() addsimps [sum_Rep_pnat]));
+qed "pnat_add_right_cancel";
+
+Goalw [pnat_add_def] "!(y::pnat). x + y ~= x";
+by (rtac (Rep_pnat_inverse RS subst) 1);
+by (auto_tac (claset() addDs [(inj_on_Abs_pnat RS inj_onD)] 
+  	               addSDs [add_eq_self_zero],
+	      simpset() addsimps [sum_Rep_pnat, Rep_pnat,Abs_pnat_inverse,
+				  Rep_pnat_gt_zero RS less_not_refl2]));
+qed "pnat_no_add_ident";
+
+
+(***) (***) (***) (***) (***) (***) (***) (***) (***)
+
+  (*** pnat_less ***)
+
+Goalw [pnat_less_def] 
+      "!!x. [| x < (y::pnat); y < z |] ==> x < z";
+by ((etac less_trans 1) THEN assume_tac 1);
+qed "pnat_less_trans";
+
+Goalw [pnat_less_def] "!!x. x < (y::pnat) ==> ~ y < x";
+by (etac less_not_sym 1);
+qed "pnat_less_not_sym";
+
+(* [| x < y; y < x |] ==> P *)
+bind_thm ("pnat_less_asym",pnat_less_not_sym RS notE);
+
+Goalw [pnat_less_def] "!!x. ~ y < (y::pnat)";
+by Auto_tac;
+qed "pnat_less_not_refl";
+
+bind_thm ("pnat_less_irrefl",pnat_less_not_refl RS notE);
+
+Goalw [pnat_less_def] 
+     "!!x. x < (y::pnat) ==> x ~= y";
+by Auto_tac;
+qed "pnat_less_not_refl2";
+
+Goal "~ Rep_pnat y < 0";
+by Auto_tac;
+qed "Rep_pnat_not_less0";
+
+(*** Rep_pnat < 0 ==> P ***)
+bind_thm ("Rep_pnat_less_zeroE",Rep_pnat_not_less0 RS notE);
+
+Goal "~ Rep_pnat y < 1";
+by (auto_tac (claset(),simpset() addsimps [less_Suc_eq,
+                  Rep_pnat_gt_zero,less_not_refl2]));
+qed "Rep_pnat_not_less_one";
+
+(*** Rep_pnat < 1 ==> P ***)
+bind_thm ("Rep_pnat_less_oneE",Rep_pnat_not_less_one RS notE);
+
+Goalw [pnat_less_def] 
+     "!!x. x < (y::pnat) ==> Rep_pnat y ~= 1";
+by (auto_tac (claset(),simpset() 
+    addsimps [Rep_pnat_not_less_one] delsimps [less_one]));
+qed "Rep_pnat_gt_implies_not0";
+
+Goalw [pnat_less_def] 
+      "(x::pnat) < y | x = y | y < x";
+by (cut_facts_tac [less_linear] 1);
+by (fast_tac (claset() addIs [inj_Rep_pnat RS injD]) 1);
+qed "pnat_less_linear";
+
+Goalw [le_def] "1 <= Rep_pnat x";
+by (rtac Rep_pnat_not_less_one 1);
+qed "Rep_pnat_le_one";
+
+Goalw [pnat_less_def]
+     "!! (z1::nat). z1 < z2  ==> ? z3. z1 + Rep_pnat z3 = z2";
+by (dtac less_imp_add_positive 1);
+by (auto_tac (claset() addSIs [Abs_pnat_inverse],
+	      simpset() addsimps [Collect_pnat_gt_0]));
+qed "lemma_less_ex_sum_Rep_pnat";
+
+
+   (*** pnat_le ***)
+
+Goalw [pnat_le_def] "!!x. ~ (x::pnat) < y ==> y <= x";
+by (assume_tac 1);
+qed "pnat_leI";
+
+Goalw [pnat_le_def] "!!x. (x::pnat) <= y ==> ~ y < x";
+by (assume_tac 1);
+qed "pnat_leD";
+
+val pnat_leE = make_elim pnat_leD;
+
+Goal "(~ (x::pnat) < y) = (y <= x)";
+by (blast_tac (claset() addIs [pnat_leI] addEs [pnat_leE]) 1);
+qed "pnat_not_less_iff_le";
+
+Goalw [pnat_le_def] "!!x. ~(x::pnat) <= y ==> y < x";
+by (Blast_tac 1);
+qed "pnat_not_leE";
+
+Goalw [pnat_le_def] "!!x. (x::pnat) < y ==> x <= y";
+by (blast_tac (claset() addEs [pnat_less_asym]) 1);
+qed "pnat_less_imp_le";
+
+(** Equivalence of m<=n and  m<n | m=n **)
+
+Goalw [pnat_le_def] "!!m. m <= n ==> m < n | m=(n::pnat)";
+by (cut_facts_tac [pnat_less_linear] 1);
+by (blast_tac (claset() addEs [pnat_less_irrefl,pnat_less_asym]) 1);
+qed "pnat_le_imp_less_or_eq";
+
+Goalw [pnat_le_def] "!!m. m<n | m=n ==> m <=(n::pnat)";
+by (cut_facts_tac [pnat_less_linear] 1);
+by (blast_tac (claset() addSEs [pnat_less_irrefl] addEs [pnat_less_asym]) 1);
+qed "pnat_less_or_eq_imp_le";
+
+Goal "(m <= (n::pnat)) = (m < n | m=n)";
+by (REPEAT(ares_tac [iffI,pnat_less_or_eq_imp_le,pnat_le_imp_less_or_eq] 1));
+qed "pnat_le_eq_less_or_eq";
+
+Goal "n <= (n::pnat)";
+by (simp_tac (simpset() addsimps [pnat_le_eq_less_or_eq]) 1);
+qed "pnat_le_refl";
+
+val prems = goal thy "!!i. [| i <= j; j < k |] ==> i < (k::pnat)";
+by (dtac pnat_le_imp_less_or_eq 1);
+by (blast_tac (claset() addIs [pnat_less_trans]) 1);
+qed "pnat_le_less_trans";
+
+Goal "!!i. [| i < j; j <= k |] ==> i < (k::pnat)";
+by (dtac pnat_le_imp_less_or_eq 1);
+by (blast_tac (claset() addIs [pnat_less_trans]) 1);
+qed "pnat_less_le_trans";
+
+Goal "!!i. [| i <= j; j <= k |] ==> i <= (k::pnat)";
+by (EVERY1[dtac pnat_le_imp_less_or_eq, 
+           dtac pnat_le_imp_less_or_eq,
+           rtac pnat_less_or_eq_imp_le, 
+           blast_tac (claset() addIs [pnat_less_trans])]);
+qed "pnat_le_trans";
+
+Goal "!!m. [| m <= n; n <= m |] ==> m = (n::pnat)";
+by (EVERY1[dtac pnat_le_imp_less_or_eq, 
+           dtac pnat_le_imp_less_or_eq,
+           blast_tac (claset() addIs [pnat_less_asym])]);
+qed "pnat_le_anti_sym";
+
+Goal "(m::pnat) < n = (m <= n & m ~= n)";
+by (rtac iffI 1);
+by (rtac conjI 1);
+by (etac pnat_less_imp_le 1);
+by (etac pnat_less_not_refl2 1);
+by (blast_tac (claset() addSDs [pnat_le_imp_less_or_eq]) 1);
+qed "pnat_less_le";
+
+(** LEAST -- the least number operator **)
+
+Goal "(! m::pnat. P m --> n <= m) = (! m. m < n --> ~ P m)";
+by (blast_tac (claset() addIs [pnat_leI] addEs [pnat_leE]) 1);
+val lemma = result();
+
+(* Comment below from NatDef.ML where Least_nat_def is proved*)
+(* This is an old def of Least for nat, which is derived for compatibility *)
+Goalw [Least_def]
+  "(LEAST n::pnat. P n) == (@n. P(n) & (ALL m. m < n --> ~P(m)))";
+by (simp_tac (simpset() addsimps [lemma]) 1);
+qed "Least_pnat_def";
+
+val [prem1,prem2] = goalw thy [Least_pnat_def]
+    "[| P(k::pnat);  !!x. x<k ==> ~P(x) |] ==> (LEAST x. P(x)) = k";
+by (rtac select_equality 1);
+by (blast_tac (claset() addSIs [prem1,prem2]) 1);
+by (cut_facts_tac [pnat_less_linear] 1);
+by (blast_tac (claset() addSIs [prem1] addSDs [prem2]) 1);
+qed "pnat_Least_equality";
+
+(***) (***) (***) (***) (***) (***) (***) (***)
+
+(*** alternative definition for pnat_le ***)
+Goalw [pnat_le_def,pnat_less_def] 
+      "((m::pnat) <= n) = (Rep_pnat m <= Rep_pnat n)";
+by (auto_tac (claset() addSIs [leI] addSEs [leD],simpset()));
+qed "pnat_le_iff_Rep_pnat_le";
+
+Goal "!!k::pnat. (k + m <= k + n) = (m<=n)";
+by (simp_tac (simpset() addsimps [pnat_le_iff_Rep_pnat_le,
+                           sum_Rep_pnat_sum RS sym]) 1);
+qed "pnat_add_left_cancel_le";
+
+Goalw [pnat_less_def] "!!k::pnat. (k + m < k + n) = (m<n)";
+by (simp_tac (simpset() addsimps [sum_Rep_pnat_sum RS sym]) 1);
+qed "pnat_add_left_cancel_less";
+
+Addsimps [pnat_add_left_cancel, pnat_add_right_cancel,
+  pnat_add_left_cancel_le, pnat_add_left_cancel_less];
+
+Goal "n <= ((m + n)::pnat)";
+by (simp_tac (simpset() addsimps [pnat_le_iff_Rep_pnat_le,
+                    sum_Rep_pnat_sum RS sym,le_add2]) 1);
+qed "pnat_le_add2";
+
+Goal "n <= ((n + m)::pnat)";
+by (simp_tac (simpset() addsimps pnat_add_ac) 1);
+by (rtac pnat_le_add2 1);
+qed "pnat_le_add1";
+
+(*** "i <= j ==> i <= j + m" ***)
+bind_thm ("pnat_trans_le_add1", pnat_le_add1 RSN (2,pnat_le_trans));
+
+(*** "i <= j ==> i <= m + j" ***)
+bind_thm ("pnat_trans_le_add2", pnat_le_add2 RSN (2,pnat_le_trans));
+
+(*"i < j ==> i < j + m"*)
+bind_thm ("pnat_trans_less_add1", pnat_le_add1 RSN (2,pnat_less_le_trans));
+
+(*"i < j ==> i < m + j"*)
+bind_thm ("pnat_trans_less_add2", pnat_le_add2 RSN (2,pnat_less_le_trans));
+
+Goalw [pnat_less_def] "!!i. i+j < (k::pnat) ==> i<k";
+by (auto_tac (claset() addEs [add_lessD1],
+    simpset() addsimps [sum_Rep_pnat_sum RS sym]));
+qed "pnat_add_lessD1";
+
+Goal "!!i::pnat. ~ (i+j < i)";
+by (rtac  notI 1);
+by (etac (pnat_add_lessD1 RS pnat_less_irrefl) 1);
+qed "pnat_not_add_less1";
+
+Goal "!!i::pnat. ~ (j+i < i)";
+by (simp_tac (simpset() addsimps [pnat_add_commute, pnat_not_add_less1]) 1);
+qed "pnat_not_add_less2";
+
+AddIffs [pnat_not_add_less1, pnat_not_add_less2];
+
+Goal "!!k::pnat. m <= n ==> m <= n + k";
+by (etac pnat_le_trans 1);
+by (rtac pnat_le_add1 1);
+qed "pnat_le_imp_add_le";
+
+Goal "!!k::pnat. m < n ==> m < n + k";
+by (etac pnat_less_le_trans 1);
+by (rtac pnat_le_add1 1);
+qed "pnat_less_imp_add_less";
+
+Goal "m + k <= n --> m <= (n::pnat)";
+by (simp_tac (simpset() addsimps [pnat_le_iff_Rep_pnat_le,
+    sum_Rep_pnat_sum RS sym]) 1);
+by (fast_tac (claset() addIs [add_leD1]) 1);
+qed_spec_mp "pnat_add_leD1";
+
+Goal "!!n::pnat. m + k <= n ==> k <= n";
+by (full_simp_tac (simpset() addsimps [pnat_add_commute]) 1);
+by (etac pnat_add_leD1 1);
+qed_spec_mp "pnat_add_leD2";
+
+Goal "!!n::pnat. m + k <= n ==> m <= n & k <= n";
+by (blast_tac (claset() addDs [pnat_add_leD1, pnat_add_leD2]) 1);
+bind_thm ("pnat_add_leE", result() RS conjE);
+
+Goalw [pnat_less_def] 
+      "!!k l::pnat. [| k < l; m + l = k + n |] ==> m < n";
+by (rtac less_add_eq_less 1 THEN assume_tac 1);
+by (auto_tac (claset(),simpset() addsimps [sum_Rep_pnat_sum]));
+qed "pnat_less_add_eq_less";
+
+(* ordering on positive naturals in terms of existence of sum *)
+(* could provide alternative definition -- Gleason *)
+Goalw [pnat_less_def,pnat_add_def] 
+      "(z1::pnat) < z2 = (? z3. z1 + z3 = z2)";
+by (rtac iffI 1);
+by (res_inst_tac [("t","z2")] (Rep_pnat_inverse RS subst) 1);
+by (dtac lemma_less_ex_sum_Rep_pnat 1);
+by (etac exE 1 THEN res_inst_tac [("x","z3")] exI 1);
+by (auto_tac (claset(),simpset() addsimps [sum_Rep_pnat_sum,Rep_pnat_inverse]));
+by (res_inst_tac [("t","Rep_pnat z1")] (add_0_right RS subst) 1);
+by (auto_tac (claset(),simpset() addsimps [sum_Rep_pnat_sum RS sym,
+               Rep_pnat_gt_zero] delsimps [add_0_right]));
+qed "pnat_less_iff";
+
+Goal "(? (x::pnat). z1 + x = z2) | z1 = z2 \
+\          |(? x. z2 + x = z1)";
+by (cut_facts_tac [pnat_less_linear] 1);
+by (asm_full_simp_tac (simpset() addsimps [pnat_less_iff]) 1);
+qed "pnat_linear_Ex_eq";
+
+Goal "!!(x::pnat). x + y = z ==> x < z";
+by (rtac (pnat_less_iff RS iffD2) 1);
+by (Blast_tac 1);
+qed "pnat_eq_lessI";
+
+(*** Monotonicity of Addition ***)
+
+(*strict, in 1st argument*)
+Goalw [pnat_less_def] "!!i j k::pnat. i < j ==> i + k < j + k";
+by (auto_tac (claset() addIs [add_less_mono1],
+       simpset() addsimps [sum_Rep_pnat_sum RS sym]));
+qed "pnat_add_less_mono1";
+
+Goalw [pnat_less_def] "!!i j k::pnat. [|i < j; k < l|] ==> i + k < j + l";
+by (auto_tac (claset() addIs [add_less_mono],
+       simpset() addsimps [sum_Rep_pnat_sum RS sym]));
+qed "pnat_add_less_mono";
+
+Goalw [pnat_less_def]
+     "!!f. [| !!i j::pnat. i<j ==> f(i) < f(j);       \
+\        i <= j                                 \
+\     |] ==> f(i) <= (f(j)::pnat)";
+by (auto_tac (claset() addSDs [inj_Rep_pnat RS injD],
+             simpset() addsimps [pnat_le_iff_Rep_pnat_le,
+                                     le_eq_less_or_eq]));
+qed "pnat_less_mono_imp_le_mono";
+
+Goal "!!i j k::pnat. i<=j ==> i + k <= j + k";
+by (res_inst_tac [("f", "%j. j+k")] pnat_less_mono_imp_le_mono 1);
+by (etac pnat_add_less_mono1 1);
+by (assume_tac 1);
+qed "pnat_add_le_mono1";
+
+Goal "!!k l::pnat. [|i<=j;  k<=l |] ==> i + k <= j + l";
+by (etac (pnat_add_le_mono1 RS pnat_le_trans) 1);
+by (simp_tac (simpset() addsimps [pnat_add_commute]) 1);
+(*j moves to the end because it is free while k, l are bound*)
+by (etac pnat_add_le_mono1 1);
+qed "pnad_add_le_mono";
+
+Goal "1 * Rep_pnat n = Rep_pnat n";
+by (Asm_simp_tac 1);
+qed "Rep_pnat_mult_1";
+
+Goal "Rep_pnat n * 1 = Rep_pnat n";
+by (Asm_simp_tac 1);
+qed "Rep_pnat_mult_1_right";
+
+Goal
+      "(Rep_pnat x * Rep_pnat y): pnat";
+by (cut_facts_tac [[Rep_pnat_gt_zero,
+    Rep_pnat_gt_zero] MRS mult_less_mono1,Collect_pnat_gt_0] 1);
+by (etac ssubst 1);
+by Auto_tac;
+qed "mult_Rep_pnat";
+
+Goalw [pnat_mult_def] 
+      "Rep_pnat x * Rep_pnat y = Rep_pnat (x * y)";
+by (simp_tac (simpset() addsimps [mult_Rep_pnat RS 
+                          Abs_pnat_inverse]) 1);
+qed "mult_Rep_pnat_mult";
+
+Goalw [pnat_mult_def] "m * n = n * (m::pnat)";
+by (full_simp_tac (simpset() addsimps [mult_commute]) 1);
+qed "pnat_mult_commute";
+
+Goalw [pnat_mult_def,pnat_add_def] "(m + n)*k = (m*k) + ((n*k)::pnat)";
+by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
+by (simp_tac (simpset() addsimps [mult_Rep_pnat RS 
+                Abs_pnat_inverse,sum_Rep_pnat RS 
+             Abs_pnat_inverse, add_mult_distrib]) 1);
+qed "pnat_add_mult_distrib";
+
+Goalw [pnat_mult_def,pnat_add_def] "k*(m + n) = (k*m) + ((k*n)::pnat)";
+by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
+by (simp_tac (simpset() addsimps [mult_Rep_pnat RS 
+                Abs_pnat_inverse,sum_Rep_pnat RS 
+             Abs_pnat_inverse, add_mult_distrib2]) 1);
+qed "pnat_add_mult_distrib2";
+
+Goalw [pnat_mult_def] 
+      "(x * y) * z = x * (y * (z::pnat))";
+by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
+by (simp_tac (simpset() addsimps [mult_Rep_pnat RS 
+                Abs_pnat_inverse,mult_assoc]) 1);
+qed "pnat_mult_assoc";
+
+Goalw [pnat_mult_def] "x * (y * z) = y * (x * (z::pnat))";
+by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
+by (simp_tac (simpset() addsimps [mult_Rep_pnat RS 
+          Abs_pnat_inverse,mult_left_commute]) 1);
+qed "pnat_mult_left_commute";
+
+Goalw [pnat_mult_def] "x * (Abs_pnat 1) = x";
+by (full_simp_tac (simpset() addsimps [one_RepI RS Abs_pnat_inverse,
+                   Rep_pnat_inverse]) 1);
+qed "pnat_mult_1";
+
+Goal "Abs_pnat 1 * x = x";
+by (full_simp_tac (simpset() addsimps [pnat_mult_1,
+                   pnat_mult_commute]) 1);
+qed "pnat_mult_1_left";
+
+(*Multiplication is an AC-operator*)
+val pnat_mult_ac = [pnat_mult_assoc, pnat_mult_commute, pnat_mult_left_commute];
+
+Goal "!!i j k::pnat. i<=j ==> i * k <= j * k";
+by (asm_full_simp_tac (simpset() addsimps [pnat_le_iff_Rep_pnat_le,
+                     mult_Rep_pnat_mult RS sym,mult_le_mono1]) 1);
+qed "pnat_mult_le_mono1";
+
+Goal "!!i::pnat. [| i<=j; k<=l |] ==> i*k<=j*l";
+by (asm_full_simp_tac (simpset() addsimps [pnat_le_iff_Rep_pnat_le,
+                     mult_Rep_pnat_mult RS sym,mult_le_mono]) 1);
+qed "pnat_mult_le_mono";
+
+Goal "!!i::pnat. i<j ==> k*i < k*j";
+by (asm_full_simp_tac (simpset() addsimps [pnat_less_def,
+    mult_Rep_pnat_mult RS sym,Rep_pnat_gt_zero,mult_less_mono2]) 1);
+qed "pnat_mult_less_mono2";
+
+Goal "!!i::pnat. i<j ==> i*k < j*k";
+by (dtac pnat_mult_less_mono2 1);
+by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [pnat_mult_commute])));
+qed "pnat_mult_less_mono1";
+
+Goalw [pnat_less_def] "(m*(k::pnat) < n*k) = (m<n)";
+by (asm_full_simp_tac (simpset() addsimps [mult_Rep_pnat_mult 
+              RS sym,Rep_pnat_gt_zero]) 1);
+qed "pnat_mult_less_cancel2";
+
+Goalw [pnat_less_def] "((k::pnat)*m < k*n) = (m<n)";
+by (asm_full_simp_tac (simpset() addsimps [mult_Rep_pnat_mult 
+              RS sym,Rep_pnat_gt_zero]) 1);
+qed "pnat_mult_less_cancel1";
+
+Addsimps [pnat_mult_less_cancel1, pnat_mult_less_cancel2];
+
+Goalw [pnat_mult_def]  "(m*(k::pnat) = n*k) = (m=n)";
+by (auto_tac (claset() addSDs [inj_on_Abs_pnat RS inj_onD, 
+    inj_Rep_pnat RS injD] addIs [mult_Rep_pnat], 
+    simpset() addsimps [Rep_pnat_gt_zero RS mult_cancel2]));
+qed "pnat_mult_cancel2";
+
+Goal "((k::pnat)*m = k*n) = (m=n)";
+by (rtac (pnat_mult_cancel2 RS subst) 1);
+by (auto_tac (claset () addIs [pnat_mult_commute RS subst],simpset()));
+qed "pnat_mult_cancel1";
+
+Addsimps [pnat_mult_cancel1, pnat_mult_cancel2];
+
+Goal
+     "!!(z1::pnat). z2*z3 = z4*z5  ==> z2*(z1*z3) = z4*(z1*z5)";
+by (auto_tac (claset() addIs [pnat_mult_cancel1 RS iffD2],
+    simpset() addsimps [pnat_mult_left_commute]));
+qed "pnat_same_multI2";
+
+val [prem] = goal thy
+    "(!!u. z = Abs_pnat(u) ==> P) ==> P";
+by (cut_inst_tac [("x1","z")] 
+    (rewrite_rule [pnat_def] (Rep_pnat RS Abs_pnat_inverse)) 1);
+by (res_inst_tac [("u","Rep_pnat z")] prem 1);
+by (dtac (inj_Rep_pnat RS injD) 1);
+by (Asm_simp_tac 1);
+qed "eq_Abs_pnat";
+
+(** embedding of naturals in positive naturals **)
+
+(* pnat_one_eq! *)
+Goalw [pnat_nat_def,pnat_one_def]"1p = *#0";
+by (Full_simp_tac 1);
+qed "pnat_one_iff";
+
+Goalw [pnat_nat_def,pnat_one_def,pnat_add_def] "1p + 1p = *#1";
+by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
+by (auto_tac (claset() addIs [(gt_0_mem_pnat RS Abs_pnat_inverse RS ssubst)],
+    simpset()));
+qed "pnat_two_eq";
+
+Goal "inj(pnat_nat)";
+by (rtac injI 1);
+by (rewtac pnat_nat_def);
+by (dtac (inj_on_Abs_pnat RS inj_onD) 1);
+by (auto_tac (claset() addSIs [gt_0_mem_pnat],simpset()));
+qed "inj_pnat_nat";
+
+Goal "0 < n + 1";
+by Auto_tac;
+qed "nat_add_one_less";
+
+Goal "0 < n1 + n2 + 1";
+by Auto_tac;
+qed "nat_add_one_less1";
+
+(* this worked with one call to auto_tac before! *)
+Goalw [pnat_add_def,pnat_nat_def,pnat_one_def] 
+          "*#n1 + *#n2 = *#(n1 + n2) + 1p";
+by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
+by (rtac (gt_0_mem_pnat RS Abs_pnat_inverse RS ssubst) 1);
+by (rtac (gt_0_mem_pnat RS Abs_pnat_inverse RS ssubst) 2);
+by (rtac (gt_0_mem_pnat RS Abs_pnat_inverse RS ssubst) 3);
+by (rtac (gt_0_mem_pnat RS Abs_pnat_inverse RS ssubst) 4);
+by (auto_tac (claset(),
+	      simpset() addsimps [sum_Rep_pnat_sum,
+				  nat_add_one_less,nat_add_one_less1]));
+qed "pnat_nat_add";
+
+Goalw [pnat_nat_def,pnat_less_def] "(n < m) = (*#n < *#m)";
+by (auto_tac (claset(),simpset() 
+    addsimps [Abs_pnat_inverse,Collect_pnat_gt_0]));
+qed "pnat_nat_less_iff";
+
+Addsimps [pnat_nat_less_iff RS sym];
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/PNat.thy	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,73 @@
+(*  Title       : PNat.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : The positive naturals
+*) 
+
+
+PNat = Arith +
+
+(** type pnat **)
+
+(* type definition *)
+
+typedef
+  pnat = "lfp(%X. {1} Un (Suc``X))"   (lfp_def)
+
+instance
+   pnat :: {ord, plus, times}
+
+consts
+
+  pSuc       :: pnat => pnat
+  "1p"       :: pnat                ("1p")
+
+constdefs
+  
+  pnat_nat  :: nat => pnat                  ("*# _" [80] 80) 
+  "*# n     == Abs_pnat(n + 1)"
+ 
+defs
+
+  pnat_one_def      "1p == Abs_pnat(1)"
+  pnat_Suc_def      "pSuc == (%n. Abs_pnat(Suc(Rep_pnat(n))))"
+
+
+  pnat_add_def
+       "x + y == Abs_pnat(Rep_pnat(x) +  Rep_pnat(y))"
+
+  pnat_mult_def
+       "x * y == Abs_pnat(Rep_pnat(x) * Rep_pnat(y))"
+
+ pnat_less_def
+       "x < (y::pnat) == Rep_pnat(x) < Rep_pnat(y)"
+
+ pnat_le_def
+       "x <= (y::pnat) ==  ~(y < x)"
+
+end
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/PRat.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,837 @@
+(*  Title       : PRat.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : The positive rationals
+*) 
+
+open PRat;
+
+Delrules [equalityI];
+
+(*** Many theorems similar to those in Integ.thy ***)
+(*** Proving that ratrel is an equivalence relation ***)
+
+Goal
+    "!! x1. [| (x1::pnat) * y2 = x2 * y1; x2 * y3 = x3 * y2 |] \
+\            ==> x1 * y3 = x3 * y1";        
+by (res_inst_tac [("k1","y2")] (pnat_mult_cancel1 RS iffD1) 1);
+by (auto_tac (claset(), simpset() addsimps [pnat_mult_assoc RS sym]));
+by (auto_tac (claset(),simpset() addsimps [pnat_mult_commute]));
+by (dres_inst_tac [("s","x2 * y3")] sym 1);
+by (asm_simp_tac (simpset() addsimps [pnat_mult_left_commute,
+    pnat_mult_commute]) 1);
+qed "prat_trans_lemma";
+
+(** Natural deduction for ratrel **)
+
+Goalw [ratrel_def]
+    "(((x1,y1),(x2,y2)): ratrel) = (x1 * y2 = x2 * y1)";
+by (Fast_tac 1);
+qed "ratrel_iff";
+
+Goalw [ratrel_def]
+    "!!x1 x2. [| x1 * y2 = x2 * y1 |] ==> ((x1,y1),(x2,y2)): ratrel";
+by (Fast_tac  1);
+qed "ratrelI";
+
+Goalw [ratrel_def]
+  "p: ratrel --> (EX x1 y1 x2 y2. \
+\                  p = ((x1,y1),(x2,y2)) & x1 *y2 = x2 *y1)";
+by (Fast_tac 1);
+qed "ratrelE_lemma";
+
+val [major,minor] = goal thy
+  "[| p: ratrel;  \
+\     !!x1 y1 x2 y2. [| p = ((x1,y1),(x2,y2));  x1*y2 = x2*y1 \
+\                    |] ==> Q |] ==> Q";
+by (cut_facts_tac [major RS (ratrelE_lemma RS mp)] 1);
+by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
+qed "ratrelE";
+
+AddSIs [ratrelI];
+AddSEs [ratrelE];
+
+Goal "(x,x): ratrel";
+by (stac surjective_pairing 1 THEN rtac (refl RS ratrelI) 1);
+qed "ratrel_refl";
+
+Goalw [equiv_def, refl_def, sym_def, trans_def]
+    "equiv {x::(pnat*pnat).True} ratrel";
+by (fast_tac (claset() addSIs [ratrel_refl] 
+                      addSEs [sym, prat_trans_lemma]) 1);
+qed "equiv_ratrel";
+
+val equiv_ratrel_iff =
+    [TrueI, TrueI] MRS 
+    ([CollectI, CollectI] MRS 
+    (equiv_ratrel RS eq_equiv_class_iff));
+
+Goalw  [prat_def,ratrel_def,quotient_def] "ratrel^^{(x,y)}:prat";
+by (Blast_tac 1);
+qed "ratrel_in_prat";
+
+Goal "inj_on Abs_prat prat";
+by (rtac inj_on_inverseI 1);
+by (etac Abs_prat_inverse 1);
+qed "inj_on_Abs_prat";
+
+Addsimps [equiv_ratrel_iff,inj_on_Abs_prat RS inj_on_iff,
+          ratrel_iff, ratrel_in_prat, Abs_prat_inverse];
+
+Addsimps [equiv_ratrel RS eq_equiv_class_iff];
+val eq_ratrelD = equiv_ratrel RSN (2,eq_equiv_class);
+
+Goal "inj(Rep_prat)";
+by (rtac inj_inverseI 1);
+by (rtac Rep_prat_inverse 1);
+qed "inj_Rep_prat";
+
+(** prat_pnat: the injection from pnat to prat **)
+Goal "inj(prat_pnat)";
+by (rtac injI 1);
+by (rewtac prat_pnat_def);
+by (dtac (inj_on_Abs_prat RS inj_onD) 1);
+by (REPEAT (rtac ratrel_in_prat 1));
+by (dtac eq_equiv_class 1);
+by (rtac equiv_ratrel 1);
+by (Fast_tac 1);
+by Safe_tac;
+by (Asm_full_simp_tac 1);
+qed "inj_prat_pnat";
+
+(* lcp's original eq_Abs_Integ *)
+val [prem] = goal thy
+    "(!!x y. z = Abs_prat(ratrel^^{(x,y)}) ==> P) ==> P";
+by (res_inst_tac [("x1","z")] 
+    (rewrite_rule [prat_def] Rep_prat RS quotientE) 1);
+by (dres_inst_tac [("f","Abs_prat")] arg_cong 1);
+by (res_inst_tac [("p","x")] PairE 1);
+by (rtac prem 1);
+by (asm_full_simp_tac (simpset() addsimps [Rep_prat_inverse]) 1);
+qed "eq_Abs_prat";
+
+(**** qinv: inverse on prat ****)
+
+Goalw [congruent_def]
+  "congruent ratrel (%p. split (%x y. ratrel^^{(y,x)}) p)";
+by Safe_tac;
+by (asm_full_simp_tac (simpset() addsimps [pnat_mult_commute]) 1);
+qed "qinv_congruent";
+
+(*Resolve th against the corresponding facts for qinv*)
+val qinv_ize = RSLIST [equiv_ratrel, qinv_congruent];
+
+Goalw [qinv_def]
+      "qinv (Abs_prat(ratrel^^{(x,y)})) = Abs_prat(ratrel ^^ {(y,x)})";
+by (res_inst_tac [("f","Abs_prat")] arg_cong 1);
+by (simp_tac (simpset() addsimps 
+   [ratrel_in_prat RS Abs_prat_inverse,qinv_ize UN_equiv_class]) 1);
+qed "qinv";
+
+Goal "qinv (qinv z) = z";
+by (res_inst_tac [("z","z")] eq_Abs_prat 1);
+by (asm_simp_tac (simpset() addsimps [qinv]) 1);
+qed "qinv_qinv";
+
+Goal "inj(qinv)";
+by (rtac injI 1);
+by (dres_inst_tac [("f","qinv")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [qinv_qinv]) 1);
+qed "inj_qinv";
+
+Goalw [prat_pnat_def] "qinv($# (Abs_pnat 1)) = $#(Abs_pnat 1)";
+by (simp_tac (simpset() addsimps [qinv]) 1);
+qed "qinv_1";
+
+Goal 
+     "!!(x1::pnat). [| x1 * y2 = x2 * y1 |] ==> \
+\     (x * y1 + x1 * ya) * (ya * y2) = (x * y2 + x2 * ya) * (ya * y1)";
+by (auto_tac (claset() addSIs [pnat_same_multI2],
+       simpset() addsimps [pnat_add_mult_distrib,
+       pnat_mult_assoc]));
+by (res_inst_tac [("n1","y2")] (pnat_mult_commute RS subst) 1);
+by (auto_tac (claset() addIs [pnat_add_left_cancel RS iffD2],simpset() addsimps pnat_mult_ac));
+by (res_inst_tac [("y1","x1")] (pnat_mult_left_commute RS subst) 1);
+by (res_inst_tac [("y1","x1")] (pnat_mult_left_commute RS ssubst) 1);
+by (auto_tac (claset(),simpset() addsimps [pnat_mult_assoc RS sym]));
+qed "prat_add_congruent2_lemma";
+
+Goal 
+    "congruent2 ratrel (%p1 p2.                  \
+\         split (%x1 y1. split (%x2 y2. ratrel^^{(x1*y2 + x2*y1, y1*y2)}) p2) p1)";
+by (rtac (equiv_ratrel RS congruent2_commuteI) 1);
+by Safe_tac;
+by (rewtac split_def);
+by (asm_simp_tac (simpset() addsimps [pnat_mult_commute,pnat_add_commute]) 1);
+by (auto_tac (claset(),simpset() addsimps [prat_add_congruent2_lemma]));
+qed "prat_add_congruent2";
+
+(*Resolve th against the corresponding facts for prat_add*)
+val prat_add_ize = RSLIST [equiv_ratrel, prat_add_congruent2];
+
+Goalw [prat_add_def]
+   "Abs_prat((ratrel^^{(x1,y1)})) + Abs_prat((ratrel^^{(x2,y2)})) =   \
+\   Abs_prat(ratrel ^^ {(x1*y2 + x2*y1, y1*y2)})";
+by (simp_tac (simpset() addsimps [prat_add_ize UN_equiv_class2]) 1);
+qed "prat_add";
+
+Goal "(z::prat) + w = w + z";
+by (res_inst_tac [("z","z")] eq_Abs_prat 1);
+by (res_inst_tac [("z","w")] eq_Abs_prat 1);
+by (asm_simp_tac (simpset() addsimps ([prat_add] @ pnat_add_ac @ pnat_mult_ac)) 1);
+qed "prat_add_commute";
+
+Goal "((z1::prat) + z2) + z3 = z1 + (z2 + z3)";
+by (res_inst_tac [("z","z1")] eq_Abs_prat 1);
+by (res_inst_tac [("z","z2")] eq_Abs_prat 1);
+by (res_inst_tac [("z","z3")] eq_Abs_prat 1);
+by (asm_simp_tac (simpset() addsimps ([pnat_add_mult_distrib2,prat_add] @ 
+                                     pnat_add_ac @ pnat_mult_ac)) 1);
+qed "prat_add_assoc";
+
+qed_goal "prat_add_left_commute" thy
+    "(z1::prat) + (z2 + z3) = z2 + (z1 + z3)"
+ (fn _ => [rtac (prat_add_commute RS trans) 1, rtac (prat_add_assoc RS trans) 1,
+           rtac (prat_add_commute RS arg_cong) 1]);
+
+(* Positive Rational addition is an AC operator *)
+val prat_add_ac = [prat_add_assoc, prat_add_commute, prat_add_left_commute];
+
+
+(*** Congruence property for multiplication ***)
+
+Goalw [congruent2_def]
+    "congruent2 ratrel (%p1 p2.                  \
+\         split (%x1 y1. split (%x2 y2. ratrel^^{(x1*x2, y1*y2)}) p2) p1)";
+
+(*Proof via congruent2_commuteI seems longer*)
+by Safe_tac;
+by (asm_simp_tac (simpset() addsimps [pnat_mult_assoc]) 1);
+(*The rest should be trivial, but rearranging terms is hard*)
+by (res_inst_tac [("x1","x1a")] (pnat_mult_left_commute RS ssubst) 1);
+by (asm_simp_tac (simpset() addsimps [pnat_mult_assoc RS sym]) 1);
+by (asm_simp_tac (simpset() addsimps pnat_mult_ac) 1);
+qed "pnat_mult_congruent2";
+
+(*Resolve th against the corresponding facts for pnat_mult*)
+val prat_mult_ize = RSLIST [equiv_ratrel, pnat_mult_congruent2];
+
+Goalw [prat_mult_def]
+  "Abs_prat(ratrel^^{(x1,y1)}) * Abs_prat(ratrel^^{(x2,y2)}) = \
+\  Abs_prat(ratrel^^{(x1*x2, y1*y2)})";
+by (asm_simp_tac
+    (simpset() addsimps [prat_mult_ize UN_equiv_class2]) 1);
+qed "prat_mult";
+
+Goal "(z::prat) * w = w * z";
+by (res_inst_tac [("z","z")] eq_Abs_prat 1);
+by (res_inst_tac [("z","w")] eq_Abs_prat 1);
+by (asm_simp_tac (simpset() addsimps (pnat_mult_ac @ [prat_mult])) 1);
+qed "prat_mult_commute";
+
+Goal "((z1::prat) * z2) * z3 = z1 * (z2 * z3)";
+by (res_inst_tac [("z","z1")] eq_Abs_prat 1);
+by (res_inst_tac [("z","z2")] eq_Abs_prat 1);
+by (res_inst_tac [("z","z3")] eq_Abs_prat 1);
+by (asm_simp_tac (simpset() addsimps [prat_mult, pnat_mult_assoc]) 1);
+qed "prat_mult_assoc";
+
+(*For AC rewriting*)
+Goal "(x::prat)*(y*z)=y*(x*z)";
+by (rtac (prat_mult_commute RS trans) 1);
+by (rtac (prat_mult_assoc RS trans) 1);
+by (rtac (prat_mult_commute RS arg_cong) 1);
+qed "prat_mult_left_commute";
+
+(*Positive Rational multiplication is an AC operator*)
+val prat_mult_ac = [prat_mult_assoc,prat_mult_commute,prat_mult_left_commute];
+
+Goalw [prat_pnat_def] "($#Abs_pnat 1) * z = z";
+by (res_inst_tac [("z","z")] eq_Abs_prat 1);
+by (asm_full_simp_tac (simpset() addsimps [prat_mult] @ pnat_mult_ac) 1);
+qed "prat_mult_1";
+
+Goalw [prat_pnat_def] "z * ($#Abs_pnat 1) = z";
+by (res_inst_tac [("z","z")] eq_Abs_prat 1);
+by (asm_full_simp_tac (simpset() addsimps [prat_mult] @ pnat_mult_ac) 1);
+qed "prat_mult_1_right";
+
+Goalw [prat_pnat_def] 
+            "$#((z1::pnat) + z2) = $#z1 + $#z2";
+by (asm_simp_tac (simpset() addsimps [prat_add,
+       pnat_add_mult_distrib,pnat_mult_1]) 1);
+qed "prat_pnat_add";
+
+Goalw [prat_pnat_def] 
+            "$#((z1::pnat) * z2) = $#z1 * $#z2";
+by (asm_simp_tac (simpset() addsimps [prat_mult,
+                              pnat_mult_1]) 1);
+qed "prat_pnat_mult";
+
+(*** prat_mult and qinv ***)
+
+Goalw [prat_def,prat_pnat_def] "qinv (q) * q = $# (Abs_pnat 1)";
+by (res_inst_tac [("z","q")] eq_Abs_prat 1);
+by (asm_full_simp_tac (simpset() addsimps [qinv,
+        prat_mult,pnat_mult_1,pnat_mult_1_left,
+                        pnat_mult_commute]) 1);
+qed "prat_mult_qinv";
+
+Goal "q * qinv (q) = $# (Abs_pnat 1)";
+by (rtac (prat_mult_commute RS subst) 1);
+by (simp_tac (simpset() addsimps [prat_mult_qinv]) 1);
+qed "prat_mult_qinv_right";
+
+Goal "? y. (x::prat) * y = $# Abs_pnat 1";
+by (fast_tac (claset() addIs [prat_mult_qinv_right]) 1);
+qed "prat_qinv_ex";
+
+Goal "?! y. (x::prat) * y = $# Abs_pnat 1";
+by (auto_tac (claset() addIs [prat_mult_qinv_right],simpset()));
+by (dres_inst_tac [("f","%x. ya*x")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [prat_mult_assoc RS sym]) 1);
+by (asm_full_simp_tac (simpset() addsimps [prat_mult_commute,
+    prat_mult_1,prat_mult_1_right]) 1);
+qed "prat_qinv_ex1";
+
+Goal "?! y. y * (x::prat) = $# Abs_pnat 1";
+by (auto_tac (claset() addIs [prat_mult_qinv],simpset()));
+by (dres_inst_tac [("f","%x. x*ya")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [prat_mult_assoc]) 1);
+by (asm_full_simp_tac (simpset() addsimps [prat_mult_commute,
+    prat_mult_1,prat_mult_1_right]) 1);
+qed "prat_qinv_left_ex1";
+
+Goal "!!y. x * y = $# Abs_pnat 1 ==> x = qinv y";
+by (cut_inst_tac [("q","y")] prat_mult_qinv 1);
+by (res_inst_tac [("x1","y")] (prat_qinv_left_ex1 RS ex1E) 1);
+by (Blast_tac 1);
+qed "prat_mult_inv_qinv";
+
+Goal "? y. x = qinv y";
+by (cut_inst_tac [("x","x")] prat_qinv_ex 1);
+by (etac exE 1 THEN dtac prat_mult_inv_qinv 1);
+by (Fast_tac 1);
+qed "prat_as_inverse_ex";
+
+Goal "qinv(x*y) = qinv(x)*qinv(y)";
+by (res_inst_tac [("z","x")] eq_Abs_prat 1);
+by (res_inst_tac [("z","y")] eq_Abs_prat 1);
+by (auto_tac (claset(),simpset() addsimps [qinv,prat_mult]));
+qed "qinv_mult_eq";
+
+(** Lemmas **)
+
+qed_goal "prat_add_assoc_cong" thy
+    "!!z. (z::prat) + v = z' + v' ==> z + (v + w) = z' + (v' + w)"
+ (fn _ => [(asm_simp_tac (simpset() addsimps [prat_add_assoc RS sym]) 1)]);
+
+qed_goal "prat_add_assoc_swap" thy "(z::prat) + (v + w) = v + (z + w)"
+ (fn _ => [(REPEAT (ares_tac [prat_add_commute RS prat_add_assoc_cong] 1))]);
+
+Goal "((z1::prat) + z2) * w = (z1 * w) + (z2 * w)";
+by (res_inst_tac [("z","z1")] eq_Abs_prat 1);
+by (res_inst_tac [("z","z2")] eq_Abs_prat 1);
+by (res_inst_tac [("z","w")] eq_Abs_prat 1);
+by (asm_simp_tac 
+    (simpset() addsimps ([pnat_add_mult_distrib2, prat_add, prat_mult] @ 
+                        pnat_add_ac @ pnat_mult_ac)) 1);
+qed "prat_add_mult_distrib";
+
+val prat_mult_commute'= read_instantiate [("z","w")] prat_mult_commute;
+
+Goal "(w::prat) * (z1 + z2) = (w * z1) + (w * z2)";
+by (simp_tac (simpset() addsimps [prat_mult_commute',prat_add_mult_distrib]) 1);
+qed "prat_add_mult_distrib2";
+
+val prat_mult_simps = [prat_mult_1, prat_mult_1_right, 
+                       prat_mult_qinv, prat_mult_qinv_right];
+Addsimps prat_mult_simps;
+
+      (*** theorems for ordering ***)
+(* prove introduction and elimination rules for prat_less *)
+
+Goalw [prat_less_def]
+    "Q1 < (Q2::prat) = (EX Q3. Q1 + Q3 = Q2)";
+by (Fast_tac 1);
+qed "prat_less_iff";
+
+Goalw [prat_less_def]
+      "!!(Q1::prat). Q1 + Q3 = Q2 ==> Q1 < Q2";
+by (Fast_tac  1);
+qed "prat_lessI";
+
+(* ordering on positive fractions in terms of existence of sum *)
+Goalw [prat_less_def]
+      "Q1 < (Q2::prat) --> (EX Q3. Q1 + Q3 = Q2)";
+by (Fast_tac 1);
+qed "prat_lessE_lemma";
+
+Goal 
+     "!! Q1. [| Q1 < (Q2::prat); \
+\        !! (Q3::prat). Q1 + Q3 = Q2 ==> P |] \
+\     ==> P";
+by (dtac (prat_lessE_lemma RS mp) 1);
+by Auto_tac;
+qed "prat_lessE";
+
+(* qless is a strong order i.e nonreflexive and transitive *)
+Goal 
+     "!!(q1::prat). [| q1 < q2; q2 < q3 |] ==> q1 < q3";
+by (REPEAT(dtac (prat_lessE_lemma RS mp) 1));
+by (REPEAT(etac exE 1));
+by (hyp_subst_tac 1);
+by (res_inst_tac [("Q3.0","Q3 + Q3a")] prat_lessI 1);
+by (auto_tac (claset(),simpset() addsimps [prat_add_assoc]));
+qed "prat_less_trans";
+
+Goal "~q < (q::prat)";
+by (EVERY1[rtac notI, dtac (prat_lessE_lemma RS mp)]);
+by (res_inst_tac [("z","q")] eq_Abs_prat 1);
+by (res_inst_tac [("z","Q3")] eq_Abs_prat 1);
+by (etac exE 1 THEN res_inst_tac [("z","Q3a")] eq_Abs_prat 1);
+by (REPEAT(hyp_subst_tac 1));
+by (asm_full_simp_tac (simpset() addsimps [prat_add,
+    pnat_no_add_ident,pnat_add_mult_distrib2] @ pnat_mult_ac) 1);
+qed "prat_less_not_refl";
+
+(*** y < y ==> P ***)
+bind_thm("prat_less_irrefl",prat_less_not_refl RS notE);
+
+Goal "!! (q1::prat). [| q1 < q2; q2 < q1 |] ==> P";
+by (dtac prat_less_trans 1 THEN assume_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [prat_less_not_refl]) 1);
+qed "prat_less_asym";
+
+Goal "!! (q1::prat). q1 < q2 ==> ~ q2 < q1";
+by (auto_tac (claset() addSDs [prat_less_asym],simpset()));
+qed "prat_less_not_sym";
+
+(* half of positive fraction exists- Gleason p. 120- Proposition 9-2.6(i)*)
+Goal "!(q::prat). ? x. x + x = q";
+by (rtac allI 1);
+by (res_inst_tac [("z","q")] eq_Abs_prat 1);
+by (res_inst_tac [("x","Abs_prat (ratrel ^^ {(x, y+y)})")] exI 1);
+by (auto_tac (claset(),simpset() addsimps 
+              [prat_add,pnat_mult_assoc RS sym,pnat_add_mult_distrib,
+               pnat_add_mult_distrib2]));
+qed "lemma_prat_dense";
+
+Goal "? (x::prat). x + x = q";
+by (res_inst_tac [("z","q")] eq_Abs_prat 1);
+by (res_inst_tac [("x","Abs_prat (ratrel ^^ {(x, y+y)})")] exI 1);
+by (auto_tac (claset(),simpset() addsimps 
+              [prat_add,pnat_mult_assoc RS sym,pnat_add_mult_distrib,
+               pnat_add_mult_distrib2]));
+qed "prat_lemma_dense";
+
+(* there exists a number between any two positive fractions *)
+(* Gleason p. 120- Proposition 9-2.6(iv) *)
+Goalw [prat_less_def] 
+      "!! (q1::prat). q1 < q2 ==> ? x. q1 < x & x < q2";
+by (auto_tac (claset() addIs [lemma_prat_dense],simpset()));
+by (res_inst_tac [("x","T")] (lemma_prat_dense RS allE) 1);
+by (etac exE 1);
+by (res_inst_tac [("x","q1 + x")] exI 1);
+by (auto_tac (claset() addIs [prat_lemma_dense],
+    simpset() addsimps [prat_add_assoc]));
+qed "prat_dense";
+
+(* ordering of addition for positive fractions *)
+Goalw [prat_less_def] 
+      "!!(q1::prat). q1 < q2 ==> q1 + x < q2 + x";
+by (Step_tac 1);
+by (res_inst_tac [("x","T")] exI 1);
+by (auto_tac (claset(),simpset() addsimps prat_add_ac));
+qed "prat_add_less2_mono1";
+
+Goal  
+      "!!(q1::prat). q1 < q2 ==> x + q1 < x + q2";
+by (auto_tac (claset() addIs [prat_add_less2_mono1],
+    simpset() addsimps [prat_add_commute]));
+qed "prat_add_less2_mono2";
+
+(* ordering of multiplication for positive fractions *)
+Goalw [prat_less_def] 
+      "!!(q1::prat). q1 < q2 ==> q1 * x < q2 * x";
+by (Step_tac 1);
+by (res_inst_tac [("x","T*x")] exI 1);
+by (auto_tac (claset(),simpset() addsimps [prat_add_mult_distrib]));
+qed "prat_mult_less2_mono1";
+
+Goal "!!(q1::prat). q1 < q2  ==> x * q1 < x * q2";
+by (auto_tac (claset() addDs [prat_mult_less2_mono1],
+    simpset() addsimps [prat_mult_commute]));
+qed "prat_mult_left_less2_mono1";
+
+(* there is no smallest positive fraction *)
+Goalw [prat_less_def] "? (x::prat). x < y";
+by (cut_facts_tac [lemma_prat_dense] 1);
+by (Fast_tac 1);
+qed "qless_Ex";
+
+(* lemma for proving $< is linear *)
+Goalw [prat_def,prat_less_def] 
+      "ratrel ^^ {(x, y * ya)} : {p::(pnat*pnat).True}/ratrel";
+by (asm_full_simp_tac (simpset() addsimps [ratrel_def,quotient_def]) 1);
+by (Blast_tac 1);
+qed "lemma_prat_less_linear";
+
+(* linearity of < -- Gleason p. 120 - Proposition 9-2.6 *)
+(*** FIXME Proof long ***)
+Goalw [prat_less_def] 
+      "(q1::prat) < q2 | q1 = q2 | q2 < q1";
+by (res_inst_tac [("z","q1")] eq_Abs_prat 1);
+by (res_inst_tac [("z","q2")] eq_Abs_prat 1);
+by (Step_tac 1 THEN REPEAT(dtac (not_ex RS iffD1) 1) 
+               THEN Auto_tac);
+by (cut_inst_tac  [("z1.0","x*ya"),
+   ("z2.0","xa*y")] pnat_linear_Ex_eq 1);
+by (EVERY1[etac disjE,etac exE]);
+by (eres_inst_tac 
+    [("x","Abs_prat(ratrel^^{(xb,ya*y)})")] allE 1);
+by (asm_full_simp_tac 
+    (simpset() addsimps [prat_add, pnat_mult_assoc 
+     RS sym,pnat_add_mult_distrib RS sym]) 1);
+by (EVERY1[asm_full_simp_tac (simpset() addsimps pnat_mult_ac),
+    etac disjE, assume_tac, etac exE]);
+by (thin_tac "!T. Abs_prat (ratrel ^^ {(x, y)}) + T ~= \
+\     Abs_prat (ratrel ^^ {(xa, ya)})" 1);
+by (eres_inst_tac [("x","Abs_prat(ratrel^^{(xb,y*ya)})")] allE 1);
+by (asm_full_simp_tac (simpset() addsimps [prat_add,
+      pnat_mult_assoc RS sym,pnat_add_mult_distrib RS sym]) 1);
+by (asm_full_simp_tac (simpset() addsimps pnat_mult_ac) 1);
+qed "prat_linear";
+
+Goal
+    "!!(q1::prat). [| q1 < q2 ==> P;  q1 = q2 ==> P; \
+\          q2 < q1 ==> P |] ==> P";
+by (cut_inst_tac [("q1.0","q1"),("q2.0","q2")] prat_linear 1);
+by Auto_tac;
+qed "prat_linear_less2";
+
+(* Gleason p. 120 -- 9-2.6 (iv) *)
+Goal 
+ "!!(q1::prat). [| q1 < q2; qinv(q1) = qinv(q2) |] ==> P";
+by (cut_inst_tac [("x","qinv (q2)"),("q1.0","q1"),
+   ("q2.0","q2")] prat_mult_less2_mono1 1);
+by (assume_tac 1);
+by (Asm_full_simp_tac 1 THEN dtac sym 1);
+by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl]));
+qed "lemma1_qinv_prat_less";
+
+Goal 
+ "!!(q1::prat). [| q1 < q2; qinv(q1) < qinv(q2) |] ==> P";
+by (cut_inst_tac [("x","qinv (q2)"),("q1.0","q1"),
+   ("q2.0","q2")] prat_mult_less2_mono1 1);
+by (assume_tac 1);
+by (cut_inst_tac [("x","q1"),("q1.0","qinv (q1)"),
+   ("q2.0","qinv (q2)")] prat_mult_left_less2_mono1 1);
+by Auto_tac;
+by (dres_inst_tac [("q2.0","$#Abs_pnat 1")] prat_less_trans 1);
+by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl]));
+qed "lemma2_qinv_prat_less";
+
+Goal 
+      "!!(q1::prat). q1 < q2  ==> qinv (q2) < qinv (q1)";
+by (res_inst_tac [("q2.0","qinv q1"),
+         ("q1.0","qinv q2")] prat_linear_less2 1);
+by (auto_tac (claset() addEs [lemma1_qinv_prat_less,
+                 lemma2_qinv_prat_less],simpset()));
+qed "qinv_prat_less";
+
+Goal "!!(q1::prat). q1 < $#Abs_pnat 1 ==> $#Abs_pnat 1 < qinv(q1)";
+by (dtac qinv_prat_less 1);
+by (full_simp_tac (simpset() addsimps [qinv_1]) 1);
+qed "prat_qinv_gt_1";
+
+Goalw [pnat_one_def] "!!(q1::prat). q1 < $#1p ==> $#1p < qinv(q1)";
+by (etac prat_qinv_gt_1 1);
+qed "prat_qinv_is_gt_1";
+
+Goalw [prat_less_def] "$#Abs_pnat 1 < $#Abs_pnat 1 + $#Abs_pnat 1";
+by (Fast_tac 1); 
+qed "prat_less_1_2";
+
+Goal "qinv($#Abs_pnat 1 + $#Abs_pnat 1) < $#Abs_pnat 1";
+by (cut_facts_tac [prat_less_1_2 RS qinv_prat_less] 1);
+by (asm_full_simp_tac (simpset() addsimps [qinv_1]) 1);
+qed "prat_less_qinv_2_1";
+
+Goal "!!(x::prat). x < y ==> x*qinv(y) < $#Abs_pnat 1";
+by (dres_inst_tac [("x","qinv(y)")] prat_mult_less2_mono1 1);
+by (Asm_full_simp_tac 1);
+qed "prat_mult_qinv_less_1";
+
+Goal "(x::prat) < x + x";
+by (cut_inst_tac [("x","x")] 
+    (prat_less_1_2 RS prat_mult_left_less2_mono1) 1);
+by (asm_full_simp_tac (simpset() addsimps 
+    [prat_add_mult_distrib2]) 1);
+qed "prat_self_less_add_self";
+
+Goalw [prat_less_def] "(x::prat) < y + x";
+by (res_inst_tac [("x","y")] exI 1);
+by (simp_tac (simpset() addsimps [prat_add_commute]) 1);
+qed "prat_self_less_add_right";
+
+Goal "(x::prat) < x + y";
+by (rtac (prat_add_commute RS subst) 1);
+by (simp_tac (simpset() addsimps [prat_self_less_add_right]) 1);
+qed "prat_self_less_add_left";
+
+Goalw [prat_less_def] "!!y. $#1p < y ==> (x::prat) < x * y";
+by (auto_tac (claset(),simpset() addsimps [pnat_one_def,
+    prat_add_mult_distrib2]));
+qed "prat_self_less_mult_right";
+
+(*** Properties of <= ***)
+
+Goalw [prat_le_def] "!!w. ~(w < z) ==> z <= (w::prat)";
+by (assume_tac 1);
+qed "prat_leI";
+
+Goalw [prat_le_def] "!!w. z<=w ==> ~(w<(z::prat))";
+by (assume_tac 1);
+qed "prat_leD";
+
+val prat_leE = make_elim prat_leD;
+
+Goal "!!w. (~(w < z)) = (z <= (w::prat))";
+by (fast_tac (claset() addSIs [prat_leI,prat_leD]) 1);
+qed "prat_less_le_iff";
+
+Goalw [prat_le_def] "!!z. ~ z <= w ==> w<(z::prat)";
+by (Fast_tac 1);
+qed "not_prat_leE";
+
+Goalw [prat_le_def] "!!z. z < w ==> z <= (w::prat)";
+by (fast_tac (claset() addEs [prat_less_asym]) 1);
+qed "prat_less_imp_le";
+
+Goalw [prat_le_def] "!!(x::prat). x <= y ==> x < y | x = y";
+by (cut_facts_tac [prat_linear] 1);
+by (fast_tac (claset() addEs [prat_less_irrefl,prat_less_asym]) 1);
+qed "prat_le_imp_less_or_eq";
+
+Goalw [prat_le_def] "!!z. z<w | z=w ==> z <=(w::prat)";
+by (cut_facts_tac [prat_linear] 1);
+by (fast_tac (claset() addEs [prat_less_irrefl,prat_less_asym]) 1);
+qed "prat_less_or_eq_imp_le";
+
+Goal "(x <= (y::prat)) = (x < y | x=y)";
+by (REPEAT(ares_tac [iffI, prat_less_or_eq_imp_le, prat_le_imp_less_or_eq] 1));
+qed "prat_le_eq_less_or_eq";
+
+Goal "w <= (w::prat)";
+by (simp_tac (simpset() addsimps [prat_le_eq_less_or_eq]) 1);
+qed "prat_le_refl";
+
+val prems = goal thy "!!i. [| i <= j; j < k |] ==> i < (k::prat)";
+by (dtac prat_le_imp_less_or_eq 1);
+by (fast_tac (claset() addIs [prat_less_trans]) 1);
+qed "prat_le_less_trans";
+
+Goal "!! (i::prat). [| i < j; j <= k |] ==> i < k";
+by (dtac prat_le_imp_less_or_eq 1);
+by (fast_tac (claset() addIs [prat_less_trans]) 1);
+qed "prat_less_le_trans";
+
+Goal "!!i. [| i <= j; j <= k |] ==> i <= (k::prat)";
+by (EVERY1 [dtac prat_le_imp_less_or_eq, dtac prat_le_imp_less_or_eq,
+            rtac prat_less_or_eq_imp_le, fast_tac (claset() addIs [prat_less_trans])]);
+qed "prat_le_trans";
+
+Goal "!!z. [| z <= w; w <= z |] ==> z = (w::prat)";
+by (EVERY1 [dtac prat_le_imp_less_or_eq, dtac prat_le_imp_less_or_eq,
+            fast_tac (claset() addEs [prat_less_irrefl,prat_less_asym])]);
+qed "prat_le_anti_sym";
+
+Goal "!!x. [| ~ y < x; y ~= x |] ==> x < (y::prat)";
+by (rtac not_prat_leE 1);
+by (fast_tac (claset() addDs [prat_le_imp_less_or_eq]) 1);
+qed "not_less_not_eq_prat_less";
+
+Goalw [prat_less_def] 
+      "!!x. [| x1 < y1; x2 < y2 |] ==> x1 + x2 < y1 + (y2::prat)";
+by (REPEAT(etac exE 1));
+by (res_inst_tac [("x","T+Ta")] exI 1);
+by (auto_tac (claset(),simpset() addsimps prat_add_ac));
+qed "prat_add_less_mono";
+
+Goalw [prat_less_def] 
+      "!!x. [| x1 < y1; x2 < y2 |] ==> x1 * x2 < y1 * (y2::prat)";
+by (REPEAT(etac exE 1));
+by (res_inst_tac [("x","T*Ta+T*x2+x1*Ta")] exI 1);
+by (auto_tac (claset(),simpset() addsimps prat_add_ac @ 
+    [prat_add_mult_distrib,prat_add_mult_distrib2]));
+qed "prat_mult_less_mono";
+
+(* more prat_le *)
+Goal "!!(q1::prat). q1 <= q2  ==> x * q1 <= x * q2";
+by (dtac prat_le_imp_less_or_eq 1);
+by (Step_tac 1);
+by (auto_tac (claset() addSIs [prat_le_refl,
+    prat_less_imp_le,prat_mult_left_less2_mono1],simpset()));
+qed "prat_mult_left_le2_mono1";
+
+Goal "!!(q1::prat). q1 <= q2  ==> q1 * x <= q2 * x";
+by (auto_tac (claset() addDs [prat_mult_left_le2_mono1],
+    simpset() addsimps [prat_mult_commute]));
+qed "prat_mult_le2_mono1";
+
+Goal 
+      "!!(q1::prat). q1 <= q2  ==> qinv (q2) <= qinv (q1)";
+by (dtac prat_le_imp_less_or_eq 1);
+by (Step_tac 1);
+by (auto_tac (claset() addSIs [prat_le_refl,
+    prat_less_imp_le,qinv_prat_less],simpset()));
+qed "qinv_prat_le";
+
+Goal "!!(q1::prat). q1 <= q2  ==> x + q1 <= x + q2";
+by (dtac prat_le_imp_less_or_eq 1);
+by (Step_tac 1);
+by (auto_tac (claset() addSIs [prat_le_refl,
+    prat_less_imp_le,prat_add_less2_mono1],
+    simpset() addsimps [prat_add_commute]));
+qed "prat_add_left_le2_mono1";
+
+Goal "!!(q1::prat). q1 <= q2  ==> q1 + x <= q2 + x";
+by (auto_tac (claset() addDs [prat_add_left_le2_mono1],
+    simpset() addsimps [prat_add_commute]));
+qed "prat_add_le2_mono1";
+
+Goal "!!k l::prat. [|i<=j;  k<=l |] ==> i + k <= j + l";
+by (etac (prat_add_le2_mono1 RS prat_le_trans) 1);
+by (simp_tac (simpset() addsimps [prat_add_commute]) 1);
+(*j moves to the end because it is free while k, l are bound*)
+by (etac prat_add_le2_mono1 1);
+qed "prat_add_le_mono";
+
+Goal "!!(x::prat). x + y < z + y ==> x < z";
+by (rtac ccontr 1);
+by (etac (prat_leI RS prat_le_imp_less_or_eq RS disjE) 1);
+by (dres_inst_tac [("x","y"),("q1.0","z")] prat_add_less2_mono1 1);
+by (auto_tac (claset() addIs [prat_less_asym],
+    simpset() addsimps [prat_less_not_refl]));
+qed "prat_add_right_less_cancel";
+
+Goal "!!(x::prat). y + x < y + z ==> x < z";
+by (res_inst_tac [("y","y")] prat_add_right_less_cancel 1);
+by (asm_full_simp_tac (simpset() addsimps [prat_add_commute]) 1);
+qed "prat_add_left_less_cancel";
+
+(*** lemmas required for lemma_gleason9_34 in PReal : w*y > y/z ***)
+Goalw [prat_pnat_def] "Abs_prat(ratrel^^{(x,y)}) = $#x*qinv($#y)";
+by (auto_tac (claset(),simpset() addsimps [prat_mult,qinv,pnat_mult_1_left,
+    pnat_mult_1]));
+qed "Abs_prat_mult_qinv";
+
+Goal "Abs_prat(ratrel^^{(x,y)}) <= Abs_prat(ratrel^^{(x,Abs_pnat 1)})";
+by (simp_tac (simpset() addsimps [Abs_prat_mult_qinv]) 1);
+by (rtac prat_mult_left_le2_mono1 1);
+by (rtac qinv_prat_le 1);
+by (pnat_ind_tac "y" 1);
+by (dres_inst_tac [("x","$#Abs_pnat 1")] prat_add_le2_mono1 2);
+by (cut_facts_tac [prat_less_1_2 RS prat_less_imp_le] 2);
+by (auto_tac (claset() addIs [prat_le_trans],
+    simpset() addsimps [prat_le_refl,
+    pSuc_is_plus_one,pnat_one_def,prat_pnat_add]));
+qed "lemma_Abs_prat_le1";
+
+Goal "Abs_prat(ratrel^^{(x,Abs_pnat 1)}) <= Abs_prat(ratrel^^{(x*y,Abs_pnat 1)})";
+by (simp_tac (simpset() addsimps [Abs_prat_mult_qinv]) 1);
+by (rtac prat_mult_le2_mono1 1);
+by (pnat_ind_tac "y" 1);
+by (dres_inst_tac [("x","$#x")] prat_add_le2_mono1 2);
+by (cut_inst_tac [("z","$#x")] (prat_self_less_add_self 
+    RS prat_less_imp_le) 2);
+by (auto_tac (claset() addIs [prat_le_trans],
+    simpset() addsimps [prat_le_refl,
+    pSuc_is_plus_one,pnat_one_def,prat_add_mult_distrib2,
+    prat_pnat_add,prat_pnat_mult]));
+qed "lemma_Abs_prat_le2";
+
+Goal "Abs_prat(ratrel^^{(x,z)}) <= Abs_prat(ratrel^^{(x*y,Abs_pnat 1)})";
+by (fast_tac (claset() addIs [prat_le_trans,lemma_Abs_prat_le1,lemma_Abs_prat_le2]) 1);
+qed "lemma_Abs_prat_le3";
+
+Goal "Abs_prat(ratrel^^{(x*y,Abs_pnat 1)}) * Abs_prat(ratrel^^{(w,x)}) = \
+\         Abs_prat(ratrel^^{(w*y,Abs_pnat 1)})";
+by (full_simp_tac (simpset() addsimps [prat_mult,
+    pnat_mult_1,pnat_mult_1_left] @ pnat_mult_ac) 1);
+qed "pre_lemma_gleason9_34";
+
+Goal "Abs_prat(ratrel^^{(y*x,Abs_pnat 1*y)}) = \
+\         Abs_prat(ratrel^^{(x,Abs_pnat 1)})";
+by (auto_tac (claset(),simpset() addsimps 
+    [pnat_mult_1,pnat_mult_1_left] @ pnat_mult_ac));
+qed "pre_lemma_gleason9_34b";
+
+Goal "($#n < $#m) = (n < m)";
+by (auto_tac (claset(),simpset() addsimps [prat_less_def,
+    pnat_less_iff,prat_pnat_add]));
+by (res_inst_tac [("z","T")] eq_Abs_prat 1);
+by (auto_tac (claset() addDs [pnat_eq_lessI],
+    simpset() addsimps [prat_add,pnat_mult_1,
+    pnat_mult_1_left,prat_pnat_def,pnat_less_iff RS sym]));
+qed "prat_pnat_less_iff";
+
+Addsimps [prat_pnat_less_iff];
+
+(***)(***)(***)(***)(***)(***)(***)(***)(***)(***)(***)(***)(***)(***)
+
+(*** prove witness that will be required to prove non-emptiness ***)
+(*** of preal type as defined using Dedekind Sections in PReal ***)
+(*** Show that exists positive real `one' ***)
+
+Goal "? q. q: {x::prat. x < $#Abs_pnat 1}";
+by (fast_tac (claset() addIs [prat_less_qinv_2_1]) 1);
+qed "lemma_prat_less_1_memEx";
+
+Goal "{x::prat. x < $#Abs_pnat 1} ~= {}";
+by (rtac notI 1);
+by (cut_facts_tac [lemma_prat_less_1_memEx] 1);
+by (Asm_full_simp_tac 1);
+qed "lemma_prat_less_1_set_non_empty";
+
+Goalw [psubset_def] "{} < {x::prat. x < $#Abs_pnat 1}";
+by (asm_full_simp_tac (simpset() addsimps 
+         [lemma_prat_less_1_set_non_empty RS not_sym]) 1);
+qed "empty_set_psubset_lemma_prat_less_1_set";
+
+(*** exists rational not in set --- $#Abs_pnat 1 itself ***)
+Goal "? q. q ~: {x::prat. x < $#Abs_pnat 1}";
+by (res_inst_tac [("x","$#Abs_pnat 1")] exI 1);
+by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl]));
+qed "lemma_prat_less_1_not_memEx";
+
+Goal "{x::prat. x < $#Abs_pnat 1} ~= {q::prat. True}";
+by (rtac notI 1);
+by (cut_facts_tac [lemma_prat_less_1_not_memEx] 1);
+by (Asm_full_simp_tac 1);
+qed "lemma_prat_less_1_set_not_rat_set";
+
+Goalw [psubset_def,subset_def] 
+      "{x::prat. x < $#Abs_pnat 1} < {q::prat. True}";
+by (asm_full_simp_tac (simpset() addsimps 
+      [lemma_prat_less_1_set_not_rat_set,
+       lemma_prat_less_1_not_memEx]) 1);
+qed "lemma_prat_less_1_set_psubset_rat_set";
+
+(*** prove non_emptiness of type ***)
+Goal "{x::prat. x < $#Abs_pnat 1} : {A. {} < A & A < {q::prat. True} & \
+\                                        (!y: A. ((!z. z < y --> z: A) & \
+\                                        (? u: A. y < u)))}";
+by (auto_tac (claset() addDs [prat_less_trans],
+    simpset() addsimps [empty_set_psubset_lemma_prat_less_1_set,
+                       lemma_prat_less_1_set_psubset_rat_set]));
+by (dtac prat_dense 1);
+by (Fast_tac 1);
+qed "preal_1";
+
+
+
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/PRat.thy	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,49 @@
+(*  Title       : PRat.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : The positive rationals
+*) 
+
+PRat = PNat + Equiv +
+
+constdefs
+    ratrel   ::  "((pnat * pnat) * (pnat * pnat)) set"
+    "ratrel  ==  {p. ? x1 y1 x2 y2. p=((x1::pnat,y1),(x2,y2)) & x1*y2 = x2*y1}" 
+
+typedef prat = "{x::(pnat*pnat).True}/ratrel"          (Equiv.quotient_def)
+
+instance
+   prat  :: {ord,plus,times}
+
+
+constdefs
+
+  prat_pnat :: pnat => prat              ("$#_" [80] 80)
+  "$# m     == Abs_prat(ratrel^^{(m,Abs_pnat 1)})"
+
+  qinv      :: prat => prat
+  "qinv(Q)  == Abs_prat(UN p:Rep_prat(Q). split (%x y. ratrel^^{(y,x)}) p)" 
+
+defs
+
+  prat_add_def  
+  "P + Q == Abs_prat(UN p1:Rep_prat(P). UN p2:Rep_prat(Q).
+                split(%x1 y1. split(%x2 y2. ratrel^^{(x1*y2 + x2*y1, y1*y2)}) p2) p1)"
+
+  prat_mult_def  
+  "P * Q == Abs_prat(UN p1:Rep_prat(P). UN p2:Rep_prat(Q).
+                split(%x1 y1. split(%x2 y2. ratrel^^{(x1*x2, y1*y2)}) p2) p1)"
+ 
+  (*** Gleason p. 119 ***)
+  prat_less_def
+  "P < (Q::prat) == ? T. P + T = Q"
+
+  prat_le_def
+  "P <= (Q::prat) == ~(Q < P)" 
+
+end
+  
+
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/PReal.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,1332 @@
+(*  Title       : PReal.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : The positive reals as Dedekind sections of positive
+                  rationals. Fundamentals of Abstract Analysis 
+                  [Gleason- p. 121] provides some of the definitions.
+*)
+
+open PReal;
+
+Goal "inj_on Abs_preal preal";
+by (rtac inj_on_inverseI 1);
+by (etac Abs_preal_inverse 1);
+qed "inj_on_Abs_preal";
+
+Addsimps [inj_on_Abs_preal RS inj_on_iff];
+
+Goal "inj(Rep_preal)";
+by (rtac inj_inverseI 1);
+by (rtac Rep_preal_inverse 1);
+qed "inj_Rep_preal";
+
+Goalw [preal_def] "{} ~: preal";
+by (Fast_tac 1);
+qed "empty_not_mem_preal";
+
+(* {} : preal ==> P *)
+bind_thm ("empty_not_mem_prealE", empty_not_mem_preal RS notE);
+
+Addsimps [empty_not_mem_preal];
+
+Goalw [preal_def] "{x::prat. x < $#Abs_pnat 1} : preal";
+by (rtac preal_1 1);
+qed "one_set_mem_preal";
+
+Addsimps [one_set_mem_preal];
+
+Goalw [preal_def] "!!x. x : preal ==> {} < x";
+by (Fast_tac 1);
+qed "preal_psubset_empty";
+
+Goal "{} < Rep_preal x";
+by (rtac (Rep_preal RS preal_psubset_empty) 1);
+qed "Rep_preal_psubset_empty";
+
+Goal "? x. x: Rep_preal X";
+by (cut_inst_tac [("x","X")]  Rep_preal_psubset_empty 1);
+by (auto_tac (claset() addIs [(equals0I RS sym)],
+              simpset() addsimps [psubset_def]));
+qed "mem_Rep_preal_Ex";
+
+Goalw [preal_def] 
+      "!!A. [| {} < A; A < {q::prat. True}; \
+\              (!y: A. ((!z. z < y --> z: A) & \
+\                        (? u: A. y < u))) |] ==> A : preal";
+by (Fast_tac 1);
+qed "prealI1";
+    
+Goalw [preal_def] 
+      "!!A. [| {} < A; A < {q::prat. True}; \
+\              !y: A. (!z. z < y --> z: A); \
+\              !y: A. (? u: A. y < u) |] ==> A : preal";
+by (Best_tac 1);
+qed "prealI2";
+
+Goalw [preal_def] 
+      "!!A. A : preal ==> {} < A & A < {q::prat. True} & \
+\                         (!y: A. ((!z. z < y --> z: A) & \
+\                                  (? u: A. y < u)))";
+by (Fast_tac 1);
+qed "prealE_lemma";
+
+
+AddSIs [prealI1,prealI2];
+
+Addsimps [Abs_preal_inverse];
+
+
+Goalw [preal_def] 
+      "!!A. A : preal ==> {} < A";
+by (Fast_tac 1);
+qed "prealE_lemma1";
+
+Goalw [preal_def] 
+      "!!A. A : preal ==> A < {q::prat. True}";
+by (Fast_tac 1);
+qed "prealE_lemma2";
+
+Goalw [preal_def] 
+      "!!A. A : preal ==> !y: A. (!z. z < y --> z: A)";
+by (Fast_tac 1);
+qed "prealE_lemma3";
+
+Goal 
+      "!!A. [| A : preal; y: A |] ==> (!z. z < y --> z: A)";
+by (fast_tac (claset() addSDs [prealE_lemma3]) 1);
+qed "prealE_lemma3a";
+
+Goal 
+      "!!A. [| A : preal; y: A; z < y |] ==> z: A";
+by (fast_tac (claset() addSDs [prealE_lemma3a]) 1);
+qed "prealE_lemma3b";
+
+Goalw [preal_def] 
+      "!!A. A : preal ==> !y: A. (? u: A. y < u)";
+by (Fast_tac 1);
+qed "prealE_lemma4";
+
+Goal 
+      "!!A. [| A : preal; y: A |] ==> ? u: A. y < u";
+by (fast_tac (claset() addSDs [prealE_lemma4]) 1);
+qed "prealE_lemma4a";
+
+Goal "? x. x~: Rep_preal X";
+by (cut_inst_tac [("x","X")] Rep_preal 1);
+by (dtac prealE_lemma2 1);
+by (rtac ccontr 1);
+by (auto_tac (claset(),simpset() addsimps [psubset_def]));
+by (blast_tac (claset() addIs [set_ext] addEs [swap]) 1);
+qed "not_mem_Rep_preal_Ex";
+
+(** prat_pnat: the injection from prat to preal **)
+(** A few lemmas **)
+Goal "{} < {xa::prat. xa < y}";
+by (cut_facts_tac [qless_Ex] 1);
+by (auto_tac (claset() addEs [equalityCE],
+              simpset() addsimps [psubset_def]));
+qed "lemma_prat_less_set_Ex";
+
+Goal "{xa::prat. xa < y} : preal";
+by (cut_facts_tac [qless_Ex] 1);
+by Safe_tac;
+by (rtac lemma_prat_less_set_Ex 1);
+by (auto_tac (claset() addIs [prat_less_trans],
+    simpset() addsimps [psubset_def]));
+by (eres_inst_tac [("c","y")] equalityCE 1);
+by (auto_tac (claset() addDs [prat_less_irrefl],simpset()));
+by (dres_inst_tac [("q1.0","ya")] prat_dense 1);
+by (Fast_tac 1);
+qed "lemma_prat_less_set_mem_preal";
+
+Goal "!!(x::prat). {xa. xa < x} = {x. x < y} ==> x = y";
+by (cut_inst_tac [("q1.0","x"),("q2.0","y")] prat_linear 1);
+by Auto_tac;
+by (dtac prat_dense 1 THEN etac exE 1);
+by (eres_inst_tac [("c","xa")] equalityCE 1);
+by (auto_tac (claset() addDs [prat_less_asym],simpset()));
+by (dtac prat_dense 1 THEN etac exE 1);
+by (eres_inst_tac [("c","xa")] equalityCE 1);
+by (auto_tac (claset() addDs [prat_less_asym],simpset()));
+qed "lemma_prat_set_eq";
+
+Goal "inj(preal_prat)";
+by (rtac injI 1);
+by (rewtac preal_prat_def);
+by (dtac (inj_on_Abs_preal RS inj_onD) 1);
+by (rtac lemma_prat_less_set_mem_preal 1);
+by (rtac lemma_prat_less_set_mem_preal 1);
+by (etac lemma_prat_set_eq 1);
+qed "inj_preal_prat";
+
+      (*** theorems for ordering ***)
+(* prove introduction and elimination rules for preal_less *)
+
+Goalw [preal_less_def]
+      "R1 < (R2::preal) = (Rep_preal(R1) < Rep_preal(R2))";
+by (Fast_tac 1);
+qed "preal_less_iff";
+
+Goalw [preal_less_def]
+      "!! (R1::preal). R1 < R2 ==> (Rep_preal(R1) < Rep_preal(R2))";
+by (Fast_tac  1);
+qed "preal_lessI";
+
+Goalw [preal_less_def]
+      "R1 < (R2::preal) --> (Rep_preal(R1) < Rep_preal(R2))";
+by (Fast_tac  1);
+qed "preal_lessE_lemma";
+
+Goal 
+     "!! R1. [| R1 < (R2::preal); \
+\               (Rep_preal(R1) < Rep_preal(R2)) ==> P |] \
+\     ==> P";
+by (dtac (preal_lessE_lemma RS mp) 1);
+by Auto_tac;
+qed "preal_lessE";
+
+(* A positive fraction not in a positive real is an upper bound *)
+(* Gleason p. 122 - Remark (1)                                  *)
+
+Goal "!!x. x ~: Rep_preal(R) ==> !y: Rep_preal(R). y < x";
+by (cut_inst_tac [("x1","R")] (Rep_preal RS prealE_lemma) 1);
+by (auto_tac (claset() addIs [not_less_not_eq_prat_less],simpset()));
+qed "not_in_preal_ub";
+
+(* preal_less is a strong order i.e nonreflexive and transitive *)
+
+Goalw [preal_less_def] "~ (x::preal) < x";
+by (simp_tac (simpset() addsimps [psubset_def]) 1);
+qed "preal_less_not_refl";
+
+(*** y < y ==> P ***)
+bind_thm("preal_less_irrefl",preal_less_not_refl RS notE);
+
+Goal "!!(x::preal). x < y ==> x ~= y";
+by (auto_tac (claset(),simpset() addsimps [preal_less_not_refl]));
+qed "preal_not_refl2";
+
+Goalw  [preal_less_def] "!!(x::preal). [| x < y; y < z |] ==> x < z";
+by (auto_tac (claset() addDs [subsetD,equalityI],
+              simpset() addsimps [psubset_def]));
+qed "preal_less_trans";
+
+Goal "!! (q1::preal). [| q1 < q2; q2 < q1 |] ==> P";
+by (dtac preal_less_trans 1 THEN assume_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [preal_less_not_refl]) 1);
+qed "preal_less_asym";
+
+Goalw [preal_less_def] 
+      "(r1::preal) < r2 | r1 = r2 | r2 < r1";
+by (auto_tac (claset() addSDs [inj_Rep_preal RS injD],
+              simpset() addsimps [psubset_def]));
+by (rtac prealE_lemma3b 1 THEN rtac Rep_preal 1);
+by (assume_tac 1);
+by (fast_tac (claset() addDs [not_in_preal_ub]) 1);
+qed "preal_linear";
+
+Goal
+    "!!(r1::preal). [| r1 < r2 ==> P;  r1 = r2 ==> P; \
+\          r2 < r1 ==> P |] ==> P";
+by (cut_inst_tac [("r1.0","r1"),("r2.0","r2")] preal_linear 1);
+by Auto_tac;
+qed "preal_linear_less2";
+
+  (*** Properties of addition ***)
+
+Goalw [preal_add_def] "(x::preal) + y = y + x";
+by (res_inst_tac [("f","Abs_preal")] arg_cong 1);
+by (rtac set_ext 1);
+by (blast_tac (claset() addIs [prat_add_commute RS subst]) 1);
+qed "preal_add_commute";
+
+(** addition of two positive reals gives a positive real **)
+(** lemmas for proving positive reals addition set in preal **)
+
+(** Part 1 of Dedekind sections def **)
+Goal "{} < {w. ? x: Rep_preal R. ? y:Rep_preal S. w = x + y}";
+by (cut_facts_tac [mem_Rep_preal_Ex,mem_Rep_preal_Ex] 1);
+by (auto_tac (claset() addSIs [psubsetI] addEs [equalityCE],simpset()));
+qed "preal_add_set_not_empty";
+
+(** Part 2 of Dedekind sections def **)
+Goal "? q. q  ~: {w. ? x: Rep_preal R. ? y:Rep_preal S. w = x + y}";
+by (cut_inst_tac [("X","R")] not_mem_Rep_preal_Ex 1);
+by (cut_inst_tac [("X","S")] not_mem_Rep_preal_Ex 1);
+by (REPEAT(etac exE 1));
+by (REPEAT(dtac not_in_preal_ub 1));
+by (res_inst_tac [("x","x+xa")] exI 1);
+by (Auto_tac THEN (REPEAT(etac ballE 1)) THEN Auto_tac);
+by (dtac prat_add_less_mono 1);
+by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl]));
+qed "preal_not_mem_add_set_Ex";
+
+Goal "{w. ? x: Rep_preal R. ? y:Rep_preal S. w = x + y} < {q. True}";
+by (auto_tac (claset() addSIs [psubsetI],simpset()));
+by (cut_inst_tac [("R","R"),("S","S")] preal_not_mem_add_set_Ex 1);
+by (etac exE 1);
+by (eres_inst_tac [("c","q")] equalityCE 1);
+by Auto_tac;
+qed "preal_add_set_not_prat_set";
+
+(** Part 3 of Dedekind sections def **)
+Goal "!y: {w. ? x: Rep_preal R. ? y: Rep_preal S. w = x + y}. \
+\         !z. z < y --> z : {w. ? x:Rep_preal R. ? y:Rep_preal S. w = x + y}";
+by Auto_tac;
+by (forward_tac [prat_mult_qinv_less_1] 1);
+by (forw_inst_tac [("x","x"),("q2.0","$#Abs_pnat 1")] 
+    prat_mult_less2_mono1 1);
+by (forw_inst_tac [("x","ya"),("q2.0","$#Abs_pnat 1")] 
+    prat_mult_less2_mono1 1);
+by (Asm_full_simp_tac 1);
+by (REPEAT(dtac (Rep_preal RS prealE_lemma3a) 1));
+by (REPEAT(etac allE 1));
+by Auto_tac;
+by (REPEAT(rtac bexI 1));
+by (auto_tac (claset(),simpset() addsimps [prat_add_mult_distrib2 
+     RS sym,prat_add_assoc RS sym,prat_mult_assoc]));
+qed "preal_add_set_lemma3";
+
+Goal "!y: {w. ? x: Rep_preal R. ? y: Rep_preal S. w = x + y}. \
+\         ? u: {w. ? x: Rep_preal R. ? y: Rep_preal S. w = x + y}. y < u";
+by Auto_tac;
+by (dtac (Rep_preal RS prealE_lemma4a) 1);
+by (auto_tac (claset() addIs [prat_add_less2_mono1],simpset()));
+qed "preal_add_set_lemma4";
+
+Goal "{w. ? x: Rep_preal R. ? y: Rep_preal S. w = x + y} : preal";
+by (rtac prealI2 1);
+by (rtac preal_add_set_not_empty 1);
+by (rtac preal_add_set_not_prat_set 1);
+by (rtac preal_add_set_lemma3 1);
+by (rtac preal_add_set_lemma4 1);
+qed "preal_mem_add_set";
+
+Goalw [preal_add_def] "((x::preal) + y) + z = x + (y + z)";
+by (res_inst_tac [("f","Abs_preal")] arg_cong 1);
+by (rtac set_ext 1);
+by (rtac (preal_mem_add_set RS Abs_preal_inverse RS ssubst) 1);
+by (rtac (preal_mem_add_set RS Abs_preal_inverse RS ssubst) 1);
+by (auto_tac (claset(),simpset() addsimps prat_add_ac));
+by (rtac bexI 1);
+by (auto_tac (claset() addSIs [exI],simpset() addsimps prat_add_ac));
+qed "preal_add_assoc";
+
+qed_goal "preal_add_left_commute" thy
+    "(z1::preal) + (z2 + z3) = z2 + (z1 + z3)"
+ (fn _ => [rtac (preal_add_commute RS trans) 1, rtac (preal_add_assoc RS trans) 1,
+           rtac (preal_add_commute RS arg_cong) 1]);
+
+(* Positive Reals addition is an AC operator *)
+val preal_add_ac = [preal_add_assoc, preal_add_commute, preal_add_left_commute];
+
+  (*** Properties of multiplication ***)
+
+(** Proofs essentially same as for addition **)
+
+Goalw [preal_mult_def] "(x::preal) * y = y * x";
+by (res_inst_tac [("f","Abs_preal")] arg_cong 1);
+by (rtac set_ext 1);
+by (blast_tac (claset() addIs [prat_mult_commute RS subst]) 1);
+qed "preal_mult_commute";
+
+(** multiplication of two positive reals gives a positive real **)
+(** lemmas for proving positive reals multiplication set in preal **)
+
+(** Part 1 of Dedekind sections def **)
+Goal "{} < {w. ? x: Rep_preal R. ? y:Rep_preal S. w = x * y}";
+by (cut_facts_tac [mem_Rep_preal_Ex,mem_Rep_preal_Ex] 1);
+by (auto_tac (claset() addSIs [psubsetI] addEs [equalityCE],simpset()));
+qed "preal_mult_set_not_empty";
+
+(** Part 2 of Dedekind sections def **)
+Goal "? q. q  ~: {w. ? x: Rep_preal R. ? y:Rep_preal S. w = x * y}";
+by (cut_inst_tac [("X","R")] not_mem_Rep_preal_Ex 1);
+by (cut_inst_tac [("X","S")] not_mem_Rep_preal_Ex 1);
+by (REPEAT(etac exE 1));
+by (REPEAT(dtac not_in_preal_ub 1));
+by (res_inst_tac [("x","x*xa")] exI 1);
+by (Auto_tac  THEN (REPEAT(etac ballE 1)) THEN Auto_tac );
+by (dtac prat_mult_less_mono 1);
+by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl]));
+qed "preal_not_mem_mult_set_Ex";
+
+Goal "{w. ? x: Rep_preal R. ? y:Rep_preal S. w = x * y} < {q. True}";
+by (auto_tac (claset() addSIs [psubsetI],simpset()));
+by (cut_inst_tac [("R","R"),("S","S")] preal_not_mem_mult_set_Ex 1);
+by (etac exE 1);
+by (eres_inst_tac [("c","q")] equalityCE 1);
+by Auto_tac;
+qed "preal_mult_set_not_prat_set";
+
+(** Part 3 of Dedekind sections def **)
+Goal "!y: {w. ? x: Rep_preal R. ? y: Rep_preal S. w = x * y}. \
+\         !z. z < y --> z : {w. ? x:Rep_preal R. ? y:Rep_preal S. w = x * y}";
+by Auto_tac;
+by (forw_inst_tac [("x","qinv(ya)"),("q1.0","z")] 
+    prat_mult_left_less2_mono1 1);
+by (asm_full_simp_tac (simpset() addsimps prat_mult_ac) 1);
+by (dtac (Rep_preal RS prealE_lemma3a) 1);
+by (etac allE 1);
+by (REPEAT(rtac bexI 1));
+by (auto_tac (claset(),simpset() addsimps [prat_mult_assoc]));
+qed "preal_mult_set_lemma3";
+
+Goal "!y: {w. ? x: Rep_preal R. ? y: Rep_preal S. w = x * y}. \
+\         ? u: {w. ? x: Rep_preal R. ? y: Rep_preal S. w = x * y}. y < u";
+by Auto_tac;
+by (dtac (Rep_preal RS prealE_lemma4a) 1);
+by (auto_tac (claset() addIs [prat_mult_less2_mono1],simpset()));
+qed "preal_mult_set_lemma4";
+
+Goal "{w. ? x: Rep_preal R. ? y: Rep_preal S. w = x * y} : preal";
+by (rtac prealI2 1);
+by (rtac preal_mult_set_not_empty 1);
+by (rtac preal_mult_set_not_prat_set 1);
+by (rtac preal_mult_set_lemma3 1);
+by (rtac preal_mult_set_lemma4 1);
+qed "preal_mem_mult_set";
+
+Goalw [preal_mult_def] "((x::preal) * y) * z = x * (y * z)";
+by (res_inst_tac [("f","Abs_preal")] arg_cong 1);
+by (rtac set_ext 1);
+by (rtac (preal_mem_mult_set RS Abs_preal_inverse RS ssubst) 1);
+by (rtac (preal_mem_mult_set RS Abs_preal_inverse RS ssubst) 1);
+by (auto_tac (claset(),simpset() addsimps prat_mult_ac));
+by (rtac bexI 1);
+by (auto_tac (claset() addSIs [exI],simpset() addsimps prat_mult_ac));
+qed "preal_mult_assoc";
+
+qed_goal "preal_mult_left_commute" thy
+    "(z1::preal) * (z2 * z3) = z2 * (z1 * z3)"
+ (fn _ => [rtac (preal_mult_commute RS trans) 1, 
+           rtac (preal_mult_assoc RS trans) 1,
+           rtac (preal_mult_commute RS arg_cong) 1]);
+
+(* Positive Reals multiplication is an AC operator *)
+val preal_mult_ac = [preal_mult_assoc, 
+                     preal_mult_commute, 
+                     preal_mult_left_commute];
+
+(* Positive Real 1 is the multiplicative identity element *) 
+(* long *)
+Goalw [preal_prat_def,preal_mult_def] "(@#($#Abs_pnat 1)) * z = z";
+by (rtac (Rep_preal_inverse RS subst) 1);
+by (res_inst_tac [("f","Abs_preal")] arg_cong 1);
+by (rtac (one_set_mem_preal RS Abs_preal_inverse RS ssubst) 1);
+by (rtac set_ext 1);
+by (auto_tac(claset(),simpset() addsimps [Rep_preal_inverse]));
+by (EVERY1[dtac (Rep_preal RS prealE_lemma4a),etac bexE]);
+by (dtac prat_mult_less_mono 1);
+by (auto_tac (claset() addDs [Rep_preal RS prealE_lemma3a],simpset()));
+by (EVERY1[forward_tac [Rep_preal RS prealE_lemma4a],etac bexE]);
+by (forw_inst_tac [("x","qinv(u)"),("q1.0","x")] 
+    prat_mult_less2_mono1 1);
+by (rtac exI 1 THEN Auto_tac THEN res_inst_tac [("x","u")] bexI 1);
+by (auto_tac (claset(),simpset() addsimps [prat_mult_assoc]));
+qed "preal_mult_1";
+
+Goal "z * (@#($#Abs_pnat 1)) = z";
+by (rtac (preal_mult_commute RS subst) 1);
+by (rtac preal_mult_1 1);
+qed "preal_mult_1_right";
+
+(** Lemmas **)
+
+qed_goal "preal_add_assoc_cong" thy
+    "!!z. (z::preal) + v = z' + v' ==> z + (v + w) = z' + (v' + w)"
+ (fn _ => [(asm_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1)]);
+
+qed_goal "preal_add_assoc_swap" thy "(z::preal) + (v + w) = v + (z + w)"
+ (fn _ => [(REPEAT (ares_tac [preal_add_commute RS preal_add_assoc_cong] 1))]);
+
+(** Distribution of multiplication across addition **)
+(** lemmas for the proof **)
+
+ (** lemmas **)
+Goalw [preal_add_def] 
+      "!!R. z: Rep_preal(R+S) ==> \
+\           ? x: Rep_preal(R). ? y: Rep_preal(S). z = x + y";
+by (dtac (preal_mem_add_set RS Abs_preal_inverse RS subst) 1);
+by (Fast_tac 1);
+qed "mem_Rep_preal_addD";
+
+Goalw [preal_add_def] 
+      "!!R. ? x: Rep_preal(R). ? y: Rep_preal(S). z = x + y \
+\      ==> z: Rep_preal(R+S)";
+by (rtac (preal_mem_add_set RS Abs_preal_inverse RS ssubst) 1);
+by (Fast_tac 1);
+qed "mem_Rep_preal_addI";
+
+Goal " z: Rep_preal(R+S) = (? x: Rep_preal(R). \
+\                                 ? y: Rep_preal(S). z = x + y)";
+by (fast_tac (claset() addSIs [mem_Rep_preal_addD,mem_Rep_preal_addI]) 1);
+qed "mem_Rep_preal_add_iff";
+
+Goalw [preal_mult_def] 
+      "!!R. z: Rep_preal(R*S) ==> \
+\           ? x: Rep_preal(R). ? y: Rep_preal(S). z = x * y";
+by (dtac (preal_mem_mult_set RS Abs_preal_inverse RS subst) 1);
+by (Fast_tac 1);
+qed "mem_Rep_preal_multD";
+
+Goalw [preal_mult_def] 
+      "!!R. ? x: Rep_preal(R). ? y: Rep_preal(S). z = x * y \
+\      ==> z: Rep_preal(R*S)";
+by (rtac (preal_mem_mult_set RS Abs_preal_inverse RS ssubst) 1);
+by (Fast_tac 1);
+qed "mem_Rep_preal_multI";
+
+Goal " z: Rep_preal(R*S) = (? x: Rep_preal(R). \
+\                                 ? y: Rep_preal(S). z = x * y)";
+by (fast_tac (claset() addSIs [mem_Rep_preal_multD,mem_Rep_preal_multI]) 1);
+qed "mem_Rep_preal_mult_iff";
+
+(** More lemmas for preal_add_mult_distrib2 **)
+goal PRat.thy "!!(a1::prat). a1 < a2 ==> a1 * b + a2 * c < a2 * (b + c)";
+by (auto_tac (claset() addSIs [prat_add_less2_mono1,prat_mult_less2_mono1],
+              simpset() addsimps [prat_add_mult_distrib2]));
+qed "lemma_prat_add_mult_mono";
+
+Goal "!!xb. [| xb: Rep_preal z1; xc: Rep_preal z2; ya: \
+\                  Rep_preal w; yb: Rep_preal w |] ==> \
+\                  xb * ya + xc * yb: Rep_preal (z1 * w + z2 * w)";
+by (fast_tac (claset() addIs [mem_Rep_preal_addI,mem_Rep_preal_multI]) 1);
+qed "lemma_add_mult_mem_Rep_preal";
+
+Goal "!!xb. [| xb: Rep_preal z1; xc: Rep_preal z2; ya: \
+\                  Rep_preal w; yb: Rep_preal w |] ==> \
+\                  yb*(xb + xc): Rep_preal (w*(z1 + z2))";
+by (fast_tac (claset() addIs [mem_Rep_preal_addI,mem_Rep_preal_multI]) 1);
+qed "lemma_add_mult_mem_Rep_preal1";
+
+Goal "!!x. x: Rep_preal (w * z1 + w * z2) ==> \
+\              x: Rep_preal (w * (z1 + z2))";
+by (auto_tac (claset() addSDs [mem_Rep_preal_addD,mem_Rep_preal_multD],
+              simpset()));
+by (forw_inst_tac [("ya","xb"),("yb","xc"),("xb","ya"),("xc","yb")] 
+                                   lemma_add_mult_mem_Rep_preal1 1);
+by Auto_tac;
+by (res_inst_tac [("q1.0","xb"),("q2.0","xc")] prat_linear_less2 1);
+by (dres_inst_tac [("b","ya"),("c","yb")] lemma_prat_add_mult_mono 1);
+by (rtac (Rep_preal RS prealE_lemma3b) 1);
+by (auto_tac (claset(),simpset() addsimps [prat_add_mult_distrib2]));
+by (dres_inst_tac [("ya","xc"),("yb","xb"),("xc","ya"),("xb","yb")] 
+                                   lemma_add_mult_mem_Rep_preal1 1);
+by Auto_tac;
+by (dres_inst_tac [("b","yb"),("c","ya")] lemma_prat_add_mult_mono 1);
+by (rtac (Rep_preal RS prealE_lemma3b) 1);
+by (thin_tac "xc * ya + xc * yb  : Rep_preal (w * (z1 + z2))" 1);
+by (auto_tac (claset(),simpset() addsimps [prat_add_mult_distrib,
+              prat_add_commute] @ preal_add_ac ));
+qed "lemma_preal_add_mult_distrib";
+
+Goal "!!x. x: Rep_preal (w * (z1 + z2)) ==> \
+\              x: Rep_preal (w * z1 + w * z2)";
+by (auto_tac (claset() addSDs [mem_Rep_preal_addD,mem_Rep_preal_multD]
+              addSIs [bexI,mem_Rep_preal_addI,mem_Rep_preal_multI],
+              simpset() addsimps [prat_add_mult_distrib2]));
+qed "lemma_preal_add_mult_distrib2";
+
+Goal "(w * ((z1::preal) + z2)) = (w * z1) + (w * z2)";
+by (rtac (inj_Rep_preal RS injD) 1);
+by (rtac set_ext 1);
+by (fast_tac (claset() addIs [lemma_preal_add_mult_distrib,
+                       lemma_preal_add_mult_distrib2]) 1);
+qed "preal_add_mult_distrib2";
+
+Goal "(((z1::preal) + z2) * w) = (z1 * w) + (z2 * w)";
+by (simp_tac (simpset() addsimps [preal_mult_commute,
+                       preal_add_mult_distrib2]) 1);
+qed "preal_add_mult_distrib";
+
+(*** Prove existence of inverse ***)
+(*** Inverse is a positive real ***)
+
+Goal "? y. qinv(y) ~:  Rep_preal X";
+by (cut_inst_tac [("X","X")] not_mem_Rep_preal_Ex 1);
+by (etac exE 1 THEN cut_inst_tac [("x","x")] prat_as_inverse_ex 1);
+by Auto_tac;
+qed "qinv_not_mem_Rep_preal_Ex";
+
+Goal "? q. q: {x. ? y. x < y & qinv y ~:  Rep_preal A}";
+by (cut_inst_tac [("X","A")] qinv_not_mem_Rep_preal_Ex 1);
+by Auto_tac;
+by (cut_inst_tac [("y","y")] qless_Ex 1);
+by (Fast_tac 1);
+qed "lemma_preal_mem_inv_set_ex";
+
+(** Part 1 of Dedekind sections def **)
+Goal "{} < {x. ? y. x < y & qinv y ~:  Rep_preal A}";
+by (cut_facts_tac [lemma_preal_mem_inv_set_ex] 1);
+by (auto_tac (claset() addSIs [psubsetI] addEs [equalityCE],simpset()));
+qed "preal_inv_set_not_empty";
+
+(** Part 2 of Dedekind sections def **)
+Goal "? y. qinv(y) :  Rep_preal X";
+by (cut_inst_tac [("X","X")] mem_Rep_preal_Ex 1);
+by (etac exE 1 THEN cut_inst_tac [("x","x")] prat_as_inverse_ex 1);
+by Auto_tac;
+qed "qinv_mem_Rep_preal_Ex";
+
+Goal "? x. x ~: {x. ? y. x < y & qinv y ~:  Rep_preal A}";
+by (rtac ccontr 1);
+by (cut_inst_tac [("X","A")] qinv_mem_Rep_preal_Ex 1);
+by Auto_tac;
+by (EVERY1[etac allE, etac exE, etac conjE]);
+by (dtac qinv_prat_less 1 THEN dtac not_in_preal_ub 1);
+by (eres_inst_tac [("x","qinv y")] ballE 1);
+by (dtac prat_less_trans 1);
+by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl]));
+qed "preal_not_mem_inv_set_Ex";
+
+Goal "{x. ? y. x < y & qinv y ~:  Rep_preal A} < {q. True}";
+by (auto_tac (claset() addSIs [psubsetI],simpset()));
+by (cut_inst_tac [("A","A")]  preal_not_mem_inv_set_Ex 1);
+by (etac exE 1);
+by (eres_inst_tac [("c","x")] equalityCE 1);
+by Auto_tac;
+qed "preal_inv_set_not_prat_set";
+
+(** Part 3 of Dedekind sections def **)
+Goal "! y: {x. ? y. x < y & qinv y ~: Rep_preal A}. \
+ \      !z. z < y --> z : {x. ? y. x < y & qinv y ~: Rep_preal A}";
+by Auto_tac;
+by (res_inst_tac [("x","ya")] exI 1);
+by (auto_tac (claset() addIs [prat_less_trans],simpset()));
+qed "preal_inv_set_lemma3";
+
+Goal "! y: {x. ? y. x < y & qinv y ~: Rep_preal A}. \
+\       Bex {x. ? y. x < y & qinv y ~: Rep_preal A} (op < y)";
+by (blast_tac (claset() addDs [prat_dense]) 1);
+qed "preal_inv_set_lemma4";
+
+Goal "{x. ? y. x < y & qinv(y) ~: Rep_preal(A)} : preal";
+by (rtac prealI2 1);
+by (rtac preal_inv_set_not_empty 1);
+by (rtac preal_inv_set_not_prat_set 1);
+by (rtac preal_inv_set_lemma3 1);
+by (rtac preal_inv_set_lemma4 1);
+qed "preal_mem_inv_set";
+
+(*more lemmas for inverse *)
+Goal "!!x. x: Rep_preal(pinv(A)*A) ==> x: Rep_preal(@#($#Abs_pnat 1))";
+by (auto_tac (claset() addSDs [mem_Rep_preal_multD],
+              simpset() addsimps [pinv_def,preal_prat_def] ));
+by (dtac (preal_mem_inv_set RS Abs_preal_inverse RS subst) 1);
+by (auto_tac (claset() addSDs [not_in_preal_ub],simpset()));
+by (dtac prat_mult_less_mono 1 THEN Blast_tac 1);
+by (auto_tac (claset(),simpset()));
+qed "preal_mem_mult_invD";
+
+(*** Gleason's Lemma 9-3.4 p 122 ***)
+Goal "!!p. ! xa : Rep_preal(A). xa + x : Rep_preal(A) ==> \
+\            ? xb : Rep_preal(A). xb + ($#p)*x : Rep_preal(A)";
+by (cut_facts_tac [mem_Rep_preal_Ex] 1);
+by (res_inst_tac [("n","p")] pnat_induct 1);
+by (auto_tac (claset(),simpset() addsimps [pnat_one_def,
+    pSuc_is_plus_one,prat_add_mult_distrib,prat_pnat_add,prat_add_assoc RS sym]));
+qed "lemma1_gleason9_34";
+
+Goal "Abs_prat (ratrel ^^ {(y, z)}) < xb + \
+\         Abs_prat (ratrel ^^ {(x*y, Abs_pnat 1)})*Abs_prat (ratrel ^^ {(w, x)})";
+by (res_inst_tac [("j","Abs_prat (ratrel ^^ {(x * y, Abs_pnat 1)}) *\
+\                   Abs_prat (ratrel ^^ {(w, x)})")] prat_le_less_trans 1);
+by (rtac prat_self_less_add_right 2);
+by (auto_tac (claset() addIs [lemma_Abs_prat_le3],
+    simpset() addsimps [prat_mult,pre_lemma_gleason9_34b,pnat_mult_assoc]));
+qed "lemma1b_gleason9_34";
+
+Goal "!!A. ! xa : Rep_preal(A). xa + x : Rep_preal(A) ==> False";
+by (cut_inst_tac [("X","A")] not_mem_Rep_preal_Ex 1);
+by (etac exE 1);
+by (dtac not_in_preal_ub 1);
+by (res_inst_tac [("z","x")] eq_Abs_prat 1);
+by (res_inst_tac [("z","xa")] eq_Abs_prat 1);
+by (dres_inst_tac [("p","y*xb")] lemma1_gleason9_34 1);
+by (etac bexE 1);
+by (cut_inst_tac [("x","y"),("y","xb"),("w","xaa"),
+    ("z","ya"),("xb","xba")] lemma1b_gleason9_34 1);
+by (dres_inst_tac [("x","xba + $#(y * xb) * x")]  bspec 1);
+by (auto_tac (claset() addIs [prat_less_asym],
+    simpset() addsimps [prat_pnat_def]));
+qed "lemma_gleason9_34a";
+
+Goal "? r: Rep_preal(R). r + x ~: Rep_preal(R)";
+by (rtac ccontr 1);
+by (blast_tac (claset() addIs [lemma_gleason9_34a]) 1);
+qed "lemma_gleason9_34";
+
+(*** Gleason's Lemma 9-3.6  ***)
+(*  lemmas for Gleason 9-3.6  *)
+(*                            *) 
+(******************************)
+
+Goal "r + r*qinv(xa)*Q3 = r*qinv(xa)*(xa + Q3)";
+by (full_simp_tac (simpset() addsimps [prat_add_mult_distrib2,
+    prat_mult_assoc]) 1);
+qed "lemma1_gleason9_36";
+
+Goal "r*qinv(xa)*(xa*x) = r*x";
+by (full_simp_tac (simpset() addsimps prat_mult_ac) 1);
+qed "lemma2_gleason9_36";
+(******)
+
+(*** FIXME: long! ***)
+Goal "!!A. $#1p < x ==> ? r: Rep_preal(A). r*x ~: Rep_preal(A)";
+by (res_inst_tac [("X1","A")] (mem_Rep_preal_Ex RS exE) 1);
+by (res_inst_tac [("Q","xa*x : Rep_preal(A)")] (excluded_middle RS disjE) 1);
+by (Fast_tac 1);
+by (dres_inst_tac [("x","xa")] prat_self_less_mult_right 1);
+by (etac prat_lessE 1);
+by (cut_inst_tac [("R","A"),("x","Q3")] lemma_gleason9_34 1);
+by (dtac sym 1 THEN Auto_tac );
+by (forward_tac [not_in_preal_ub] 1);
+by (dres_inst_tac [("x","xa + Q3")] bspec 1 THEN assume_tac 1);
+by (dtac prat_add_right_less_cancel 1);
+by (dres_inst_tac [("x","qinv(xa)*Q3")] prat_mult_less2_mono1 1);
+by (dres_inst_tac [("x","r")] prat_add_less2_mono2 1);
+by (asm_full_simp_tac (simpset() addsimps
+    [prat_mult_assoc RS sym,lemma1_gleason9_36]) 1);
+by (dtac sym 1);
+by (auto_tac (claset(),simpset() addsimps [lemma2_gleason9_36]));
+by (res_inst_tac [("x","r")] bexI 1);
+by (rtac notI 1);
+by (dres_inst_tac [("y","r*x")] (Rep_preal RS prealE_lemma3b) 1);
+by Auto_tac;
+qed "lemma_gleason9_36";
+
+Goal "!!A. $#Abs_pnat 1 < x ==> ? r: Rep_preal(A). r*x ~: Rep_preal(A)";
+by (rtac lemma_gleason9_36 1);
+by (asm_simp_tac (simpset() addsimps [pnat_one_def]) 1);
+qed "lemma_gleason9_36a";
+
+(*** Part 2 of existence of inverse ***)
+Goal "!!x. x: Rep_preal(@#($#Abs_pnat 1)) ==> x: Rep_preal(pinv(A)*A)";
+by (auto_tac (claset() addSIs [mem_Rep_preal_multI],
+              simpset() addsimps [pinv_def,preal_prat_def] ));
+by (rtac (preal_mem_inv_set RS Abs_preal_inverse RS ssubst) 1);
+by (dtac prat_qinv_gt_1 1);
+by (dres_inst_tac [("A","A")] lemma_gleason9_36a 1);
+by Auto_tac;
+by (dtac (Rep_preal RS prealE_lemma4a) 1);
+by (Auto_tac  THEN dtac qinv_prat_less 1);
+by (res_inst_tac [("x","qinv(u)*x")] exI 1);
+by (rtac conjI 1);
+by (res_inst_tac [("x","qinv(r)*x")] exI 1);
+by (auto_tac (claset() addIs [prat_mult_less2_mono1],
+    simpset() addsimps [qinv_mult_eq,qinv_qinv]));
+by (res_inst_tac [("x","u")] bexI 1);
+by (auto_tac (claset(),simpset() addsimps [prat_mult_assoc,
+    prat_mult_left_commute]));
+qed "preal_mem_mult_invI";
+
+Goal "pinv(A)*A = (@#($#Abs_pnat 1))";
+by (rtac (inj_Rep_preal RS injD) 1);
+by (rtac set_ext 1);
+by (fast_tac (claset() addDs [preal_mem_mult_invD,preal_mem_mult_invI]) 1);
+qed "preal_mult_inv";
+
+Goal "A*pinv(A) = (@#($#Abs_pnat 1))";
+by (rtac (preal_mult_commute RS subst) 1);
+by (rtac preal_mult_inv 1);
+qed "preal_mult_inv_right";
+
+val [prem] = goal thy
+    "(!!u. z = Abs_preal(u) ==> P) ==> P";
+by (cut_inst_tac [("x1","z")] 
+    (rewrite_rule [preal_def] (Rep_preal RS Abs_preal_inverse)) 1);
+by (res_inst_tac [("u","Rep_preal z")] prem 1);
+by (dtac (inj_Rep_preal RS injD) 1);
+by (Asm_simp_tac 1);
+qed "eq_Abs_preal";
+
+(*** Lemmas/Theorem(s) need lemma_gleason9_34 ***)
+Goal "Rep_preal (R1) <= Rep_preal(R1 + R2)";
+by (cut_inst_tac [("X","R2")] mem_Rep_preal_Ex 1);
+by (auto_tac (claset() addSIs [bexI] addIs [(Rep_preal RS prealE_lemma3b),
+   prat_self_less_add_left,mem_Rep_preal_addI],simpset()));
+qed "Rep_preal_self_subset";
+
+Goal "~ Rep_preal (R1 + R2) <= Rep_preal(R1)";
+by (cut_inst_tac [("X","R2")] mem_Rep_preal_Ex 1);
+by (etac exE 1);
+by (cut_inst_tac [("R","R1")] lemma_gleason9_34 1);
+by (auto_tac (claset() addIs [mem_Rep_preal_addI],simpset()));
+qed "Rep_preal_sum_not_subset";
+
+Goal "Rep_preal (R1 + R2) ~= Rep_preal(R1)";
+by (rtac notI 1);
+by (etac equalityE 1);
+by (asm_full_simp_tac (simpset() addsimps [Rep_preal_sum_not_subset]) 1);
+qed "Rep_preal_sum_not_eq";
+
+(*** at last --- Gleason prop. 9-3.5(iii) p. 123 ***)
+Goalw [preal_less_def,psubset_def] "(R1::preal) < R1 + R2";
+by (simp_tac (simpset() addsimps [Rep_preal_self_subset,
+    Rep_preal_sum_not_eq RS not_sym]) 1);
+qed "preal_self_less_add_left";
+
+Goal "(R1::preal) < R2 + R1";
+by (simp_tac (simpset() addsimps [preal_add_commute,preal_self_less_add_left]) 1);
+qed "preal_self_less_add_right";
+
+(*** Properties of <= ***)
+
+Goalw [preal_le_def,psubset_def,preal_less_def] 
+                     "!!w. z<=w ==> ~(w<(z::preal))";
+by (auto_tac  (claset() addDs [equalityI],simpset()));
+qed "preal_leD";
+
+val preal_leE = make_elim preal_leD;
+
+Goalw [preal_le_def,psubset_def,preal_less_def]
+                   "!!z. ~ z <= w ==> w<(z::preal)";
+by (cut_inst_tac [("r1.0","w"),("r2.0","z")] preal_linear 1);
+by (auto_tac  (claset(),simpset() addsimps [preal_less_def,psubset_def]));
+qed "not_preal_leE";
+		       
+Goal "!!w. ~(w < z) ==> z <= (w::preal)";
+by (fast_tac (claset() addIs [not_preal_leE]) 1);
+qed "preal_leI";
+
+Goal "!!w. (~(w < z)) = (z <= (w::preal))";
+by (fast_tac (claset() addSIs [preal_leI,preal_leD]) 1);
+qed "preal_less_le_iff";
+
+Goalw [preal_le_def,preal_less_def,psubset_def] 
+                  "!!z. z < w ==> z <= (w::preal)";
+by (Fast_tac 1);
+qed "preal_less_imp_le";
+
+Goalw [preal_le_def,preal_less_def,psubset_def] 
+                      "!!(x::preal). x <= y ==> x < y | x = y";
+by (auto_tac (claset() addIs [inj_Rep_preal RS injD],simpset()));
+qed "preal_le_imp_less_or_eq";
+
+Goalw [preal_le_def,preal_less_def,psubset_def] 
+                       "!!(x::preal). x < y | x = y ==> x <=y";
+by Auto_tac;
+qed "preal_less_or_eq_imp_le";
+
+Goal "(x <= (y::preal)) = (x < y | x=y)";
+by (REPEAT(ares_tac [iffI, preal_less_or_eq_imp_le, preal_le_imp_less_or_eq] 1));
+qed "preal_le_eq_less_or_eq";
+
+Goalw [preal_le_def] "w <= (w::preal)";
+by (Simp_tac 1);
+qed "preal_le_refl";
+
+val prems = goal thy "!!i. [| i <= j; j < k |] ==> i < (k::preal)";
+by (dtac preal_le_imp_less_or_eq 1);
+by (fast_tac (claset() addIs [preal_less_trans]) 1);
+qed "preal_le_less_trans";
+
+val prems = goal thy "!!i. [| i < j; j <= k |] ==> i < (k::preal)";
+by (dtac preal_le_imp_less_or_eq 1);
+by (fast_tac (claset() addIs [preal_less_trans]) 1);
+qed "preal_less_le_trans";
+
+Goal "!!i. [| i <= j; j <= k |] ==> i <= (k::preal)";
+by (EVERY1 [dtac preal_le_imp_less_or_eq, dtac preal_le_imp_less_or_eq,
+            rtac preal_less_or_eq_imp_le, fast_tac (claset() addIs [preal_less_trans])]);
+qed "preal_le_trans";
+
+Goal "!!z. [| z <= w; w <= z |] ==> z = (w::preal)";
+by (EVERY1 [dtac preal_le_imp_less_or_eq, dtac preal_le_imp_less_or_eq,
+            fast_tac (claset() addEs [preal_less_irrefl,preal_less_asym])]);
+qed "preal_le_anti_sym";
+
+Goal "!!x. [| ~ y < x; y ~= x |] ==> x < (y::preal)";
+by (rtac not_preal_leE 1);
+by (fast_tac (claset() addDs [preal_le_imp_less_or_eq]) 1);
+qed "not_less_not_eq_preal_less";
+
+(****)(****)(****)(****)(****)(****)(****)(****)(****)(****)(****)(****)(****)(****)
+
+(**** Set up all lemmas for proving A < B ==> ?D. A + D = B ****)
+(**** Gleason prop. 9-3.5(iv) p. 123 ****)
+(**** Define the D required and show that it is a positive real ****)
+
+(* useful lemmas - proved elsewhere? *)
+Goalw [psubset_def] "!!A. A < B ==> ? x. x ~: A & x : B";
+by (etac conjE 1);
+by (etac swap 1);
+by (etac equalityI 1);
+by Auto_tac;
+qed "lemma_psubset_mem";
+
+Goalw [psubset_def] "~ (A::'a set) < A";
+by (Fast_tac 1);
+qed "lemma_psubset_not_refl";
+
+Goalw [psubset_def] "!!(A::'a set). [| A < B; B < C |] ==> A < C";
+by (auto_tac (claset() addDs [subset_antisym],simpset()));
+qed "psubset_trans";
+
+Goalw [psubset_def] "!!(A::'a set). [| A <= B; B < C |] ==> A < C";
+by (auto_tac (claset() addDs [subset_antisym],simpset()));
+qed "subset_psubset_trans";
+
+Goalw [psubset_def] "!!(A::'a set). [| A < B; B <= C |] ==> A < C";
+by (auto_tac (claset() addDs [subset_antisym],simpset()));
+qed "subset_psubset_trans2";
+
+Goalw [psubset_def] "!!(A::'a set). [| A < B; c : A |] ==> c : B";
+by (auto_tac (claset() addDs [subsetD],simpset()));
+qed "psubsetD";
+
+(** Part 1 of Dedekind sections def **)
+Goalw [preal_less_def]
+     "!!A. A < B ==> \
+\     ? q. q : {d. ? n. n ~: Rep_preal(A) & n + d : Rep_preal(B)}";
+by (EVERY1[dtac lemma_psubset_mem, etac exE, etac conjE]);
+by (dres_inst_tac [("x1","B")] (Rep_preal RS prealE_lemma4a) 1);
+by (auto_tac (claset(),simpset() addsimps [prat_less_def]));
+qed "lemma_ex_mem_less_left_add1";
+
+Goal
+     "!!A. A < B ==> \
+\       {} < {d. ? n. n ~: Rep_preal(A) & n + d : Rep_preal(B)}";
+by (dtac lemma_ex_mem_less_left_add1 1);
+by (auto_tac (claset() addSIs [psubsetI] addEs [equalityCE],simpset()));
+qed "preal_less_set_not_empty";
+
+(** Part 2 of Dedekind sections def **)
+Goal "? q. q ~: {d. ? n. n ~: Rep_preal(A) & n + d : Rep_preal(B)}";
+by (cut_inst_tac [("X","B")] not_mem_Rep_preal_Ex 1);
+by (etac exE 1);
+by (res_inst_tac [("x","x")] exI 1);
+by Auto_tac;
+by (cut_inst_tac [("x","x"),("y","n")] prat_self_less_add_right 1);
+by (auto_tac (claset() addDs [Rep_preal RS prealE_lemma3b],simpset()));
+qed "lemma_ex_not_mem_less_left_add1";
+
+Goal "{d. ? n. n ~: Rep_preal(A) & n + d : Rep_preal(B)} < {q. True}";
+by (auto_tac (claset() addSIs [psubsetI],simpset()));
+by (cut_inst_tac [("A","A"),("B","B")] lemma_ex_not_mem_less_left_add1 1);
+by (etac exE 1);
+by (eres_inst_tac [("c","q")] equalityCE 1);
+by Auto_tac;
+qed "preal_less_set_not_prat_set";
+
+(** Part 3 of Dedekind sections def **)
+Goal "!!A. A < B ==> ! y: {d. ? n. n ~: Rep_preal(A) & n + d : Rep_preal(B)}. \
+ \      !z. z < y --> z : {d. ? n. n ~: Rep_preal(A) & n + d : Rep_preal(B)}";
+by Auto_tac;
+by (dres_inst_tac [("x","n")] prat_add_less2_mono2 1);
+by (dtac (Rep_preal RS prealE_lemma3b) 1);
+by Auto_tac;
+qed "preal_less_set_lemma3";
+
+Goal "!!A. A < B ==> ! y: {d. ? n. n ~: Rep_preal(A) & n + d : Rep_preal(B)}. \
+\       Bex {d. ? n. n ~: Rep_preal(A) & n + d : Rep_preal(B)} (op < y)";
+by Auto_tac;
+by (dtac (Rep_preal RS prealE_lemma4a) 1);
+by (auto_tac (claset(),simpset() addsimps [prat_less_def,prat_add_assoc]));
+qed "preal_less_set_lemma4";
+
+Goal 
+     "!! (A::preal). A < B ==> \
+\     {d. ? n. n ~: Rep_preal(A) & n + d : Rep_preal(B)}: preal";
+by (rtac prealI2 1);
+by (rtac preal_less_set_not_empty 1);
+by (rtac preal_less_set_not_prat_set 2);
+by (rtac preal_less_set_lemma3 2);
+by (rtac preal_less_set_lemma4 3);
+by Auto_tac;
+qed "preal_mem_less_set";
+
+(** proving that A + D <= B **)
+Goalw [preal_le_def] 
+       "!! (A::preal). A < B ==> \
+\         A + Abs_preal({d. ? n. n ~: Rep_preal(A) & n + d : Rep_preal(B)}) <= B";
+by (rtac subsetI 1);
+by (dtac mem_Rep_preal_addD 1);
+by (auto_tac (claset(),simpset() addsimps [
+    preal_mem_less_set RS Abs_preal_inverse]));
+by (dtac not_in_preal_ub 1);
+by (dtac bspec 1 THEN assume_tac 1);
+by (dres_inst_tac [("x","y")] prat_add_less2_mono1 1);
+by (dres_inst_tac [("x1","B")] (Rep_preal RS prealE_lemma3b) 1);
+by Auto_tac;
+qed "preal_less_add_left_subsetI";
+
+(** proving that B <= A + D  --- trickier **)
+(** lemma **)
+Goal "!!x. x : Rep_preal(B) ==> ? e. x + e : Rep_preal(B)";
+by (dtac (Rep_preal RS prealE_lemma4a) 1);
+by (auto_tac (claset(),simpset() addsimps [prat_less_def]));
+qed "lemma_sum_mem_Rep_preal_ex";
+
+Goalw [preal_le_def] 
+       "!! (A::preal). A < B ==> \
+\         B <= A + Abs_preal({d. ? n. n ~: Rep_preal(A) & n + d : Rep_preal(B)})";
+by (rtac subsetI 1);
+by (res_inst_tac [("Q","x: Rep_preal(A)")] (excluded_middle RS disjE) 1);
+by (rtac mem_Rep_preal_addI 1);
+by (dtac lemma_sum_mem_Rep_preal_ex 1);
+by (etac exE 1);
+by (cut_inst_tac [("R","A"),("x","e")] lemma_gleason9_34 1 THEN etac bexE 1);
+by (dtac not_in_preal_ub 1 THEN dtac bspec 1 THEN assume_tac 1);
+by (etac prat_lessE 1);
+by (res_inst_tac [("x","r")] bexI 1);
+by (res_inst_tac [("x","Q3")] bexI 1);
+by (cut_facts_tac [Rep_preal_self_subset] 4);
+by (auto_tac (claset(),simpset() addsimps [
+    preal_mem_less_set RS Abs_preal_inverse]));
+by (res_inst_tac [("x","r+e")] exI 1);
+by (asm_full_simp_tac (simpset() addsimps prat_add_ac) 1);
+qed "preal_less_add_left_subsetI2";
+
+(*** required proof ***)
+Goal "!! (A::preal). A < B ==> \
+\         A + Abs_preal({d. ? n. n ~: Rep_preal(A) & n + d : Rep_preal(B)}) = B";
+by (blast_tac (claset() addIs [preal_le_anti_sym,
+                preal_less_add_left_subsetI,preal_less_add_left_subsetI2]) 1);
+qed "preal_less_add_left";
+
+Goal "!! (A::preal). A < B ==> ? D. A + D = B";
+by (fast_tac (claset() addDs [preal_less_add_left]) 1);
+qed "preal_less_add_left_Ex";        
+
+Goal "!!(A::preal). A < B ==> A + C < B + C";
+by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
+    simpset() addsimps [preal_add_assoc]));
+by (res_inst_tac [("y1","D")] (preal_add_commute RS subst) 1);
+by (auto_tac (claset() addIs [preal_self_less_add_left],
+          simpset() addsimps [preal_add_assoc RS sym]));
+qed "preal_add_less2_mono1";
+
+Goal "!!(A::preal). A < B ==> C + A < C + B";
+by (auto_tac (claset() addIs [preal_add_less2_mono1],
+    simpset() addsimps [preal_add_commute]));
+qed "preal_add_less2_mono2";
+
+Goal 
+      "!!(q1::preal). q1 < q2 ==> q1 * x < q2 * x";
+by (dtac preal_less_add_left_Ex 1);
+by (auto_tac (claset(),simpset() addsimps [preal_add_mult_distrib,
+    preal_self_less_add_left]));
+qed "preal_mult_less_mono1";
+
+Goal "!!(q1::preal). q1 < q2  ==> x * q1 < x * q2";
+by (auto_tac (claset() addDs [preal_mult_less_mono1],
+    simpset() addsimps [preal_mult_commute]));
+qed "preal_mult_left_less_mono1";
+
+Goal "!!(q1::preal). q1 <= q2  ==> x * q1 <= x * q2";
+by (dtac preal_le_imp_less_or_eq 1);
+by (Step_tac 1);
+by (auto_tac (claset() addSIs [preal_le_refl,
+    preal_less_imp_le,preal_mult_left_less_mono1],simpset()));
+qed "preal_mult_left_le_mono1";
+ 
+Goal "!!(q1::preal). q1 <= q2  ==> q1 * x <= q2 * x";
+by (auto_tac (claset() addDs [preal_mult_left_le_mono1],
+    simpset() addsimps [preal_mult_commute]));
+qed "preal_mult_le_mono1";
+ 
+Goal "!!(q1::preal). q1 <= q2  ==> x + q1 <= x + q2";
+by (dtac preal_le_imp_less_or_eq 1);
+by (Step_tac 1);
+by (auto_tac (claset() addSIs [preal_le_refl,
+    preal_less_imp_le,preal_add_less2_mono1],
+    simpset() addsimps [preal_add_commute]));
+qed "preal_add_left_le_mono1";
+
+Goal "!!(q1::preal). q1 <= q2  ==> q1 + x <= q2 + x";
+by (auto_tac (claset() addDs [preal_add_left_le_mono1],
+    simpset() addsimps [preal_add_commute]));
+qed "preal_add_le_mono1";
+ 
+Goal "!!k l::preal. [|i<=j;  k<=l |] ==> i + k <= j + l";
+by (etac (preal_add_le_mono1 RS preal_le_trans) 1);
+by (simp_tac (simpset() addsimps [preal_add_commute]) 1);
+(*j moves to the end because it is free while k, l are bound*)
+by (etac preal_add_le_mono1 1);
+qed "preal_add_le_mono";
+
+Goal "!!(A::preal). A + C < B + C ==> A < B";
+by (cut_facts_tac [preal_linear] 1);
+by (auto_tac (claset() addEs [preal_less_irrefl],simpset()));
+by (dres_inst_tac [("A","B"),("C","C")] preal_add_less2_mono1 1);
+by (fast_tac (claset() addDs [preal_less_trans] 
+                addEs [preal_less_irrefl]) 1);
+qed "preal_add_right_less_cancel";
+
+Goal "!!(A::preal). C + A < C + B ==> A < B";
+by (auto_tac (claset() addEs [preal_add_right_less_cancel],
+              simpset() addsimps [preal_add_commute]));
+qed "preal_add_left_less_cancel";
+
+Goal "((A::preal) + C < B + C) = (A < B)";
+by (REPEAT(ares_tac [iffI,preal_add_less2_mono1,
+    preal_add_right_less_cancel] 1));
+qed "preal_add_less_iff1";
+
+Addsimps [preal_add_less_iff1];
+
+Goal "(C + (A::preal) < C + B) = (A < B)";
+by (REPEAT(ares_tac [iffI,preal_add_less2_mono2,
+    preal_add_left_less_cancel] 1));
+qed "preal_add_less_iff2";
+
+Addsimps [preal_add_less_iff2];
+
+Goal 
+      "!!x1. [| x1 < y1; x2 < y2 |] ==> x1 + x2 < y1 + (y2::preal)";
+by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
+    simpset() addsimps  preal_add_ac));
+by (rtac (preal_add_assoc RS subst) 1);
+by (rtac preal_self_less_add_right 1);
+qed "preal_add_less_mono";
+
+Goal 
+      "!!x1. [| x1 < y1; x2 < y2 |] ==> x1 * x2 < y1 * (y2::preal)";
+by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
+              simpset() addsimps [preal_add_mult_distrib,
+              preal_add_mult_distrib2,preal_self_less_add_left,
+              preal_add_assoc] @ preal_mult_ac));
+qed "preal_mult_less_mono";
+
+Goal "!!(A::preal). A + C = B + C ==> A = B";
+by (cut_facts_tac [preal_linear] 1);
+by Auto_tac;
+by (ALLGOALS(dres_inst_tac [("C","C")] preal_add_less2_mono1));
+by (auto_tac (claset() addEs [preal_less_irrefl],simpset()));
+qed "preal_add_right_cancel";
+
+Goal "!!(A::preal). C + A = C + B ==> A = B";
+by (auto_tac (claset() addIs [preal_add_right_cancel],
+              simpset() addsimps [preal_add_commute]));
+qed "preal_add_left_cancel";
+
+Goal "(C + A = C + B) = ((A::preal) = B)";
+by (fast_tac (claset() addIs [preal_add_left_cancel]) 1);
+qed "preal_add_left_cancel_iff";
+
+Goal "(A + C = B + C) = ((A::preal) = B)";
+by (fast_tac (claset() addIs [preal_add_right_cancel]) 1);
+qed "preal_add_right_cancel_iff";
+
+Addsimps [preal_add_left_cancel_iff,preal_add_right_cancel_iff];
+
+(*** Completeness of preal ***)
+
+(*** prove that supremum is a cut ***)
+Goal "!!P. ? (X::preal). X: P ==> \
+\         ? q.  q: {w. ? X. X : P & w : Rep_preal X}";
+by Safe_tac;
+by (cut_inst_tac [("X","X")] mem_Rep_preal_Ex 1);
+by Auto_tac;
+qed "preal_sup_mem_Ex";
+
+(** Part 1 of Dedekind def **)
+Goal "!!P. ? (X::preal). X: P ==> \
+\         {} < {w. ? X : P. w : Rep_preal X}";
+by (dtac preal_sup_mem_Ex 1);
+by (auto_tac (claset() addSIs [psubsetI] addEs [equalityCE],simpset()));
+qed "preal_sup_set_not_empty";
+
+(** Part 2 of Dedekind sections def **) 
+Goalw [preal_less_def] "!!P. ? Y. (! X: P. X < Y)  \             
+\         ==> ? q. q ~: {w. ? X. X: P & w: Rep_preal(X)}"; (**)
+by (auto_tac (claset(),simpset() addsimps [psubset_def]));
+by (cut_inst_tac [("X","Y")] not_mem_Rep_preal_Ex 1);
+by (etac exE 1);
+by (res_inst_tac [("x","x")] exI 1);
+by (auto_tac (claset() addSDs [bspec],simpset()));
+qed "preal_sup_not_mem_Ex";
+
+Goalw [preal_le_def] "!!P. ? Y. (! X: P. X <= Y)  \
+\         ==> ? q. q ~: {w. ? X. X: P & w: Rep_preal(X)}";
+by (Step_tac 1);
+by (cut_inst_tac [("X","Y")] not_mem_Rep_preal_Ex 1);
+by (etac exE 1);
+by (res_inst_tac [("x","x")] exI 1);
+by (auto_tac (claset() addSDs [bspec],simpset()));
+qed "preal_sup_not_mem_Ex1";
+
+Goal "!!P. ? Y. (! X: P. X < Y)  \                                    
+\         ==> {w. ? X: P. w: Rep_preal(X)} < {q. True}";       (**)
+by (dtac preal_sup_not_mem_Ex 1);
+by (auto_tac (claset() addSIs [psubsetI],simpset()));
+by (eres_inst_tac [("c","q")] equalityCE 1);
+by Auto_tac;
+qed "preal_sup_set_not_prat_set";
+
+Goal "!!P. ? Y. (! X: P. X <= Y)  \
+\         ==> {w. ? X: P. w: Rep_preal(X)} < {q. True}";
+by (dtac preal_sup_not_mem_Ex1 1);
+by (auto_tac (claset() addSIs [psubsetI],simpset()));
+by (eres_inst_tac [("c","q")] equalityCE 1);
+by Auto_tac;
+qed "preal_sup_set_not_prat_set1";
+
+(** Part 3 of Dedekind sections def **)
+Goal "!!P. [|? (X::preal). X: P; ? Y. (! X:P. X < Y) |] \              
+\         ==> ! y: {w. ? X: P. w: Rep_preal X}. \
+\             !z. z < y --> z: {w. ? X: P. w: Rep_preal X}";         (**)
+by (auto_tac(claset() addEs [Rep_preal RS prealE_lemma3b],simpset()));
+qed "preal_sup_set_lemma3";
+
+Goal "!!P. [|? (X::preal). X: P; ? Y. (! X:P. X <= Y) |] \
+\         ==> ! y: {w. ? X: P. w: Rep_preal X}. \
+\             !z. z < y --> z: {w. ? X: P. w: Rep_preal X}";
+by (auto_tac(claset() addEs [Rep_preal RS prealE_lemma3b],simpset()));
+qed "preal_sup_set_lemma3_1";
+
+Goal "!!P. [|? (X::preal). X: P; ? Y. (! X:P. X < Y) |] \              
+\         ==>  !y: {w. ? X: P. w: Rep_preal X}. \                        
+\             Bex {w. ? X: P. w: Rep_preal X} (op < y)";                (**)
+by (blast_tac (claset() addDs [(Rep_preal RS prealE_lemma4a)]) 1);
+qed "preal_sup_set_lemma4";
+
+Goal "!!P. [|? (X::preal). X: P; ? Y. (! X:P. X <= Y) |] \
+\         ==>  !y: {w. ? X: P. w: Rep_preal X}. \
+\             Bex {w. ? X: P. w: Rep_preal X} (op < y)";
+by (blast_tac (claset() addDs [(Rep_preal RS prealE_lemma4a)]) 1);
+qed "preal_sup_set_lemma4_1";
+
+Goal "!!P. [|? (X::preal). X: P; ? Y. (! X:P. X < Y) |] \            
+\         ==> {w. ? X: P. w: Rep_preal(X)}: preal";                      (**)
+by (rtac prealI2 1);
+by (rtac preal_sup_set_not_empty 1);
+by (rtac preal_sup_set_not_prat_set 2);
+by (rtac preal_sup_set_lemma3 3);
+by (rtac preal_sup_set_lemma4 5);
+by Auto_tac;
+qed "preal_sup";
+
+Goal "!!P. [|? (X::preal). X: P; ? Y. (! X:P. X <= Y) |] \
+\         ==> {w. ? X: P. w: Rep_preal(X)}: preal";
+by (rtac prealI2 1);
+by (rtac preal_sup_set_not_empty 1);
+by (rtac preal_sup_set_not_prat_set1 2);
+by (rtac preal_sup_set_lemma3_1 3);
+by (rtac preal_sup_set_lemma4_1 5);
+by Auto_tac;
+qed "preal_sup1";
+
+Goalw [psup_def] "!!P. ? Y. (! X:P. X < Y) ==> ! x: P. x <= psup P";      (**) 
+by (auto_tac (claset(),simpset() addsimps [preal_le_def]));
+by (rtac (preal_sup RS Abs_preal_inverse RS ssubst) 1);
+by Auto_tac;
+qed "preal_psup_leI";
+
+Goalw [psup_def] "!!P. ? Y. (! X:P. X <= Y) ==> ! x: P. x <= psup P";
+by (auto_tac (claset(),simpset() addsimps [preal_le_def]));
+by (rtac (preal_sup1 RS Abs_preal_inverse RS ssubst) 1);
+by (auto_tac (claset(),simpset() addsimps [preal_le_def]));
+qed "preal_psup_leI2";
+
+Goal "!!P. [| ? Y. (! X:P. X < Y); x : P |] ==> x <= psup P";              (**)
+by (blast_tac (claset() addSDs [preal_psup_leI]) 1);
+qed "preal_psup_leI2b";
+
+Goal "!!P. [| ? Y. (! X:P. X <= Y); x : P |] ==> x <= psup P";
+by (blast_tac (claset() addSDs [preal_psup_leI2]) 1);
+qed "preal_psup_leI2a";
+
+Goalw [psup_def] "!!P. [| ? X. X : P; ! X:P. X < Y |] ==> psup P <= Y";   (**)
+by (auto_tac (claset(),simpset() addsimps [preal_le_def]));
+by (dtac (([exI,exI] MRS preal_sup) RS Abs_preal_inverse RS subst) 1);
+by (rotate_tac 1 2);
+by (assume_tac 2);
+by (auto_tac (claset() addSDs [bspec],simpset() addsimps [preal_less_def,psubset_def]));
+qed "psup_le_ub";
+
+Goalw [psup_def] "!!P. [| ? X. X : P; ! X:P. X <= Y |] ==> psup P <= Y";
+by (auto_tac (claset(),simpset() addsimps [preal_le_def]));
+by (dtac (([exI,exI] MRS preal_sup1) RS Abs_preal_inverse RS subst) 1);
+by (rotate_tac 1 2);
+by (assume_tac 2);
+by (auto_tac (claset() addSDs [bspec],
+    simpset() addsimps [preal_less_def,psubset_def,preal_le_def]));
+qed "psup_le_ub1";
+
+(** supremum property **)
+Goal "!!P. [|? (X::preal). X: P; ? Y. (! X:P. X < Y) |] \                  
+\         ==> (!Y. (? X: P. Y < X) = (Y < psup P))";              
+by (forward_tac [preal_sup RS Abs_preal_inverse] 1);
+by (Fast_tac 1);
+by (auto_tac (claset() addSIs [psubsetI],simpset() addsimps [psup_def,preal_less_def]));
+by (blast_tac (claset() addDs [psubset_def RS meta_eq_to_obj_eq RS iffD1]) 1);
+by (rotate_tac 4 1);
+by (asm_full_simp_tac (simpset() addsimps [psubset_def]) 1);
+by (dtac bspec 1 THEN assume_tac 1);
+by (REPEAT(etac conjE 1));
+by (EVERY1[rtac swap, assume_tac, rtac set_ext]);
+by (auto_tac (claset() addSDs [lemma_psubset_mem],simpset()));
+by (cut_inst_tac [("r1.0","Xa"),("r2.0","Ya")] preal_linear 1);
+by (auto_tac (claset() addDs [psubsetD],simpset() addsimps [preal_less_def]));
+qed "preal_complete";
+
+(****)(****)(****)(****)(****)(****)(****)(****)(****)(****)
+    (****** Embedding ******)
+(*** mapping from prat into preal ***)
+
+Goal "!!z1. x < z1 + z2 ==> x * z1 * qinv (z1 + z2) < z1";
+by (dres_inst_tac [("x","z1 * qinv (z1 + z2)")] prat_mult_less2_mono1 1);
+by (asm_full_simp_tac (simpset() addsimps prat_mult_ac) 1);
+qed "lemma_preal_rat_less";
+
+Goal "!!z1. x < z1 + z2 ==> x * z2 * qinv (z1 + z2) < z2";
+by (stac prat_add_commute 1);
+by (dtac (prat_add_commute RS subst) 1);
+by (etac lemma_preal_rat_less 1);
+qed "lemma_preal_rat_less2";
+
+Goalw [preal_prat_def,preal_add_def] 
+            "@#((z1::prat) + z2) = @#z1 + @#z2";
+by (res_inst_tac [("f","Abs_preal")] arg_cong 1);
+by (auto_tac (claset() addIs [prat_add_less_mono] addSIs [set_ext],simpset() addsimps 
+    [lemma_prat_less_set_mem_preal RS Abs_preal_inverse]));
+by (res_inst_tac [("x","x*z1*qinv(z1+z2)")] exI 1 THEN rtac conjI 1);
+by (etac lemma_preal_rat_less 1);
+by (res_inst_tac [("x","x*z2*qinv(z1+z2)")] exI 1 THEN rtac conjI 1);
+by (etac lemma_preal_rat_less2 1);
+by (asm_full_simp_tac (simpset() addsimps [prat_add_mult_distrib RS sym,
+     prat_add_mult_distrib2 RS sym] @ prat_mult_ac) 1);
+qed "preal_prat_add";
+
+Goal "!!xa. x < xa ==> x*z1*qinv(xa) < z1";
+by (dres_inst_tac [("x","z1 * qinv xa")] prat_mult_less2_mono1 1);
+by (dtac (prat_mult_left_commute RS subst) 1);
+by (asm_full_simp_tac (simpset() addsimps prat_mult_ac) 1);
+qed "lemma_preal_rat_less3";
+
+Goal "!!xa. xa < z1 * z2 ==> xa*z2*qinv(z1*z2) < z2";
+by (dres_inst_tac [("x","z2 * qinv(z1*z2)")] prat_mult_less2_mono1 1);
+by (dtac (prat_mult_left_commute RS subst) 1);
+by (asm_full_simp_tac (simpset() addsimps prat_mult_ac) 1);
+qed "lemma_preal_rat_less4";
+
+Goalw [preal_prat_def,preal_mult_def] 
+            "@#((z1::prat) * z2) = @#z1 * @#z2";
+by (res_inst_tac [("f","Abs_preal")] arg_cong 1);
+by (auto_tac (claset() addIs [prat_mult_less_mono] addSIs [set_ext],simpset() addsimps 
+    [lemma_prat_less_set_mem_preal RS Abs_preal_inverse]));
+by (dtac prat_dense 1);
+by (Step_tac 1);
+by (res_inst_tac [("x","x*z1*qinv(xa)")] exI 1 THEN rtac conjI 1);
+by (etac lemma_preal_rat_less3 1);
+by (res_inst_tac [("x"," xa*z2*qinv(z1*z2)")] exI 1 THEN rtac conjI 1);
+by (etac lemma_preal_rat_less4 1);
+by (asm_full_simp_tac (simpset() 
+    addsimps [qinv_mult_eq RS sym] @ prat_mult_ac) 1);
+by (asm_full_simp_tac (simpset() 
+    addsimps [prat_mult_assoc RS sym]) 1);
+qed "preal_prat_mult";
+
+Goalw [preal_prat_def,preal_less_def] 
+      "(@#p < @#q) = (p < q)";
+by (auto_tac (claset() addSDs [lemma_prat_set_eq] addEs [prat_less_trans],
+    simpset() addsimps [lemma_prat_less_set_mem_preal,
+    psubset_def,prat_less_not_refl]));
+by (res_inst_tac [("q1.0","p"),("q2.0","q")] prat_linear_less2 1);
+by (auto_tac (claset() addIs [prat_less_irrefl],simpset()));
+qed "preal_prat_less_iff";
+
+Addsimps [preal_prat_less_iff];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/PReal.thy	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,42 @@
+(*   Title       : PReal.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : The positive reals as Dedekind sections of positive
+                  rationals. Fundamentals of Abstract Analysis [Gleason- p. 121] 
+                  provides some of the definitions.
+*)
+
+PReal = PRat +
+
+typedef preal = "{A::prat set. {} < A & A < {q::prat. True} &
+                               (!y: A. ((!z. z < y --> z: A) &
+                                        (? u: A. y < u)))}"      (preal_1)
+instance
+   preal :: {ord, plus, times}
+
+constdefs
+  preal_prat :: prat => preal              ("@#_" [80] 80)
+   "@# q     == Abs_preal({x::prat. x < q})"
+
+  pinv       :: preal => preal
+  "pinv(R)   == Abs_preal({w. ? y. w < y & qinv y ~: Rep_preal(R)})" 
+
+  psup       :: preal set => preal
+  "psup(P)   == Abs_preal({w. ? X: P. w: Rep_preal(X)})"
+
+defs
+
+  preal_add_def
+        "R + S == Abs_preal({w. ? x: Rep_preal(R). ? y: Rep_preal(S). w = x + y})"
+
+  preal_mult_def
+        "R * S == Abs_preal({w. ? x: Rep_preal(R). ? y: Rep_preal(S). w = x * y})"
+
+  preal_less_def
+        "R < (S::preal) == Rep_preal(R) < Rep_preal(S)"
+
+  preal_le_def
+        "R <= (S::preal) == Rep_preal(R) <= Rep_preal(S)"
+ 
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/RComplete.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,218 @@
+(*  Title       : RComplete.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Completeness theorems for positive
+                  reals and reals 
+*) 
+
+
+open RComplete;
+
+Goal "!!(x::real). [| isLub R S x; isLub R S y |] ==> x = y";
+by (forward_tac [isLub_isUb] 1);
+by (forw_inst_tac [("x","y")] isLub_isUb 1);
+by (blast_tac (claset() addSIs [real_le_anti_sym]
+                addSDs [isLub_le_isUb]) 1);
+qed "real_isLub_unique";
+
+Goalw [setle_def,setge_def] 
+         "!!x::real. [| x <=* S'; S <= S' |] ==> x <=* S";
+by (Blast_tac 1);
+qed "real_order_restrict";
+
+(*----------------------------------------------------------------
+           Completeness theorem for the positive reals(again)
+ ----------------------------------------------------------------*)
+
+Goal "!!S. [| ALL x: S. 0r < x; \
+\                 EX x. x: S; \
+\                 EX u. isUb (UNIV::real set) S u \
+\              |] ==> EX t. isLub (UNIV::real set) S t";
+by (res_inst_tac [("x","%#psup({w. %#w : S})")] exI 1);
+by (auto_tac (claset(),simpset() addsimps [isLub_def,leastP_def,isUb_def]));
+by (auto_tac (claset() addSIs [setleI,setgeI] 
+    addSDs [real_gt_zero_preal_Ex RS iffD1],simpset()));
+by (forw_inst_tac [("x","y")] bspec 1 THEN assume_tac 1);
+by (dtac (real_gt_zero_preal_Ex RS iffD1) 1);
+by (auto_tac (claset(),simpset() addsimps [real_preal_le_iff]));
+by (rtac preal_psup_leI2a 1);
+by (forw_inst_tac [("y","%#ya")] setleD 1 THEN assume_tac 1);
+by (forward_tac  [real_ge_preal_preal_Ex] 1);
+by (Step_tac 1);
+by (res_inst_tac [("x","y")] exI 1);
+by (blast_tac (claset() addSDs [setleD] addIs [real_preal_le_iff RS iffD1]) 1);
+by (forw_inst_tac [("x","x")] bspec 1 THEN assume_tac 1);
+by (forward_tac [isUbD2] 1);
+by (dtac (real_gt_zero_preal_Ex RS iffD1) 1);
+by (auto_tac (claset() addSDs [isUbD,
+    real_ge_preal_preal_Ex],simpset() addsimps [real_preal_le_iff]));
+by (blast_tac (claset() addSDs [setleD] addSIs 
+    [psup_le_ub1] addIs [real_preal_le_iff RS iffD1]) 1);
+qed "posreals_complete";
+
+
+(*-------------------------------
+    Lemmas
+ -------------------------------*)
+Goal "! y : {z. ? x: P. z = x + %~xa + 1r} Int {x. 0r < x}. 0r < y";
+by Auto_tac;
+qed "real_sup_lemma3";
+ 
+(* lemmas re-arranging the terms *)
+Goal "(S <= Y + %~X + Z) = (S + X + %~Z <= Y)";
+by (Step_tac 1);
+by (dres_inst_tac [("x","%~Z")] real_add_le_mono1 1);
+by (dres_inst_tac [("x","Z")] real_add_le_mono1 2);
+by (auto_tac (claset(),simpset() addsimps [real_add_assoc,
+    real_add_minus,real_add_zero_right,real_add_minus_left]));
+by (dres_inst_tac [("x","X")] real_add_le_mono1 1);
+by (dres_inst_tac [("x","%~X")] real_add_le_mono1 2);
+by (auto_tac (claset(),simpset() addsimps [real_add_assoc,
+    real_add_minus,real_add_zero_right,real_add_minus_left]));
+by (auto_tac (claset(),simpset() addsimps [real_add_commute]));
+qed "lemma_le_swap";
+
+Goal "(xa <= S + X + %~Z) = (xa + %~X + Z <= S)";
+by (Step_tac 1);
+by (dres_inst_tac [("x","Z")] real_add_le_mono1 1);
+by (dres_inst_tac [("x","%~Z")] real_add_le_mono1 2);
+by (auto_tac (claset(),simpset() addsimps [real_add_assoc,
+    real_add_minus,real_add_zero_right,real_add_minus_left]));
+by (dres_inst_tac [("x","%~X")] real_add_le_mono1 1);
+by (dres_inst_tac [("x","X")] real_add_le_mono1 2);
+by (auto_tac (claset(),simpset() addsimps [real_add_assoc,
+    real_add_minus,real_add_zero_right,real_add_minus_left]));
+by (auto_tac (claset(),simpset() addsimps [real_add_commute]));
+qed "lemma_le_swap2";
+
+Goal "!!x. [| 0r < x + %~X + 1r; x < xa |] ==> 0r < xa + %~X + 1r";
+by (dtac real_add_less_mono 1);
+by (assume_tac 1);
+by (dres_inst_tac [("C","%~x"),("A","0r + x")] real_add_less_mono2 1);
+by (asm_full_simp_tac (simpset() addsimps [real_add_zero_right,
+    real_add_assoc RS sym,real_add_minus_left,real_add_zero_left]) 1);
+by (asm_full_simp_tac (simpset() addsimps real_add_ac) 1);
+qed "lemma_real_complete1";
+
+Goal "!!x. [| x + %~X + 1r <= S; xa < x |] ==> xa + %~X + 1r <= S";
+by (dtac real_less_imp_le 1);
+by (dtac real_add_le_mono 1);
+by (assume_tac 1);
+by (asm_full_simp_tac (simpset() addsimps real_add_ac) 1);
+by (dres_inst_tac [("x","%~x"),("q2.0","x + S")] real_add_left_le_mono1 1);
+by (asm_full_simp_tac (simpset() addsimps [real_add_assoc RS sym,
+        real_add_minus_left,real_add_zero_left]) 1);
+qed "lemma_real_complete2";
+
+Goal "!!x. [| x + %~X + 1r <= S; xa < x |] ==> xa <= S + X + %~1r"; (**)
+by (rtac (lemma_le_swap2 RS iffD2) 1);
+by (etac lemma_real_complete2 1);
+by (assume_tac 1);
+qed "lemma_real_complete2a";
+
+Goal "!!x. [| x + %~X + 1r <= S; xa <= x |] ==> xa <= S + X + %~1r";
+by (rotate_tac 1 1);
+by (etac (real_le_imp_less_or_eq RS disjE) 1);
+by (blast_tac (claset() addIs [lemma_real_complete2a]) 1);
+by (blast_tac (claset() addIs [(lemma_le_swap2 RS iffD2)]) 1);
+qed "lemma_real_complete2b";
+
+(*------------------------------------
+      reals Completeness (again!)
+ ------------------------------------*)
+Goal "!!(S::real set). [| EX X. X: S; \
+\                             EX Y. isUb (UNIV::real set) S Y \
+\                          |] ==> EX t. isLub (UNIV :: real set) S t";
+by (Step_tac 1);
+by (subgoal_tac "? u. u: {z. ? x: S. z = x + %~X + 1r} \
+\                Int {x. 0r < x}" 1);
+by (subgoal_tac "isUb (UNIV::real set) ({z. ? x: S. z = x + %~X + 1r} \
+\                Int {x. 0r < x})  (Y + %~X + 1r)" 1); 
+by (cut_inst_tac [("P","S"),("xa","X")] real_sup_lemma3 1);
+by (EVERY1[forward_tac [exI RSN (3,posreals_complete)], Blast_tac, Blast_tac, Step_tac]);
+by (res_inst_tac [("x","t + X + %~1r")] exI 1);
+by (rtac isLubI2 1);
+by (rtac setgeI 2 THEN Step_tac 2);
+by (subgoal_tac "isUb (UNIV:: real set) ({z. ? x: S. z = x + %~X + 1r} \
+\                Int {x. 0r < x})  (y + %~X + 1r)" 2); 
+by (dres_inst_tac [("y","(y + %~ X + 1r)")] isLub_le_isUb 2 
+      THEN assume_tac 2);
+by (etac (lemma_le_swap RS subst) 2);
+by (rtac (setleI RS isUbI) 1);
+by (Step_tac 1);
+by (res_inst_tac [("R1.0","x"),("R2.0","y")] real_linear_less2 1);
+by (stac lemma_le_swap2 1);
+by (forward_tac [isLubD2] 1 THEN assume_tac 2);
+by (Step_tac 1);
+by (Blast_tac 1);
+by (dtac lemma_real_complete1 1 THEN REPEAT(assume_tac 1));
+by (stac lemma_le_swap2 1);
+by (forward_tac [isLubD2] 1 THEN assume_tac 2);
+by (Blast_tac 1);
+by (rtac lemma_real_complete2b 1);
+by (etac real_less_imp_le 2);
+by (blast_tac (claset() addSIs [isLubD2]) 1 THEN Step_tac 1);
+by (blast_tac (claset() addDs [isUbD] addSIs [(setleI RS isUbI)]
+    addIs [real_add_le_mono1,real_add_assoc RS ssubst]) 1);
+by (blast_tac (claset() addDs [isUbD] addSIs [(setleI RS isUbI)]
+    addIs [real_add_le_mono1,real_add_assoc RS ssubst]) 1);
+by (auto_tac (claset(),simpset() addsimps [real_add_assoc RS sym,
+     real_add_minus,real_add_zero_left,real_zero_less_one]));
+qed "reals_complete";
+
+(*----------------------------------------------------------------
+        Related property: Archimedean property of reals
+ ----------------------------------------------------------------*)
+
+Goal "(ALL m. x*%%#m + x <= t) = (ALL m. x*%%#m <= t + %~x)";
+by Auto_tac;
+by (ALLGOALS(dres_inst_tac [("x","m")] spec));
+by (dres_inst_tac [("x","%~x")] real_add_le_mono1 1);
+by (dres_inst_tac [("x","x")] real_add_le_mono1 2);
+by (auto_tac (claset(),simpset() addsimps [real_add_assoc,
+      real_add_minus,real_add_minus_left,real_add_zero_right]));
+qed "lemma_arch";
+
+Goal "!!x. 0r < x ==> EX n. rinv(%%#n) < x";
+by (stac real_nat_rinv_Ex_iff 1);
+by (EVERY1[rtac ccontr, Asm_full_simp_tac]);
+by (fold_tac [real_le_def]);
+by (subgoal_tac "isUb (UNIV::real set) {z. EX n. z = x*%%#n} 1r" 1);
+by (subgoal_tac "EX X. X : {z. EX n. z = x*%%#n}" 1);
+by (dtac reals_complete 1);
+by (auto_tac (claset() addIs [isUbI,setleI],simpset()));
+by (subgoal_tac "ALL m. x*(%%#Suc m) <= t" 1);
+by (asm_full_simp_tac (simpset() addsimps 
+   [real_nat_Suc,real_add_mult_distrib2]) 1);
+by (blast_tac (claset() addIs [isLubD2]) 2);
+by (asm_full_simp_tac (simpset() addsimps [lemma_arch]) 1);
+by (subgoal_tac "isUb (UNIV::real set) {z. EX n. z = x*%%#n} (t + %~x)" 1);
+by (blast_tac (claset() addSIs [isUbI,setleI]) 2);
+by (dres_inst_tac [("y","t+%~x")] isLub_le_isUb 1);
+by (dres_inst_tac [("x","%~t")] real_add_left_le_mono1 2);
+by (auto_tac (claset() addDs [real_le_less_trans,
+    (real_minus_zero_less_iff2 RS iffD2)], simpset() 
+    addsimps [real_less_not_refl,real_add_assoc RS sym,
+    real_add_minus_left,real_add_zero_left]));
+qed "reals_Archimedean";
+
+Goal "EX n. (x::real) < %%#n";
+by (res_inst_tac [("R1.0","x"),("R2.0","0r")] real_linear_less2 1);
+by (res_inst_tac [("x","0")] exI 1);
+by (res_inst_tac [("x","0")] exI 2);
+by (auto_tac (claset() addEs [real_less_trans],
+    simpset() addsimps [real_nat_one,real_zero_less_one]));
+by (forward_tac [(real_rinv_gt_zero RS reals_Archimedean)] 1);
+by (Step_tac 1 THEN res_inst_tac [("x","n")] exI 1);
+by (forw_inst_tac [("y","rinv x")] real_mult_less_mono1 1);
+by (auto_tac (claset(),simpset() addsimps [real_not_refl2 RS not_sym]));
+by (dres_inst_tac [("n1","n"),("y","1r")] 
+     (real_nat_less_zero RS real_mult_less_mono2)  1);
+by (auto_tac (claset(),simpset() addsimps [real_nat_less_zero,
+    real_not_refl2 RS not_sym,real_mult_assoc RS sym]));
+qed "reals_Archimedean2";
+
+
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/RComplete.thy	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,9 @@
+(*  Title       : RComplete.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Completeness theorems for positive
+                  reals and reals 
+*) 
+
+RComplete = Lubs + Real
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/README.html	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,29 @@
+<!-- $Id$ -->
+<HTML><HEAD><TITLE>HOL/Real/README</TITLE></HEAD><BODY>
+
+<H2>Real--Dedekind Cut Construction of the Real Line</H2>
+
+<P>Requires <A HREF="../Integ/Equiv.thy">Equiv.thy</A> in the subdirectory <A
+HREF="../Integ">HOL/Integ</A>.
+
+<UL>
+<LI><A HREF="PNat.thy">PNat</A>  The positive integers (very much the same as <A HREF="../Nat.thy">Nat.thy</A>!) 
+<LI><A HREF="PRat.thy">PRat</A>  The positive rationals
+<LI><A HREF="PReal.thy">PReal</A> The positive reals constructed using Dedekind cuts
+<LI><A HREF="Real.thy">Real</A>  The real numbers
+<LI><A HREF="Lubs.thy">Lubs</A>  Definition of upper bounds, lubs and so on. 
+     (Useful e.g. in Fleuriot's NSA theory)
+<LI><A HREF="RComplete.thy">RComplete</A> Proof of completeness of reals in form of the supremum 
+            property. Also proofs that the reals have the Archimedean
+            property.
+<LI><A HREF="RealAbs.thy">RealAbs</A> The absolute value function defined for the reals
+</UL>
+
+<P>Last modified on $Date$
+
+<HR>
+
+<ADDRESS>
+<A NAME="lcp@cl.cam.ac.uk" HREF="mailto:lcp@cl.cam.ac.uk">lcp@cl.cam.ac.uk</A>
+</ADDRESS>
+</BODY></HTML>
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/ROOT.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,15 @@
+(*  Title:      HOL/Real/ROOT
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1998  University of Cambridge
+
+Construction of the Reals using Dedekind Cuts, by Jacques Fleuriot
+*)
+
+HOL_build_completed;    (*Make examples fail if HOL did*)
+
+writeln"Root file for HOL/Real";
+
+set proof_timing;
+time_use_thy "RealAbs";
+time_use_thy "RComplete";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/Real.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,1464 @@
+(*  Title       : Real.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : The reals
+*)
+
+open Real;
+
+(*** Proving that realrel is an equivalence relation ***)
+
+Goal
+    "!! x1. [| (x1::preal) + y2 = x2 + y1; x2 + y3 = x3 + y2 |] \
+\            ==> x1 + y3 = x3 + y1";        
+by (res_inst_tac [("C","y2")] preal_add_right_cancel 1);
+by (rotate_tac 1 1 THEN dtac sym 1);
+by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
+by (rtac (preal_add_left_commute RS subst) 1);
+by (res_inst_tac [("x1","x1")] (preal_add_assoc RS subst) 1);
+by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
+qed "preal_trans_lemma";
+
+(** Natural deduction for realrel **)
+
+Goalw [realrel_def]
+    "(((x1,y1),(x2,y2)): realrel) = (x1 + y2 = x2 + y1)";
+by (Fast_tac 1);
+qed "realrel_iff";
+
+Goalw [realrel_def]
+    "!!x1 x2. [| x1 + y2 = x2 + y1 |] ==> ((x1,y1),(x2,y2)): realrel";
+by (Fast_tac  1);
+qed "realrelI";
+
+Goalw [realrel_def]
+  "p: realrel --> (EX x1 y1 x2 y2. \
+\                  p = ((x1,y1),(x2,y2)) & x1 + y2 = x2 + y1)";
+by (Fast_tac 1);
+qed "realrelE_lemma";
+
+val [major,minor] = goal thy
+  "[| p: realrel;  \
+\     !!x1 y1 x2 y2. [| p = ((x1,y1),(x2,y2));  x1+y2 = x2+y1 \
+\                    |] ==> Q |] ==> Q";
+by (cut_facts_tac [major RS (realrelE_lemma RS mp)] 1);
+by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
+qed "realrelE";
+
+AddSIs [realrelI];
+AddSEs [realrelE];
+
+Goal "(x,x): realrel";
+by (stac surjective_pairing 1 THEN rtac (refl RS realrelI) 1);
+qed "realrel_refl";
+
+Goalw [equiv_def, refl_def, sym_def, trans_def]
+    "equiv {x::(preal*preal).True} realrel";
+by (fast_tac (claset() addSIs [realrel_refl] 
+                      addSEs [sym,preal_trans_lemma]) 1);
+qed "equiv_realrel";
+
+val equiv_realrel_iff =
+    [TrueI, TrueI] MRS 
+    ([CollectI, CollectI] MRS 
+    (equiv_realrel RS eq_equiv_class_iff));
+
+Goalw  [real_def,realrel_def,quotient_def] "realrel^^{(x,y)}:real";
+by (Blast_tac 1);
+qed "realrel_in_real";
+
+Goal "inj_on Abs_real real";
+by (rtac inj_on_inverseI 1);
+by (etac Abs_real_inverse 1);
+qed "inj_on_Abs_real";
+
+Addsimps [equiv_realrel_iff,inj_on_Abs_real RS inj_on_iff,
+          realrel_iff, realrel_in_real, Abs_real_inverse];
+
+Addsimps [equiv_realrel RS eq_equiv_class_iff];
+val eq_realrelD = equiv_realrel RSN (2,eq_equiv_class);
+
+Goal "inj(Rep_real)";
+by (rtac inj_inverseI 1);
+by (rtac Rep_real_inverse 1);
+qed "inj_Rep_real";
+
+(** real_preal: the injection from preal to real **)
+Goal "inj(real_preal)";
+by (rtac injI 1);
+by (rewtac real_preal_def);
+by (dtac (inj_on_Abs_real RS inj_onD) 1);
+by (REPEAT (rtac realrel_in_real 1));
+by (dtac eq_equiv_class 1);
+by (rtac equiv_realrel 1);
+by (Fast_tac 1);
+by Safe_tac;
+by (Asm_full_simp_tac 1);
+qed "inj_real_preal";
+
+val [prem] = goal thy
+    "(!!x y. z = Abs_real(realrel^^{(x,y)}) ==> P) ==> P";
+by (res_inst_tac [("x1","z")] 
+    (rewrite_rule [real_def] Rep_real RS quotientE) 1);
+by (dres_inst_tac [("f","Abs_real")] arg_cong 1);
+by (res_inst_tac [("p","x")] PairE 1);
+by (rtac prem 1);
+by (asm_full_simp_tac (simpset() addsimps [Rep_real_inverse]) 1);
+qed "eq_Abs_real";
+
+(**** real_minus: additive inverse on real ****)
+
+Goalw [congruent_def]
+  "congruent realrel (%p. split (%x y. realrel^^{(y,x)}) p)";
+by Safe_tac;
+by (asm_full_simp_tac (simpset() addsimps [preal_add_commute]) 1);
+qed "real_minus_congruent";
+
+(*Resolve th against the corresponding facts for real_minus*)
+val real_minus_ize = RSLIST [equiv_realrel, real_minus_congruent];
+
+Goalw [real_minus_def]
+      "%~ (Abs_real(realrel^^{(x,y)})) = Abs_real(realrel ^^ {(y,x)})";
+by (res_inst_tac [("f","Abs_real")] arg_cong 1);
+by (simp_tac (simpset() addsimps 
+   [realrel_in_real RS Abs_real_inverse,real_minus_ize UN_equiv_class]) 1);
+qed "real_minus";
+
+Goal "%~ (%~ z) = z";
+by (res_inst_tac [("z","z")] eq_Abs_real 1);
+by (asm_simp_tac (simpset() addsimps [real_minus]) 1);
+qed "real_minus_minus";
+
+Addsimps [real_minus_minus];
+
+Goal "inj(real_minus)";
+by (rtac injI 1);
+by (dres_inst_tac [("f","real_minus")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [real_minus_minus]) 1);
+qed "inj_real_minus";
+
+Goalw [real_zero_def] "%~0r = 0r";
+by (simp_tac (simpset() addsimps [real_minus]) 1);
+qed "real_minus_zero";
+
+Addsimps [real_minus_zero];
+
+Goal "(%~x = 0r) = (x = 0r)"; 
+by (res_inst_tac [("z","x")] eq_Abs_real 1);
+by (auto_tac (claset(),simpset() addsimps [real_zero_def,
+    real_minus] @ preal_add_ac));
+qed "real_minus_zero_iff";
+
+Addsimps [real_minus_zero_iff];
+
+Goal "(%~x ~= 0r) = (x ~= 0r)"; 
+by Auto_tac;
+qed "real_minus_not_zero_iff";
+
+(*** Congruence property for addition ***)
+Goalw [congruent2_def]
+    "congruent2 realrel (%p1 p2.                  \
+\         split (%x1 y1. split (%x2 y2. realrel^^{(x1+x2, y1+y2)}) p2) p1)";
+by Safe_tac;
+by (asm_simp_tac (simpset() addsimps [preal_add_assoc]) 1);
+by (res_inst_tac [("z1.1","x1a")] (preal_add_left_commute RS ssubst) 1);
+by (asm_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1);
+by (asm_simp_tac (simpset() addsimps preal_add_ac) 1);
+qed "real_add_congruent2";
+
+(*Resolve th against the corresponding facts for real_add*)
+val real_add_ize = RSLIST [equiv_realrel, real_add_congruent2];
+
+Goalw [real_add_def]
+  "Abs_real(realrel^^{(x1,y1)}) + Abs_real(realrel^^{(x2,y2)}) = \
+\  Abs_real(realrel^^{(x1+x2, y1+y2)})";
+by (asm_simp_tac
+    (simpset() addsimps [real_add_ize UN_equiv_class2]) 1);
+qed "real_add";
+
+Goal "(z::real) + w = w + z";
+by (res_inst_tac [("z","z")] eq_Abs_real 1);
+by (res_inst_tac [("z","w")] eq_Abs_real 1);
+by (asm_simp_tac (simpset() addsimps (preal_add_ac @ [real_add])) 1);
+qed "real_add_commute";
+
+Goal "((z1::real) + z2) + z3 = z1 + (z2 + z3)";
+by (res_inst_tac [("z","z1")] eq_Abs_real 1);
+by (res_inst_tac [("z","z2")] eq_Abs_real 1);
+by (res_inst_tac [("z","z3")] eq_Abs_real 1);
+by (asm_simp_tac (simpset() addsimps [real_add, preal_add_assoc]) 1);
+qed "real_add_assoc";
+
+(*For AC rewriting*)
+Goal "(x::real)+(y+z)=y+(x+z)";
+by (rtac (real_add_commute RS trans) 1);
+by (rtac (real_add_assoc RS trans) 1);
+by (rtac (real_add_commute RS arg_cong) 1);
+qed "real_add_left_commute";
+
+(* real addition is an AC operator *)
+val real_add_ac = [real_add_assoc,real_add_commute,real_add_left_commute];
+
+Goalw [real_preal_def,real_zero_def] "0r + z = z";
+by (res_inst_tac [("z","z")] eq_Abs_real 1);
+by (asm_full_simp_tac (simpset() addsimps [real_add] @ preal_add_ac) 1);
+qed "real_add_zero_left";
+
+Goal "z + 0r = z";
+by (simp_tac (simpset() addsimps [real_add_zero_left,real_add_commute]) 1);
+qed "real_add_zero_right";
+
+Goalw [real_zero_def] "z + %~z = 0r";
+by (res_inst_tac [("z","z")] eq_Abs_real 1);
+by (asm_full_simp_tac (simpset() addsimps [real_minus,
+        real_add, preal_add_commute]) 1);
+qed "real_add_minus";
+
+Goal "%~z + z = 0r";
+by (simp_tac (simpset() addsimps 
+    [real_add_commute,real_add_minus]) 1);
+qed "real_add_minus_left";
+
+Goal "? y. (x::real) + y = 0r";
+by (fast_tac (claset() addIs [real_add_minus]) 1);
+qed "real_minus_ex";
+
+Goal "?! y. (x::real) + y = 0r";
+by (auto_tac (claset() addIs [real_add_minus],simpset()));
+by (dres_inst_tac [("f","%x. ya+x")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1);
+by (asm_full_simp_tac (simpset() addsimps [real_add_commute,
+    real_add_zero_right,real_add_zero_left]) 1);
+qed "real_minus_ex1";
+
+Goal "?! y. y + (x::real) = 0r";
+by (auto_tac (claset() addIs [real_add_minus_left],simpset()));
+by (dres_inst_tac [("f","%x. x+ya")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [real_add_assoc]) 1);
+by (asm_full_simp_tac (simpset() addsimps [real_add_commute,
+    real_add_zero_right,real_add_zero_left]) 1);
+qed "real_minus_left_ex1";
+
+Goal "!!y. x + y = 0r ==> x = %~y";
+by (cut_inst_tac [("z","y")] real_add_minus_left 1);
+by (res_inst_tac [("x1","y")] (real_minus_left_ex1 RS ex1E) 1);
+by (Blast_tac 1);
+qed "real_add_minus_eq_minus";
+
+Goal "? y. x = %~y";
+by (cut_inst_tac [("x","x")] real_minus_ex 1);
+by (etac exE 1 THEN dtac real_add_minus_eq_minus 1);
+by (Fast_tac 1);
+qed "real_as_add_inverse_ex";
+
+(* real_minus_add_distrib *)
+Goal "%~(x + y) = %~x + %~y";
+by (res_inst_tac [("z","x")] eq_Abs_real 1);
+by (res_inst_tac [("z","y")] eq_Abs_real 1);
+by (auto_tac (claset(),simpset() addsimps [real_minus,real_add]));
+qed "real_minus_add_eq";
+
+val real_minus_add_distrib = real_minus_add_eq;
+
+Goal "((x::real) + y = x + z) = (y = z)";
+by (Step_tac 1);
+by (dres_inst_tac [("f","%t.%~x + t")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [real_add_minus_left,
+                 real_add_assoc RS sym,real_add_zero_left]) 1);
+qed "real_add_left_cancel";
+
+Goal "(y + (x::real)= z + x) = (y = z)";
+by (simp_tac (simpset() addsimps [real_add_commute,real_add_left_cancel]) 1);
+qed "real_add_right_cancel";
+
+(*** Congruence property for multiplication ***)
+Goal "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==> \
+\         x * x1 + y * y1 + (x * y2 + x2 * y) = \
+\         x * x2 + y * y2 + (x * y1 + x1 * y)";
+by (asm_full_simp_tac (simpset() addsimps [preal_add_left_commute,
+    preal_add_assoc RS sym,preal_add_mult_distrib2 RS sym]) 1);
+by (rtac (preal_mult_commute RS subst) 1);
+by (res_inst_tac [("y1","x2")] (preal_mult_commute RS subst) 1);
+by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc,
+    preal_add_mult_distrib2 RS sym]) 1);
+by (asm_full_simp_tac (simpset() addsimps [preal_add_commute]) 1);
+qed "real_mult_congruent2_lemma";
+
+Goal 
+    "congruent2 realrel (%p1 p2.                  \
+\         split (%x1 y1. split (%x2 y2. realrel^^{(x1*x2 + y1*y2, x1*y2+x2*y1)}) p2) p1)";
+by (rtac (equiv_realrel RS congruent2_commuteI) 1);
+by Safe_tac;
+by (rewtac split_def);
+by (asm_simp_tac (simpset() addsimps [preal_mult_commute,preal_add_commute]) 1);
+by (auto_tac (claset(),simpset() addsimps [real_mult_congruent2_lemma]));
+qed "real_mult_congruent2";
+
+(*Resolve th against the corresponding facts for real_mult*)
+val real_mult_ize = RSLIST [equiv_realrel, real_mult_congruent2];
+
+Goalw [real_mult_def]
+   "Abs_real((realrel^^{(x1,y1)})) * Abs_real((realrel^^{(x2,y2)})) =   \
+\   Abs_real(realrel ^^ {(x1*x2+y1*y2,x1*y2+x2*y1)})";
+by (simp_tac (simpset() addsimps [real_mult_ize UN_equiv_class2]) 1);
+qed "real_mult";
+
+Goal "(z::real) * w = w * z";
+by (res_inst_tac [("z","z")] eq_Abs_real 1);
+by (res_inst_tac [("z","w")] eq_Abs_real 1);
+by (asm_simp_tac (simpset() addsimps ([real_mult] @ preal_add_ac @ preal_mult_ac)) 1);
+qed "real_mult_commute";
+
+Goal "((z1::real) * z2) * z3 = z1 * (z2 * z3)";
+by (res_inst_tac [("z","z1")] eq_Abs_real 1);
+by (res_inst_tac [("z","z2")] eq_Abs_real 1);
+by (res_inst_tac [("z","z3")] eq_Abs_real 1);
+by (asm_simp_tac (simpset() addsimps ([preal_add_mult_distrib2,real_mult] @ 
+                                     preal_add_ac @ preal_mult_ac)) 1);
+qed "real_mult_assoc";
+
+qed_goal "real_mult_left_commute" thy
+    "(z1::real) * (z2 * z3) = z2 * (z1 * z3)"
+ (fn _ => [rtac (real_mult_commute RS trans) 1, rtac (real_mult_assoc RS trans) 1,
+           rtac (real_mult_commute RS arg_cong) 1]);
+
+(* real multiplication is an AC operator *)
+val real_mult_ac = [real_mult_assoc, real_mult_commute, real_mult_left_commute];
+
+Goalw [real_one_def,pnat_one_def] "1r * z = z";
+by (res_inst_tac [("z","z")] eq_Abs_real 1);
+by (asm_full_simp_tac (simpset() addsimps [real_mult,
+    preal_add_mult_distrib2,preal_mult_1_right] 
+    @ preal_mult_ac @ preal_add_ac) 1);
+qed "real_mult_1";
+
+Goal "z * 1r = z";
+by (simp_tac (simpset() addsimps [real_mult_commute,
+    real_mult_1]) 1);
+qed "real_mult_1_right";
+
+Goalw [real_zero_def,pnat_one_def] "0r * z = 0r";
+by (res_inst_tac [("z","z")] eq_Abs_real 1);
+by (asm_full_simp_tac (simpset() addsimps [real_mult,
+    preal_add_mult_distrib2,preal_mult_1_right] 
+    @ preal_mult_ac @ preal_add_ac) 1);
+qed "real_mult_0";
+
+Goal "z * 0r = 0r";
+by (simp_tac (simpset() addsimps [real_mult_commute,
+    real_mult_0]) 1);
+qed "real_mult_0_right";
+
+Addsimps [real_mult_0_right,real_mult_0];
+
+Goal "%~(x * y) = %~x * y";
+by (res_inst_tac [("z","x")] eq_Abs_real 1);
+by (res_inst_tac [("z","y")] eq_Abs_real 1);
+by (auto_tac (claset(),simpset() addsimps [real_minus,real_mult] 
+    @ preal_mult_ac @ preal_add_ac));
+qed "real_minus_mult_eq1";
+
+Goal "%~(x * y) = x * %~y";
+by (res_inst_tac [("z","x")] eq_Abs_real 1);
+by (res_inst_tac [("z","y")] eq_Abs_real 1);
+by (auto_tac (claset(),simpset() addsimps [real_minus,real_mult] 
+    @ preal_mult_ac @ preal_add_ac));
+qed "real_minus_mult_eq2";
+
+Goal "%~x*%~y = x*y";
+by (full_simp_tac (simpset() addsimps [real_minus_mult_eq2 RS sym,
+    real_minus_mult_eq1 RS sym]) 1);
+qed "real_minus_mult_cancel";
+
+Addsimps [real_minus_mult_cancel];
+
+Goal "%~x*y = x*%~y";
+by (full_simp_tac (simpset() addsimps [real_minus_mult_eq2 RS sym,
+    real_minus_mult_eq1 RS sym]) 1);
+qed "real_minus_mult_commute";
+
+(*-----------------------------------------------------------------------------
+
+ -----------------------------------------------------------------------------*)
+
+(** Lemmas **)
+
+qed_goal "real_add_assoc_cong" thy
+    "!!z. (z::real) + v = z' + v' ==> z + (v + w) = z' + (v' + w)"
+ (fn _ => [(asm_simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1)]);
+
+qed_goal "real_add_assoc_swap" thy "(z::real) + (v + w) = v + (z + w)"
+ (fn _ => [(REPEAT (ares_tac [real_add_commute RS real_add_assoc_cong] 1))]);
+
+Goal "((z1::real) + z2) * w = (z1 * w) + (z2 * w)";
+by (res_inst_tac [("z","z1")] eq_Abs_real 1);
+by (res_inst_tac [("z","z2")] eq_Abs_real 1);
+by (res_inst_tac [("z","w")] eq_Abs_real 1);
+by (asm_simp_tac 
+    (simpset() addsimps ([preal_add_mult_distrib2, real_add, real_mult] @ 
+                        preal_add_ac @ preal_mult_ac)) 1);
+qed "real_add_mult_distrib";
+
+val real_mult_commute'= read_instantiate [("z","w")] real_mult_commute;
+
+Goal "(w::real) * (z1 + z2) = (w * z1) + (w * z2)";
+by (simp_tac (simpset() addsimps [real_mult_commute',real_add_mult_distrib]) 1);
+qed "real_add_mult_distrib2";
+
+val real_mult_simps = [real_mult_1, real_mult_1_right];
+Addsimps real_mult_simps;
+
+(*** one and zero are distinct ***)
+Goalw [real_zero_def,real_one_def] "0r ~= 1r";
+by (auto_tac (claset(),simpset() addsimps 
+   [preal_self_less_add_left RS preal_not_refl2]));
+qed "real_zero_not_eq_one";
+
+(*** existence of inverse ***)
+(** lemma -- alternative definition for 0r **)
+Goalw [real_zero_def] "0r = Abs_real (realrel ^^ {(x, x)})";
+by (auto_tac (claset(),simpset() addsimps [preal_add_commute]));
+qed "real_zero_iff";
+
+Goalw [real_zero_def,real_one_def] 
+          "!!(x::real). x ~= 0r ==> ? y. x*y = 1r";
+by (res_inst_tac [("z","x")] eq_Abs_real 1);
+by (cut_inst_tac [("r1.0","xa"),("r2.0","y")] preal_linear 1);
+by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
+           simpset() addsimps [real_zero_iff RS sym]));
+by (res_inst_tac [("x","Abs_real (realrel ^^ {(@#$#1p,pinv(D)+@#$#1p)})")] exI 1);
+by (res_inst_tac [("x","Abs_real (realrel ^^ {(pinv(D)+@#$#1p,@#$#1p)})")] exI 2);
+by (auto_tac (claset(),simpset() addsimps [real_mult,
+    pnat_one_def,preal_mult_1_right,preal_add_mult_distrib2,
+    preal_add_mult_distrib,preal_mult_1,preal_mult_inv_right] 
+    @ preal_add_ac @ preal_mult_ac));
+qed "real_mult_inv_right_ex";
+
+Goal "!!(x::real). x ~= 0r ==> ? y. y*x = 1r";
+by (asm_simp_tac (simpset() addsimps [real_mult_commute,
+    real_mult_inv_right_ex]) 1);
+qed "real_mult_inv_left_ex";
+
+Goalw [rinv_def] "!!(x::real). x ~= 0r ==> rinv(x)*x = 1r";
+by (forward_tac [real_mult_inv_left_ex] 1);
+by (Step_tac 1);
+by (rtac selectI2 1);
+by Auto_tac;
+qed "real_mult_inv_left";
+
+Goal "!!(x::real). x ~= 0r ==> x*rinv(x) = 1r";
+by (auto_tac (claset() addIs [real_mult_commute RS subst],
+              simpset() addsimps [real_mult_inv_left]));
+qed "real_mult_inv_right";
+
+Goal "!!a. (c::real) ~= 0r ==> (c*a=c*b) = (a=b)";
+by Auto_tac;
+by (dres_inst_tac [("f","%x. x*rinv c")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [real_mult_inv_right] @ real_mult_ac)  1);
+qed "real_mult_left_cancel";
+    
+Goal "!!a. (c::real) ~= 0r ==> (a*c=b*c) = (a=b)";
+by (Step_tac 1);
+by (dres_inst_tac [("f","%x. x*rinv c")] arg_cong 1);
+by (asm_full_simp_tac (simpset() addsimps [real_mult_inv_right] @ real_mult_ac)  1);
+qed "real_mult_right_cancel";
+
+Goalw [rinv_def] "!!x. x ~= 0r ==> rinv(x) ~= 0r";
+by (forward_tac [real_mult_inv_left_ex] 1);
+by (etac exE 1);
+by (rtac selectI2 1);
+by (auto_tac (claset(),simpset() addsimps [real_mult_0,
+    real_zero_not_eq_one]));
+qed "rinv_not_zero";
+
+Addsimps [real_mult_inv_left,real_mult_inv_right];
+
+Goal "!!x. x ~= 0r ==> rinv(rinv x) = x";
+by (res_inst_tac [("c1","rinv x")] (real_mult_right_cancel RS iffD1) 1);
+by (etac rinv_not_zero 1);
+by (auto_tac (claset() addDs [rinv_not_zero],simpset()));
+qed "real_rinv_rinv";
+
+Goalw [rinv_def] "rinv(1r) = 1r";
+by (cut_facts_tac [real_zero_not_eq_one RS 
+       not_sym RS real_mult_inv_left_ex] 1);
+by (etac exE 1);
+by (rtac selectI2 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [real_zero_not_eq_one RS not_sym]));
+qed "real_rinv_1";
+
+Goal "!!x. x ~= 0r ==> rinv(%~x) = %~rinv(x)";
+by (res_inst_tac [("c1","%~x")] (real_mult_right_cancel RS iffD1) 1);
+by Auto_tac;
+qed "real_minus_rinv";
+
+      (*** theorems for ordering ***)
+(* prove introduction and elimination rules for real_less *)
+
+Goalw [real_less_def]
+ "P < (Q::real) = (EX x1 y1 x2 y2. x1 + y2 < x2 + y1 & \
+\                                  (x1,y1::preal):Rep_real(P) & \
+\                                  (x2,y2):Rep_real(Q))";
+by (Fast_tac 1);
+qed "real_less_iff";
+
+Goalw [real_less_def]
+ "!!P. [| x1 + y2 < x2 + y1; (x1,y1::preal):Rep_real(P); \
+\         (x2,y2):Rep_real(Q) |] ==> P < (Q::real)";
+by (Fast_tac 1);
+qed "real_lessI";
+
+Goalw [real_less_def]
+     "!! R1. [| R1 < (R2::real); \
+\         !!x1 x2 y1 y2. x1 + y2 < x2 + y1 ==> P; \
+\         !!x1 y1. (x1,y1::preal):Rep_real(R1) ==> P; \ 
+\         !!x2 y2. (x2,y2::preal):Rep_real(R2) ==> P |] \
+\     ==> P";
+by Auto_tac;
+qed "real_lessE";
+
+Goalw [real_less_def]
+ "!!R1. R1 < (R2::real) ==> (EX x1 y1 x2 y2. x1 + y2 < x2 + y1 & \
+\                                  (x1,y1::preal):Rep_real(R1) & \
+\                                  (x2,y2):Rep_real(R2))";
+by (Fast_tac 1);
+qed "real_lessD";
+
+(* real_less is a strong order i.e nonreflexive and transitive *)
+(*** lemmas ***)
+Goal "!!(x::preal). [| x = y; x1 = y1 |] ==> x + y1 = x1 + y";
+by (asm_simp_tac (simpset() addsimps [preal_add_commute]) 1);
+qed "preal_lemma_eq_rev_sum";
+
+Goal "!!(b::preal). x + (b + y) = x1 + (b + y1) ==> x + y = x1 + y1";
+by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
+qed "preal_add_left_commute_cancel";
+
+Goal 
+     "!!(x::preal). [| x + y2a = x2a + y; \
+\                      x + y2b = x2b + y |] \
+\                   ==> x2a + y2b = x2b + y2a";
+by (dtac preal_lemma_eq_rev_sum 1);
+by (assume_tac 1);
+by (thin_tac "x + y2b = x2b + y" 1);
+by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
+by (dtac preal_add_left_commute_cancel 1);
+by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
+qed "preal_lemma_for_not_refl";
+
+Goal "~ (R::real) < R";
+by (res_inst_tac [("z","R")] eq_Abs_real 1);
+by (auto_tac (claset(),simpset() addsimps [real_less_def]));
+by (dtac preal_lemma_for_not_refl 1);
+by (assume_tac 1 THEN rotate_tac 2 1);
+by (auto_tac (claset(),simpset() addsimps [preal_less_not_refl]));
+qed "real_less_not_refl";
+
+(*** y < y ==> P ***)
+bind_thm("real_less_irrefl",real_less_not_refl RS notE);
+
+Goal "!!(x::real). x < y ==> x ~= y";
+by (auto_tac (claset(),simpset() addsimps [real_less_not_refl]));
+qed "real_not_refl2";
+
+(* lemma re-arranging and eliminating terms *)
+Goal "!! (a::preal). [| a + b = c + d; \
+\            x2b + d + (c + y2e) < a + y2b + (x2e + b) |] \
+\         ==> x2b + y2e < x2e + y2b";
+by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
+by (res_inst_tac [("C","c+d")] preal_add_left_less_cancel 1);
+by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1);
+qed "preal_lemma_trans";
+
+(** heavy re-writing involved*)
+Goal "!!(R1::real). [| R1 < R2; R2 < R3 |] ==> R1 < R3";
+by (res_inst_tac [("z","R1")] eq_Abs_real 1);
+by (res_inst_tac [("z","R2")] eq_Abs_real 1);
+by (res_inst_tac [("z","R3")] eq_Abs_real 1);
+by (auto_tac (claset(),simpset() addsimps [real_less_def]));
+by (REPEAT(rtac exI 1));
+by (EVERY[rtac conjI 1, rtac conjI 2]);
+by (REPEAT(Blast_tac 2));
+by (dtac preal_lemma_for_not_refl 1 THEN assume_tac 1);
+by (blast_tac (claset() addDs [preal_add_less_mono] 
+    addIs [preal_lemma_trans]) 1);
+qed "real_less_trans";
+
+Goal "!! (R1::real). [| R1 < R2; R2 < R1 |] ==> P";
+by (dtac real_less_trans 1 THEN assume_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [real_less_not_refl]) 1);
+qed "real_less_asym";
+
+(****)(****)(****)(****)(****)(****)(****)(****)(****)(****)
+    (****** Map and more real_less ******)
+(*** mapping from preal into real ***)
+Goalw [real_preal_def] 
+            "%#((z1::preal) + z2) = %#z1 + %#z2";
+by (asm_simp_tac (simpset() addsimps [real_add,
+       preal_add_mult_distrib,preal_mult_1] addsimps preal_add_ac) 1);
+qed "real_preal_add";
+
+Goalw [real_preal_def] 
+            "%#((z1::preal) * z2) = %#z1* %#z2";
+by (full_simp_tac (simpset() addsimps [real_mult,
+        preal_add_mult_distrib2,preal_mult_1,
+        preal_mult_1_right,pnat_one_def] 
+        @ preal_add_ac @ preal_mult_ac) 1);
+qed "real_preal_mult";
+
+Goalw [real_preal_def]
+      "!!(x::preal). y < x ==> ? m. Abs_real (realrel ^^ {(x,y)}) = %#m";
+by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
+    simpset() addsimps preal_add_ac));
+qed "real_preal_ExI";
+
+Goalw [real_preal_def]
+      "!!(x::preal). ? m. Abs_real (realrel ^^ {(x,y)}) = %#m ==> y < x";
+by (auto_tac (claset(),simpset() addsimps 
+    [preal_add_commute,preal_add_assoc]));
+by (asm_full_simp_tac (simpset() addsimps 
+    [preal_add_assoc RS sym,preal_self_less_add_left]) 1);
+qed "real_preal_ExD";
+
+Goal "(? m. Abs_real (realrel ^^ {(x,y)}) = %#m) = (y < x)";
+by (fast_tac (claset() addSIs [real_preal_ExI,real_preal_ExD]) 1);
+qed "real_preal_iff";
+
+(*** Gleason prop 9-4.4 p 127 ***)
+Goalw [real_preal_def,real_zero_def] 
+      "? m. (x::real) = %#m | x = 0r | x = %~(%#m)";
+by (res_inst_tac [("z","x")] eq_Abs_real 1);
+by (auto_tac (claset(),simpset() addsimps [real_minus] @ preal_add_ac));
+by (cut_inst_tac [("r1.0","x"),("r2.0","y")] preal_linear 1);
+by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
+    simpset() addsimps [preal_add_assoc RS sym]));
+by (auto_tac (claset(),simpset() addsimps [preal_add_commute]));
+qed "real_preal_trichotomy";
+
+Goal "!!x. [| !!m. x = %#m ==> P; \
+\            x = 0r ==> P; \
+\            !!m. x = %~(%#m) ==> P |] ==> P";
+by (cut_inst_tac [("x","x")] real_preal_trichotomy 1);
+by Auto_tac;
+qed "real_preal_trichotomyE";
+
+Goalw [real_preal_def] "!!m1 m2. %#m1 < %#m2 ==> m1 < m2";
+by (auto_tac (claset(),simpset() addsimps [real_less_def] @ preal_add_ac));
+by (auto_tac (claset(),simpset() addsimps [preal_add_assoc RS sym]));
+by (auto_tac (claset(),simpset() addsimps preal_add_ac));
+qed "real_preal_lessD";
+
+Goal "!!m1 m2. m1 < m2 ==> %#m1 < %#m2";
+by (dtac preal_less_add_left_Ex 1);
+by (auto_tac (claset(),simpset() addsimps [real_preal_add,
+    real_preal_def,real_less_def]));
+by (REPEAT(rtac exI 1));
+by (EVERY[rtac conjI 1, rtac conjI 2]);
+by (REPEAT(Fast_tac 2));
+by (simp_tac (simpset() addsimps [preal_self_less_add_left] 
+    delsimps [preal_add_less_iff2]) 1);
+qed "real_preal_lessI";
+
+Goal "(%#m1 < %#m2) = (m1 < m2)";
+by (fast_tac (claset() addIs [real_preal_lessI,real_preal_lessD]) 1);
+qed "real_preal_less_iff1";
+
+Addsimps [real_preal_less_iff1];
+
+Goal "%~ %#m < %#m";
+by (auto_tac (claset(),simpset() addsimps 
+    [real_preal_def,real_less_def,real_minus]));
+by (REPEAT(rtac exI 1));
+by (EVERY[rtac conjI 1, rtac conjI 2]);
+by (REPEAT(Fast_tac 2));
+by (full_simp_tac (simpset() addsimps preal_add_ac) 1);
+by (full_simp_tac (simpset() addsimps [preal_self_less_add_right,
+    preal_add_assoc RS sym]) 1);
+qed "real_preal_minus_less_self";
+
+Goalw [real_zero_def] "%~ %#m < 0r";
+by (auto_tac (claset(),simpset() addsimps 
+    [real_preal_def,real_less_def,real_minus]));
+by (REPEAT(rtac exI 1));
+by (EVERY[rtac conjI 1, rtac conjI 2]);
+by (REPEAT(Fast_tac 2));
+by (full_simp_tac (simpset() addsimps 
+  [preal_self_less_add_right] @ preal_add_ac) 1);
+qed "real_preal_minus_less_zero";
+
+Goal "~ 0r < %~ %#m";
+by (cut_facts_tac [real_preal_minus_less_zero] 1);
+by (fast_tac (claset() addDs [real_less_trans] 
+               addEs [real_less_irrefl]) 1);
+qed "real_preal_not_minus_gt_zero";
+
+Goalw [real_zero_def] " 0r < %#m";
+by (auto_tac (claset(),simpset() addsimps 
+    [real_preal_def,real_less_def,real_minus]));
+by (REPEAT(rtac exI 1));
+by (EVERY[rtac conjI 1, rtac conjI 2]);
+by (REPEAT(Fast_tac 2));
+by (full_simp_tac (simpset() addsimps 
+  [preal_self_less_add_right] @ preal_add_ac) 1);
+qed "real_preal_zero_less";
+
+Goal "~ %#m < 0r";
+by (cut_facts_tac [real_preal_zero_less] 1);
+by (fast_tac (claset() addDs [real_less_trans] 
+               addEs [real_less_irrefl]) 1);
+qed "real_preal_not_less_zero";
+
+Goal "0r < %~ %~ %#m";
+by (simp_tac (simpset() addsimps 
+    [real_preal_zero_less]) 1);
+qed "real_minus_minus_zero_less";
+
+(* another lemma *)
+Goalw [real_zero_def] " 0r < %#m + %#m1";
+by (auto_tac (claset(),simpset() addsimps 
+    [real_preal_def,real_less_def,real_add]));
+by (REPEAT(rtac exI 1));
+by (EVERY[rtac conjI 1, rtac conjI 2]);
+by (REPEAT(Fast_tac 2));
+by (full_simp_tac (simpset() addsimps preal_add_ac) 1);
+by (full_simp_tac (simpset() addsimps [preal_self_less_add_right,
+    preal_add_assoc RS sym]) 1);
+qed "real_preal_sum_zero_less";
+
+Goal "%~ %#m < %#m1";
+by (auto_tac (claset(),simpset() addsimps 
+    [real_preal_def,real_less_def,real_minus]));
+by (REPEAT(rtac exI 1));
+by (EVERY[rtac conjI 1, rtac conjI 2]);
+by (REPEAT(Fast_tac 2));
+by (full_simp_tac (simpset() addsimps preal_add_ac) 1);
+by (full_simp_tac (simpset() addsimps [preal_self_less_add_right,
+    preal_add_assoc RS sym]) 1);
+qed "real_preal_minus_less_all";
+
+Goal "~ %#m < %~ %#m1";
+by (cut_facts_tac [real_preal_minus_less_all] 1);
+by (fast_tac (claset() addDs [real_less_trans] 
+               addEs [real_less_irrefl]) 1);
+qed "real_preal_not_minus_gt_all";
+
+Goal "!!m1 m2. %~ %#m1 < %~ %#m2 ==> %#m2 < %#m1";
+by (auto_tac (claset(),simpset() addsimps 
+    [real_preal_def,real_less_def,real_minus]));
+by (REPEAT(rtac exI 1));
+by (EVERY[rtac conjI 1, rtac conjI 2]);
+by (REPEAT(Fast_tac 2));
+by (auto_tac (claset(),simpset() addsimps preal_add_ac));
+by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1);
+by (auto_tac (claset(),simpset() addsimps preal_add_ac));
+qed "real_preal_minus_less_rev1";
+
+Goal "!!m1 m2. %#m1 < %#m2 ==> %~ %#m2 < %~ %#m1";
+by (auto_tac (claset(),simpset() addsimps 
+    [real_preal_def,real_less_def,real_minus]));
+by (REPEAT(rtac exI 1));
+by (EVERY[rtac conjI 1, rtac conjI 2]);
+by (REPEAT(Fast_tac 2));
+by (auto_tac (claset(),simpset() addsimps preal_add_ac));
+by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1);
+by (auto_tac (claset(),simpset() addsimps preal_add_ac));
+qed "real_preal_minus_less_rev2";
+
+Goal "(%~ %#m1 < %~ %#m2) = (%#m2 < %#m1)";
+by (blast_tac (claset() addSIs [real_preal_minus_less_rev1,
+               real_preal_minus_less_rev2]) 1);
+qed "real_preal_minus_less_rev_iff";
+
+Addsimps [real_preal_minus_less_rev_iff];
+
+(*** linearity ***)
+Goal "(R1::real) < R2 | R1 = R2 | R2 < R1";
+by (res_inst_tac [("x","R1")]  real_preal_trichotomyE 1);
+by (ALLGOALS(res_inst_tac [("x","R2")]  real_preal_trichotomyE));
+by (auto_tac (claset() addSDs [preal_le_anti_sym],
+              simpset() addsimps [preal_less_le_iff,real_preal_minus_less_zero,
+               real_preal_zero_less,real_preal_minus_less_all]));
+qed "real_linear";
+
+Goal
+    "!!(R1::real). [| R1 < R2 ==> P;  R1 = R2 ==> P; \
+\          R2 < R1 ==> P |] ==> P";
+by (cut_inst_tac [("R1.0","R1"),("R2.0","R2")] real_linear 1);
+by Auto_tac;
+qed "real_linear_less2";
+
+(*** Properties of <= ***)
+
+Goalw [real_le_def] "!!w. ~(w < z) ==> z <= (w::real)";
+by (assume_tac 1);
+qed "real_leI";
+
+Goalw [real_le_def] "!!w. z<=w ==> ~(w<(z::real))";
+by (assume_tac 1);
+qed "real_leD";
+
+val real_leE = make_elim real_leD;
+
+Goal "!!w. (~(w < z)) = (z <= (w::real))";
+by (fast_tac (claset() addSIs [real_leI,real_leD]) 1);
+qed "real_less_le_iff";
+
+Goalw [real_le_def] "!!z. ~ z <= w ==> w<(z::real)";
+by (Fast_tac 1);
+qed "not_real_leE";
+
+Goalw [real_le_def] "!!z. z < w ==> z <= (w::real)";
+by (fast_tac (claset() addEs [real_less_asym]) 1);
+qed "real_less_imp_le";
+
+Goalw [real_le_def] "!!(x::real). x <= y ==> x < y | x = y";
+by (cut_facts_tac [real_linear] 1);
+by (fast_tac (claset() addEs [real_less_irrefl,real_less_asym]) 1);
+qed "real_le_imp_less_or_eq";
+
+Goalw [real_le_def] "!!z. z<w | z=w ==> z <=(w::real)";
+by (cut_facts_tac [real_linear] 1);
+by (fast_tac (claset() addEs [real_less_irrefl,real_less_asym]) 1);
+qed "real_less_or_eq_imp_le";
+
+Goal "(x <= (y::real)) = (x < y | x=y)";
+by (REPEAT(ares_tac [iffI, real_less_or_eq_imp_le, real_le_imp_less_or_eq] 1));
+qed "real_le_eq_less_or_eq";
+
+Goal "w <= (w::real)";
+by (simp_tac (simpset() addsimps [real_le_eq_less_or_eq]) 1);
+qed "real_le_refl";
+
+val prems = goal Real.thy "!!i. [| i <= j; j < k |] ==> i < (k::real)";
+by (dtac real_le_imp_less_or_eq 1);
+by (fast_tac (claset() addIs [real_less_trans]) 1);
+qed "real_le_less_trans";
+
+Goal "!! (i::real). [| i < j; j <= k |] ==> i < k";
+by (dtac real_le_imp_less_or_eq 1);
+by (fast_tac (claset() addIs [real_less_trans]) 1);
+qed "real_less_le_trans";
+
+Goal "!!i. [| i <= j; j <= k |] ==> i <= (k::real)";
+by (EVERY1 [dtac real_le_imp_less_or_eq, dtac real_le_imp_less_or_eq,
+            rtac real_less_or_eq_imp_le, fast_tac (claset() addIs [real_less_trans])]);
+qed "real_le_trans";
+
+Goal "!!z. [| z <= w; w <= z |] ==> z = (w::real)";
+by (EVERY1 [dtac real_le_imp_less_or_eq, dtac real_le_imp_less_or_eq,
+            fast_tac (claset() addEs [real_less_irrefl,real_less_asym])]);
+qed "real_le_anti_sym";
+
+Goal "!!x. [| ~ y < x; y ~= x |] ==> x < (y::real)";
+by (rtac not_real_leE 1);
+by (fast_tac (claset() addDs [real_le_imp_less_or_eq]) 1);
+qed "not_less_not_eq_real_less";
+
+Goal "(0r < %~R) = (R < 0r)";
+by (res_inst_tac [("x","R")]  real_preal_trichotomyE 1);
+by (auto_tac (claset(),simpset() addsimps [real_preal_not_minus_gt_zero,
+                        real_preal_not_less_zero,real_preal_zero_less,
+                        real_preal_minus_less_zero]));
+qed "real_minus_zero_less_iff";
+
+Addsimps [real_minus_zero_less_iff];
+
+Goal "(%~R < 0r) = (0r < R)";
+by (res_inst_tac [("x","R")]  real_preal_trichotomyE 1);
+by (auto_tac (claset(),simpset() addsimps [real_preal_not_minus_gt_zero,
+                        real_preal_not_less_zero,real_preal_zero_less,
+                        real_preal_minus_less_zero]));
+qed "real_minus_zero_less_iff2";
+
+(** lemma **)
+Goal "(0r < x) = (? y. x = %#y)";
+by (auto_tac (claset(),simpset() addsimps [real_preal_zero_less]));
+by (cut_inst_tac [("x","x")] real_preal_trichotomy 1);
+by (blast_tac (claset() addSEs [real_less_irrefl,
+     real_preal_not_minus_gt_zero RS notE]) 1);
+qed "real_gt_zero_preal_Ex";
+
+Goal "!!x. %#z < x ==> ? y. x = %#y";
+by (blast_tac (claset() addSDs [real_preal_zero_less RS real_less_trans]
+               addIs [real_gt_zero_preal_Ex RS iffD1]) 1);
+qed "real_gt_preal_preal_Ex";
+
+Goal "!!x. %#z <= x ==> ? y. x = %#y";
+by (blast_tac (claset() addDs [real_le_imp_less_or_eq,
+              real_gt_preal_preal_Ex]) 1);
+qed "real_ge_preal_preal_Ex";
+
+Goal "!!y. y <= 0r ==> !x. y < %#x";
+by (auto_tac (claset() addEs [real_le_imp_less_or_eq RS disjE]
+              addIs [real_preal_zero_less RSN(2,real_less_trans)],
+              simpset() addsimps [real_preal_zero_less]));
+qed "real_less_all_preal";
+
+Goal "!!y. ~ 0r < y ==> !x. y < %#x";
+by (blast_tac (claset() addSIs [real_less_all_preal,real_leI]) 1);
+qed "real_less_all_real2";
+
+(**** Derive alternative definition for real_less ****)
+(** lemma **)
+Goal "!!(R::real). ? A. S + A = R";
+by (res_inst_tac [("x","%~S + R")] exI 1);
+by (simp_tac (simpset() addsimps [real_add_minus,
+    real_add_zero_right] @ real_add_ac) 1);
+qed "real_lemma_add_left_ex";
+
+Goal "!!(R::real). R < S ==> ? T. R + T = S";
+by (res_inst_tac [("x","R")]  real_preal_trichotomyE 1);
+by (ALLGOALS(res_inst_tac [("x","S")]  real_preal_trichotomyE));
+by (auto_tac (claset() addSDs [preal_le_anti_sym] addSDs [preal_less_add_left_Ex],
+              simpset() addsimps [preal_less_le_iff,real_preal_add,real_minus_add_eq,
+               real_preal_minus_less_zero,real_less_not_refl,real_minus_ex,real_add_assoc,
+               real_preal_zero_less,real_preal_minus_less_all,real_add_minus_left,
+               real_preal_not_less_zero,real_add_zero_left,real_lemma_add_left_ex]));
+qed "real_less_add_left_Ex";
+
+Goal "!!(R::real). R < S ==> ? T. 0r < T & R + T = S";
+by (res_inst_tac [("x","R")]  real_preal_trichotomyE 1);
+by (ALLGOALS(res_inst_tac [("x","S")]  real_preal_trichotomyE));
+by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
+                         simpset() addsimps [real_preal_not_minus_gt_all,
+            real_preal_add, real_preal_not_less_zero,real_less_not_refl,
+    real_preal_not_minus_gt_zero,real_add_zero_left,real_minus_add_eq]));
+by (res_inst_tac [("x","%#D")] exI 1);
+by (res_inst_tac [("x","%#m+%#ma")] exI 2);
+by (res_inst_tac [("x","%#m")] exI 3);
+by (res_inst_tac [("x","%#D")] exI 4);
+by (auto_tac (claset(),simpset() addsimps [real_preal_zero_less,
+    real_preal_sum_zero_less,real_add_minus_left,real_add_assoc,
+                          real_add_minus,real_add_zero_right]));
+by (simp_tac (simpset() addsimps [real_add_assoc RS sym, 
+            real_add_minus_left,real_add_zero_left]) 1);
+qed "real_less_add_positive_left_Ex";
+
+(* lemmas *)
+(** change naff name(s)! **)
+Goal "!!S. (W < S) ==> (0r < S + %~W)";
+by (dtac real_less_add_positive_left_Ex 1);
+by (auto_tac (claset(),simpset() addsimps [real_add_minus,
+    real_add_zero_right] @ real_add_ac));
+qed "real_less_sum_gt_zero";
+
+Goal "!!S. T = S + W ==> S = T + %~W";
+by (asm_simp_tac (simpset() addsimps [real_add_minus,
+    real_add_zero_right] @ real_add_ac) 1);
+qed "real_lemma_change_eq_subj";
+
+(* FIXME: long! *)
+Goal "!!W. (0r < S + %~W) ==> (W < S)";
+by (rtac ccontr 1);
+by (dtac (real_leI RS real_le_imp_less_or_eq) 1);
+by (auto_tac (claset(),
+    simpset() addsimps [real_less_not_refl,real_add_minus]));
+by (EVERY1[dtac real_less_add_positive_left_Ex, etac exE, etac conjE]);
+by (asm_full_simp_tac (simpset() addsimps [real_add_zero_left]) 1);
+by (dtac real_lemma_change_eq_subj 1);
+by (auto_tac (claset(),simpset() addsimps [real_minus_minus]));
+by (dtac real_less_sum_gt_zero 1);
+by (asm_full_simp_tac (simpset() addsimps [real_minus_add_eq] @ real_add_ac) 1);
+by (EVERY1[rotate_tac 1, dtac (real_add_left_commute RS ssubst)]);
+by (auto_tac (claset() addEs [real_less_asym],
+              simpset() addsimps [real_add_minus,real_add_zero_right]));
+qed "real_sum_gt_zero_less";
+
+Goal "(0r < S + %~W) = (W < S)";
+by (fast_tac (claset() addIs [real_less_sum_gt_zero,
+    real_sum_gt_zero_less]) 1);
+qed "real_less_sum_gt_0_iff";
+
+Goal "((x::real) < y) = (%~y < %~x)";
+by (rtac (real_less_sum_gt_0_iff RS subst) 1);
+by (res_inst_tac [("W1","x")] (real_less_sum_gt_0_iff RS subst) 1);
+by (simp_tac (simpset() addsimps [real_add_commute]) 1);
+qed "real_less_swap_iff";
+
+Goal "!!T. [| R + L = S; 0r < L |] ==> R < S";
+by (rtac (real_less_sum_gt_0_iff RS iffD1) 1);
+by (auto_tac (claset(),simpset() addsimps [
+    real_add_minus,real_add_zero_right] @ real_add_ac));
+qed "real_lemma_add_positive_imp_less";
+
+Goal "!!(R::real). ? T. 0r < T & R + T = S ==> R < S";
+by (blast_tac (claset() addIs [real_lemma_add_positive_imp_less]) 1);
+qed "real_ex_add_positive_left_less";
+
+(*** alternative definition for real_less ***)
+Goal "!!(R::real). (? T. 0r < T & R + T = S) = (R < S)";
+by (fast_tac (claset() addSIs [real_less_add_positive_left_Ex,
+    real_ex_add_positive_left_less]) 1);
+qed "real_less_iffdef";
+
+Goal "(0r < x) = (%~x < x)";
+by (Step_tac 1);
+by (rtac ccontr 2 THEN forward_tac 
+    [real_leI RS real_le_imp_less_or_eq] 2);
+by (Step_tac 2);
+by (dtac (real_minus_zero_less_iff RS iffD2) 2);
+by (fast_tac (claset() addDs [real_less_trans]) 2);
+by (auto_tac (claset(),simpset() addsimps 
+    [real_gt_zero_preal_Ex,real_preal_minus_less_self]));
+qed "real_gt_zero_iff";
+
+Goal "(x < 0r) = (x < %~x)";
+by (rtac (real_minus_zero_less_iff RS subst) 1);
+by (stac real_gt_zero_iff 1);
+by (Full_simp_tac 1);
+qed "real_lt_zero_iff";
+
+Goalw [real_le_def] "(0r <= x) = (%~x <= x)";
+by (auto_tac (claset(),simpset() addsimps [real_lt_zero_iff RS sym]));
+qed "real_ge_zero_iff";
+
+Goalw [real_le_def] "(x <= 0r) = (x <= %~x)";
+by (auto_tac (claset(),simpset() addsimps [real_gt_zero_iff RS sym]));
+qed "real_le_zero_iff";
+
+Goal "(%#m1 <= %#m2) = (m1 <= m2)";
+by (auto_tac (claset() addSIs [preal_leI],
+    simpset() addsimps [real_less_le_iff RS sym]));
+by (dtac preal_le_less_trans 1 THEN assume_tac 1);
+by (etac preal_less_irrefl 1);
+qed "real_preal_le_iff";
+
+Goal "!!(x::real). [| 0r < x; 0r < y |] ==> 0r < x * y";
+by (auto_tac (claset(),simpset() addsimps [real_gt_zero_preal_Ex]));  
+by (res_inst_tac [("x","y*ya")] exI 1);
+by (full_simp_tac (simpset() addsimps [real_preal_mult]) 1);
+qed "real_mult_order";
+
+Goal "!!(x::real). [| x < 0r; y < 0r |] ==> 0r < x * y";
+by (REPEAT(dtac (real_minus_zero_less_iff RS iffD2) 1));
+by (dtac real_mult_order 1 THEN assume_tac 1);
+by (Asm_full_simp_tac 1);
+qed "real_mult_less_zero1";
+
+Goal "!!(x::real). [| 0r <= x; 0r <= y |] ==> 0r <= x * y";
+by (REPEAT(dtac real_le_imp_less_or_eq 1));
+by (auto_tac (claset() addIs [real_mult_order,
+    real_less_imp_le],simpset() addsimps [real_le_refl]));
+qed "real_le_mult_order";
+
+Goal "!!(x::real). [| x <= 0r; y <= 0r |] ==> 0r <= x * y";
+by (rtac real_less_or_eq_imp_le 1);
+by (dtac real_le_imp_less_or_eq 1 THEN etac disjE 1);
+by Auto_tac;
+by (dtac real_le_imp_less_or_eq 1);
+by (auto_tac (claset() addDs [real_mult_less_zero1],simpset()));
+qed "real_mult_le_zero1";
+
+Goal "!!(x::real). [| 0r <= x; y < 0r |] ==> x * y <= 0r";
+by (rtac real_less_or_eq_imp_le 1);
+by (dtac real_le_imp_less_or_eq 1 THEN etac disjE 1);
+by Auto_tac;
+by (dtac (real_minus_zero_less_iff RS iffD2) 1);
+by (rtac (real_minus_zero_less_iff RS subst) 1);
+by (blast_tac (claset() addDs [real_mult_order] 
+    addIs [real_minus_mult_eq2 RS ssubst]) 1);
+qed "real_mult_le_zero";
+
+Goal "!!(x::real). [| 0r < x; y < 0r |] ==> x*y < 0r";
+by (dtac (real_minus_zero_less_iff RS iffD2) 1);
+by (dtac real_mult_order 1 THEN assume_tac 1);
+by (rtac (real_minus_zero_less_iff RS iffD1) 1);
+by (asm_full_simp_tac (simpset() addsimps [real_minus_mult_eq2]) 1);
+qed "real_mult_less_zero";
+
+Goalw [real_one_def] "0r < 1r";
+by (auto_tac (claset() addIs [real_gt_zero_preal_Ex RS iffD2],
+    simpset() addsimps [real_preal_def]));
+qed "real_zero_less_one";
+
+(*** Completeness of reals ***)
+(** use supremum property of preal and theorems about real_preal **)
+              (*** a few lemmas ***)
+Goal "!!P y. ! x:P. 0r < x ==> ((? x:P. y < x) = (? X. %#X : P & y < %#X))";
+by (blast_tac (claset() addSDs [bspec,real_gt_zero_preal_Ex RS iffD1]) 1);
+qed "real_sup_lemma1";
+
+Goal "!!P. [| ! x:P. 0r < x; ? x. x: P; ? y. !x: P. x < y |] \
+\         ==> (? X. X: {w. %#w : P}) & (? Y. !X: {w. %#w : P}. X < Y)";
+by (rtac conjI 1);
+by (blast_tac (claset() addDs [bspec,real_gt_zero_preal_Ex RS iffD1]) 1);
+by Auto_tac;
+by (dtac bspec 1 THEN assume_tac 1);
+by (forward_tac [bspec] 1  THEN assume_tac 1);
+by (dtac real_less_trans 1 THEN assume_tac 1);
+by (dtac (real_gt_zero_preal_Ex RS iffD1) 1 THEN etac exE 1);
+by (res_inst_tac [("x","ya")] exI 1);
+by Auto_tac;
+by (dres_inst_tac [("x","%#X")] bspec 1 THEN assume_tac 1);
+by (etac real_preal_lessD 1);
+qed "real_sup_lemma2";
+
+(*-------------------------------------------------------------
+            Completeness of Positive Reals
+ -------------------------------------------------------------*)
+
+(* Supremum property for the set of positive reals *)
+(* FIXME: long proof - can be improved - need only have one case split *)
+(* will do for now *)
+Goal "!!P. [| ! x:P. 0r < x; ? x. x: P; ? y. !x: P. x < y |] \
+\         ==> (? S. !y. (? x: P. y < x) = (y < S))";
+by (res_inst_tac [("x","%#psup({w. %#w : P})")] exI 1);
+by Auto_tac;
+by (forward_tac [real_sup_lemma2] 1 THEN Auto_tac);
+by (case_tac "0r < ya" 1);
+by (dtac (real_gt_zero_preal_Ex RS iffD1) 1);
+by (dtac real_less_all_real2 2);
+by Auto_tac;
+by (rtac (preal_complete RS spec RS iffD1) 1);
+by Auto_tac;
+by (forward_tac [real_gt_preal_preal_Ex] 1);
+by Auto_tac;
+(* second part *)
+by (rtac (real_sup_lemma1 RS iffD2) 1 THEN assume_tac 1);
+by (case_tac "0r < ya" 1);
+by (auto_tac (claset() addSDs [real_less_all_real2,
+        real_gt_zero_preal_Ex RS iffD1],simpset()));
+by (forward_tac [real_sup_lemma2] 2 THEN Auto_tac);
+by (forward_tac [real_sup_lemma2] 1 THEN Auto_tac);
+by (rtac (preal_complete RS spec RS iffD2 RS bexE) 1);
+by (Fast_tac 3);
+by (Fast_tac 1);
+by (Fast_tac 1);
+by (Blast_tac 1);
+qed "posreal_complete";
+
+(*------------------------------------------------------------------
+
+ ------------------------------------------------------------------*)
+
+Goal "!!(A::real). A < B ==> A + C < B + C";
+by (dtac (real_less_iffdef RS iffD2) 1);
+by (rtac (real_less_iffdef RS iffD1) 1);
+by (REPEAT(Step_tac 1));
+by (full_simp_tac (simpset() addsimps real_add_ac) 1);
+qed "real_add_less_mono1";
+
+Goal "!!(A::real). A < B ==> C + A < C + B";
+by (auto_tac (claset() addIs [real_add_less_mono1],
+    simpset() addsimps [real_add_commute]));
+qed "real_add_less_mono2";
+
+Goal "!!(A::real). A + C < B + C ==> A < B";
+by (dres_inst_tac [("C","%~C")] real_add_less_mono1 1);
+by (asm_full_simp_tac (simpset() addsimps [real_add_assoc,
+    real_add_minus,real_add_zero_right]) 1);
+qed "real_less_add_right_cancel";
+
+Goal "!!(A::real). C + A < C + B ==> A < B";
+by (dres_inst_tac [("C","%~C")] real_add_less_mono2 1);
+by (asm_full_simp_tac (simpset() addsimps [real_add_assoc RS sym,
+    real_add_minus_left,real_add_zero_left]) 1);
+qed "real_less_add_left_cancel";
+
+Goal "!!x. [| 0r < x; 0r < y |] ==> 0r < x + y";
+by (REPEAT(dtac (real_gt_zero_preal_Ex RS iffD1) 1));
+by (rtac (real_gt_zero_preal_Ex RS iffD2) 1);
+by (Step_tac 1);
+by (res_inst_tac [("x","y + ya")] exI 1);
+by (full_simp_tac (simpset() addsimps [real_preal_add]) 1);
+qed "real_add_order";
+
+Goal "!!(x::real). [| 0r <= x; 0r <= y |] ==> 0r <= x + y";
+by (REPEAT(dtac real_le_imp_less_or_eq 1));
+by (auto_tac (claset() addIs [real_add_order,
+    real_less_imp_le],simpset() addsimps [real_add_zero_left,
+    real_add_zero_right,real_le_refl]));
+qed "real_le_add_order";
+
+Goal 
+      "!!x. [| R1 < S1; R2 < S2 |] ==> R1 + R2 < S1 + (S2::real)";
+by (dtac (real_less_iffdef RS iffD2) 1);
+by (dtac (real_less_iffdef RS iffD2) 1);
+by (rtac (real_less_iffdef RS iffD1) 1);
+by Auto_tac;
+by (res_inst_tac [("x","T + Ta")] exI 1);
+by (auto_tac (claset(),simpset() addsimps [real_add_order] @ real_add_ac));
+qed "real_add_less_mono";
+
+Goal "!!(x::real). [| 0r <= x; 0r <= y |] ==> 0r <= x + y";
+by (REPEAT(dtac real_le_imp_less_or_eq 1));
+by (auto_tac (claset() addIs [real_add_order,
+    real_less_imp_le],simpset() addsimps [real_add_zero_left,
+    real_add_zero_right,real_le_refl]));
+qed "real_le_add_order";
+
+Goal "!!(q1::real). q1 <= q2  ==> x + q1 <= x + q2";
+by (dtac real_le_imp_less_or_eq 1);
+by (Step_tac 1);
+by (auto_tac (claset() addSIs [real_le_refl,
+    real_less_imp_le,real_add_less_mono1],
+    simpset() addsimps [real_add_commute]));
+qed "real_add_left_le_mono1";
+
+Goal "!!(q1::real). q1 <= q2  ==> q1 + x <= q2 + x";
+by (auto_tac (claset() addDs [real_add_left_le_mono1],
+    simpset() addsimps [real_add_commute]));
+qed "real_add_le_mono1";
+
+Goal "!!k l::real. [|i<=j;  k<=l |] ==> i + k <= j + l";
+by (etac (real_add_le_mono1 RS real_le_trans) 1);
+by (simp_tac (simpset() addsimps [real_add_commute]) 1);
+(*j moves to the end because it is free while k, l are bound*)
+by (etac real_add_le_mono1 1);
+qed "real_add_le_mono";
+
+Goal "EX (x::real). x < y";
+by (rtac (real_add_zero_right RS subst) 1);
+by (res_inst_tac [("x","y + %~1r")] exI 1);
+by (auto_tac (claset() addSIs [real_add_less_mono2],
+    simpset() addsimps [real_minus_zero_less_iff2,
+    real_zero_less_one]));
+qed "real_less_Ex";
+(*---------------------------------------------------------------------------------
+             An embedding of the naturals in the reals
+ ---------------------------------------------------------------------------------*)
+
+Goalw [real_nat_def] "%%#0 = 1r";
+by (full_simp_tac (simpset() addsimps [pnat_one_iff RS sym,real_preal_def]) 1);
+by (fold_tac [real_one_def]);
+by (rtac refl 1);
+qed "real_nat_one";
+
+Goalw [real_nat_def] "%%#1 = 1r + 1r";
+by (full_simp_tac (simpset() addsimps [real_preal_def,real_one_def,
+    pnat_two_eq,real_add,prat_pnat_add RS sym,preal_prat_add RS sym
+    ] @ pnat_add_ac) 1);
+qed "real_nat_two";
+
+Goalw [real_nat_def]
+          "%%#n1 + %%#n2 = %%#(n1 + n2) + 1r";
+by (full_simp_tac (simpset() addsimps [real_nat_one RS sym,
+    real_nat_def,real_preal_add RS sym,preal_prat_add RS sym,
+    prat_pnat_add RS sym,pnat_nat_add]) 1);
+qed "real_nat_add";
+
+Goal "%%#(n + 1) = %%#n + 1r";
+by (res_inst_tac [("x1","1r")] (real_add_right_cancel RS iffD1) 1);
+by (rtac (real_nat_add RS subst) 1);
+by (full_simp_tac (simpset() addsimps [real_nat_two,real_add_assoc]) 1);
+qed "real_nat_add_one";
+
+Goal "Suc n = n + 1";
+by Auto_tac;
+qed "lemma";
+
+Goal "%%#Suc n = %%#n + 1r";
+by (stac lemma 1);
+by (rtac real_nat_add_one 1);
+qed "real_nat_Suc";
+
+Goal "inj(real_nat)";
+by (rtac injI 1);
+by (rewtac real_nat_def);
+by (dtac (inj_real_preal RS injD) 1);
+by (dtac (inj_preal_prat RS injD) 1);
+by (dtac (inj_prat_pnat RS injD) 1);
+by (etac (inj_pnat_nat RS injD) 1);
+qed "inj_real_nat";
+
+Goalw [real_nat_def] "0r < %%#n";
+by (rtac (real_gt_zero_preal_Ex RS iffD2) 1);
+by (Blast_tac 1);
+qed "real_nat_less_zero";
+
+Goal "!!n. 1r <= %%#n";
+by (simp_tac (simpset() addsimps [real_nat_one RS sym]) 1);
+by (nat_ind_tac "n" 1);
+by (auto_tac (claset(),simpset () 
+    addsimps [real_nat_Suc,real_le_refl,real_nat_one]));
+by (res_inst_tac [("t","1r")] (real_add_zero_left RS subst) 1);
+by (rtac real_add_le_mono 1);
+by (auto_tac (claset(),simpset () 
+    addsimps [real_le_refl,real_nat_less_zero,
+    real_less_imp_le,real_add_zero_left]));
+qed "real_nat_less_one";
+
+Goal "rinv(%%#n) ~= 0r";
+by (rtac ((real_nat_less_zero RS 
+    real_not_refl2 RS not_sym) RS rinv_not_zero) 1);
+qed "real_nat_rinv_not_zero";
+
+Goal "!!x. rinv(%%#x) = rinv(%%#y) ==> x = y";
+by (rtac (inj_real_nat RS injD) 1);
+by (res_inst_tac [("n2","x")] 
+    (real_nat_rinv_not_zero RS real_mult_left_cancel RS iffD1) 1);
+by (full_simp_tac (simpset() addsimps [(real_nat_less_zero RS 
+    real_not_refl2 RS not_sym) RS real_mult_inv_left]) 1);
+by (asm_full_simp_tac (simpset() addsimps [(real_nat_less_zero RS 
+    real_not_refl2 RS not_sym)]) 1);
+qed "real_nat_rinv_inj";
+
+Goal "!!x. 0r < x ==> 0r < rinv x";
+by (EVERY1[rtac ccontr, dtac real_leI]);
+by (forward_tac [real_minus_zero_less_iff2 RS iffD2] 1);
+by (forward_tac [real_not_refl2 RS not_sym] 1);
+by (dtac (real_not_refl2 RS not_sym RS rinv_not_zero) 1);
+by (EVERY1[dtac real_le_imp_less_or_eq, Step_tac]); 
+by (dtac real_mult_less_zero1 1 THEN assume_tac 1);
+by (auto_tac (claset() addIs [real_zero_less_one RS real_less_asym],
+    simpset() addsimps [real_minus_mult_eq1 RS sym]));
+qed "real_rinv_gt_zero";
+
+Goal "!!x. x < 0r ==> rinv x < 0r";
+by (forward_tac [real_not_refl2] 1);
+by (dtac (real_minus_zero_less_iff RS iffD2) 1);
+by (rtac (real_minus_zero_less_iff RS iffD1) 1);
+by (dtac (real_minus_rinv RS sym) 1);
+by (auto_tac (claset() addIs [real_rinv_gt_zero],
+    simpset()));
+qed "real_rinv_less_zero";
+
+Goal "x+x=x*(1r+1r)";
+by (simp_tac (simpset() addsimps [real_add_mult_distrib2]) 1);
+qed "real_add_self";
+
+Goal "x < x + 1r";
+by (rtac (real_less_sum_gt_0_iff RS iffD1) 1);
+by (full_simp_tac (simpset() addsimps [real_zero_less_one,
+    real_add_assoc,real_add_minus,real_add_zero_right,
+    real_add_left_commute]) 1);
+qed "real_self_less_add_one";
+
+Goal "1r < 1r + 1r";
+by (rtac real_self_less_add_one 1);
+qed "real_one_less_two";
+
+Goal "0r < 1r + 1r";
+by (rtac ([real_zero_less_one,
+          real_one_less_two] MRS real_less_trans) 1);
+qed "real_zero_less_two";
+
+Goal "1r + 1r ~= 0r";
+by (rtac (real_zero_less_two RS real_not_refl2 RS not_sym) 1);
+qed "real_two_not_zero";
+
+Addsimps [real_two_not_zero];
+
+Goal "x*rinv(1r + 1r) + x*rinv(1r + 1r) = x";
+by (stac real_add_self 1);
+by (full_simp_tac (simpset() addsimps [real_mult_assoc]) 1);
+qed "real_sum_of_halves";
+
+Goal "!!(x::real). [| 0r<z; x<y |] ==> x*z<y*z";       
+by (rotate_tac 1 1);
+by (dtac real_less_sum_gt_zero 1);
+by (rtac real_sum_gt_zero_less 1);
+by (dtac real_mult_order 1 THEN assume_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [real_add_mult_distrib2,
+    real_minus_mult_eq2 RS sym, real_mult_commute ]) 1);
+qed "real_mult_less_mono1";
+
+Goal "!!(y::real). [| 0r<z; x<y |] ==> z*x<z*y";       
+by (asm_simp_tac (simpset() addsimps [real_mult_commute,real_mult_less_mono1]) 1);
+qed "real_mult_less_mono2";
+
+Goal "!!(x::real). [| 0r<z; x*z<y*z |] ==> x<y";
+by (forw_inst_tac [("x","x*z")] (real_rinv_gt_zero 
+                       RS real_mult_less_mono1) 1);
+by (auto_tac (claset(),simpset() addsimps 
+     [real_mult_assoc,real_not_refl2 RS not_sym]));
+qed "real_mult_less_cancel1";
+
+Goal "!!(x::real). [| 0r<z; z*x<z*y |] ==> x<y";
+by (etac real_mult_less_cancel1 1);
+by (asm_full_simp_tac (simpset() addsimps [real_mult_commute]) 1);
+qed "real_mult_less_cancel2";
+
+Goal "!!z. 0r < z ==> (x*z < y*z) = (x < y)";
+by (blast_tac (claset() addIs [real_mult_less_mono1,
+    real_mult_less_cancel1]) 1);
+qed "real_mult_less_iff1";
+
+Goal "!!z. 0r < z ==> (z*x < z*y) = (x < y)";
+by (blast_tac (claset() addIs [real_mult_less_mono2,
+    real_mult_less_cancel2]) 1);
+qed "real_mult_less_iff2";
+
+Addsimps [real_mult_less_iff1,real_mult_less_iff2];
+
+Goal "!!(x::real). [| 0r<=z; x<y |] ==> x*z<=y*z";
+by (EVERY1 [rtac real_less_or_eq_imp_le, dtac real_le_imp_less_or_eq]);
+by (auto_tac (claset() addIs [real_mult_less_mono1],simpset()));
+qed "real_mult_le_less_mono1";
+
+Goal "!!(x::real). [| 0r<=z; x<y |] ==> z*x<=z*y";
+by (asm_simp_tac (simpset() addsimps [real_mult_commute,real_mult_le_less_mono1]) 1);
+qed "real_mult_le_less_mono2";
+
+Goal "!!x y (z::real). [| 0r<=z; x<=y |] ==> z*x<=z*y";
+by (dres_inst_tac [("x","x")] real_le_imp_less_or_eq 1);
+by (auto_tac (claset() addIs [real_mult_le_less_mono2,real_le_refl],simpset()));
+qed "real_mult_le_le_mono1";
+
+Goal "!!(x::real). x < y ==> x < (x + y)*rinv(1r + 1r)";
+by (dres_inst_tac [("C","x")] real_add_less_mono2 1);
+by (dtac (real_add_self RS subst) 1);
+by (dtac (real_zero_less_two RS real_rinv_gt_zero RS 
+          real_mult_less_mono1) 1);
+by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc]) 1);
+qed "real_less_half_sum";
+
+Goal "!!(x::real). x < y ==> (x + y)*rinv(1r + 1r) < y";
+by (dres_inst_tac [("C","y")] real_add_less_mono1 1);
+by (dtac (real_add_self RS subst) 1);
+by (dtac (real_zero_less_two RS real_rinv_gt_zero RS 
+          real_mult_less_mono1) 1);
+by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc]) 1);
+qed "real_gt_half_sum";
+
+Goal "!!(x::real). x < y ==> EX r. x < r & r < y";
+by (blast_tac (claset() addSIs [real_less_half_sum,real_gt_half_sum]) 1);
+qed "real_dense";
+
+Goal "(EX n. rinv(%%#n) < r) = (EX n. 1r < r * %%#n)";
+by (Step_tac 1);
+by (dres_inst_tac [("n1","n")] (real_nat_less_zero 
+                       RS real_mult_less_mono1) 1);
+by (dres_inst_tac [("n2","n")] (real_nat_less_zero RS 
+        real_rinv_gt_zero RS real_mult_less_mono1) 2);
+by (auto_tac (claset(),simpset() addsimps [(real_nat_less_zero RS 
+    real_not_refl2 RS not_sym),real_mult_assoc]));
+qed "real_nat_rinv_Ex_iff";
+
+Goalw [real_nat_def] "(%%#n < %%#m) = (n < m)";
+by Auto_tac;
+qed "real_nat_less_iff";
+
+Addsimps [real_nat_less_iff];
+
+Goal "!!u. 0r < u  ==> (u < rinv (%%#n)) = (%%#n < rinv(u))";
+by (Step_tac 1);
+by (res_inst_tac [("n2","n")] (real_nat_less_zero RS 
+    real_rinv_gt_zero RS real_mult_less_cancel1) 1);
+by (res_inst_tac [("x1","u")] ( real_rinv_gt_zero
+   RS real_mult_less_cancel1) 2);
+by (auto_tac (claset(),simpset() addsimps [real_nat_less_zero, 
+    real_not_refl2 RS not_sym]));
+by (res_inst_tac [("z","u")] real_mult_less_cancel2 1);
+by (res_inst_tac [("n1","n")] (real_nat_less_zero RS 
+    real_mult_less_cancel2) 3);
+by (auto_tac (claset(),simpset() addsimps [real_nat_less_zero, 
+    real_not_refl2 RS not_sym,real_mult_assoc RS sym]));
+qed "real_nat_less_rinv_iff";
+
+Goal "!!x. 0r < u ==> (u = rinv(%%#n)) = (%%#n = rinv u)";
+by (auto_tac (claset(),simpset() addsimps [real_rinv_rinv,
+    real_nat_less_zero,real_not_refl2 RS not_sym]));
+qed "real_nat_rinv_eq_iff";
+
+(*
+(*------------------------------------------------------------------
+     lemmas about upper bounds and least upper bound
+ ------------------------------------------------------------------*)
+Goalw [real_ub_def] 
+          "!!S. [| real_ub u S; x : S |] ==> x <= u";
+by Auto_tac;
+qed "real_ubD";
+
+*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/Real.thy	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,61 @@
+(*  Title       : Real.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : The reals
+*) 
+
+Real = PReal +
+
+constdefs
+    realrel   ::  "((preal * preal) * (preal * preal)) set"
+    "realrel  ==  {p. ? x1 y1 x2 y2. p=((x1::preal,y1),(x2,y2)) & x1+y2 = x2+y1}" 
+
+typedef real = "{x::(preal*preal).True}/realrel"          (Equiv.quotient_def)
+
+
+instance
+   real  :: {ord,plus,times}
+
+consts 
+
+  "0r"       :: real               ("0r")   
+  "1r"       :: real               ("1r")  
+
+defs
+
+  real_zero_def      "0r == Abs_real(realrel^^{(@#($#1p),@#($#1p))})"
+  real_one_def       "1r == Abs_real(realrel^^{(@#($#1p) + @#($#1p),@#($#1p))})"
+
+constdefs
+
+  real_preal :: preal => real              ("%#_" [80] 80)
+  "%# m     == Abs_real(realrel^^{(m+@#($#1p),@#($#1p))})"
+
+  real_minus :: real => real               ("%~ _" [80] 80) 
+  "%~ R     ==  Abs_real(UN p:Rep_real(R). split (%x y. realrel^^{(y,x)}) p)"
+
+  rinv       :: real => real
+  "rinv(R)   == (@S. R ~= 0r & S*R = 1r)"
+
+  real_nat :: nat => real                  ("%%# _" [80] 80) 
+  "%%# n      == %#(@#($#(*# n)))"
+
+defs
+
+  real_add_def  
+  "P + Q == Abs_real(UN p1:Rep_real(P). UN p2:Rep_real(Q).
+                split(%x1 y1. split(%x2 y2. realrel^^{(x1+x2, y1+y2)}) p2) p1)"
+  
+  real_mult_def  
+  "P * Q == Abs_real(UN p1:Rep_real(P). UN p2:Rep_real(Q).
+                split(%x1 y1. split(%x2 y2. realrel^^{(x1*x2+y1*y2,x1*y2+x2*y1)}) p2) p1)"
+
+  real_less_def
+  "P < (Q::real) == EX x1 y1 x2 y2. x1 + y2 < x2 + y1 & 
+                                   (x1,y1::preal):Rep_real(P) &
+                                   (x2,y2):Rep_real(Q)" 
+
+  real_le_def
+  "P <= (Q::real) == ~(Q < P)"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/RealAbs.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,239 @@
+(*  Title       : RealAbs.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Absolute value function for the reals
+*) 
+
+open RealAbs;
+
+(*----------------------------------------------------------------------------
+       Properties of the absolute value function over the reals
+       (adapted version of previously proved theorems about abs)
+ ----------------------------------------------------------------------------*)
+Goalw [rabs_def] "rabs r = (if 0r<=r then r else %~r)";
+by (Step_tac 1);
+qed "rabs_iff";
+
+Goalw [rabs_def] "rabs 0r = 0r";
+by (rtac (real_le_refl RS if_P) 1);
+qed "rabs_zero";
+
+Addsimps [rabs_zero];
+
+Goalw [rabs_def] "rabs 0r = %~0r";
+by (stac real_minus_zero 1);
+by (rtac if_cancel 1);
+qed "rabs_minus_zero";
+
+val [prem] = goalw thy [rabs_def] "0r<=x ==> rabs x = x";
+by (rtac (prem RS if_P) 1);
+qed "rabs_eqI1";
+
+val [prem] = goalw thy [rabs_def] "0r<x ==> rabs x = x";
+by (simp_tac (simpset() addsimps [(prem RS real_less_imp_le),rabs_eqI1]) 1);
+qed "rabs_eqI2";
+
+val [prem] = goalw thy [rabs_def,real_le_def] "x<0r ==> rabs x = %~x";
+by (simp_tac (simpset() addsimps [prem,if_not_P]) 1);
+qed "rabs_minus_eqI2";
+
+Goal "!!x. x<=0r ==> rabs x = %~x";
+by (dtac real_le_imp_less_or_eq 1);
+by (fast_tac (HOL_cs addIs [rabs_minus_zero,rabs_minus_eqI2]) 1);
+qed "rabs_minus_eqI1";
+
+Goalw [rabs_def,real_le_def] "0r<= rabs x";
+by (full_simp_tac (simpset()  setloop (split_tac [expand_if])) 1);
+by (blast_tac (claset() addDs [real_minus_zero_less_iff RS iffD2,
+    real_less_asym]) 1);
+qed "rabs_ge_zero";
+
+Goal "rabs(rabs x)=rabs x";
+by (res_inst_tac [("r1","rabs x")] (rabs_iff RS ssubst) 1);
+by (blast_tac (claset() addIs [if_P,rabs_ge_zero]) 1);
+qed "rabs_idempotent";
+
+Goalw [rabs_def] "(x=0r) = (rabs x = 0r)";
+by (full_simp_tac (simpset() setloop (split_tac [expand_if])) 1);
+qed "rabs_zero_iff";
+
+Goal  "(x ~= 0r) = (rabs x ~= 0r)";
+by (full_simp_tac (simpset() addsimps [rabs_zero_iff RS sym] 
+    setloop (split_tac [expand_if])) 1);
+qed "rabs_not_zero_iff";
+
+Goalw [rabs_def] "x<=rabs x";
+by (full_simp_tac (simpset() addsimps [real_le_refl] setloop (split_tac [expand_if])) 1);
+by (auto_tac (claset() addDs [not_real_leE RS real_less_imp_le],
+    simpset() addsimps [real_le_zero_iff]));
+qed "rabs_ge_self";
+
+Goalw [rabs_def] "%~x<=rabs x";
+by (full_simp_tac (simpset() addsimps [real_le_refl,
+    real_ge_zero_iff] setloop (split_tac [expand_if])) 1);
+qed "rabs_ge_minus_self";
+
+(* case splits nightmare *)
+Goalw [rabs_def] "rabs(x*y) = (rabs x)*(rabs y)";
+by (auto_tac (claset(),simpset() addsimps [real_minus_mult_eq1,
+   real_minus_mult_commute,real_minus_mult_eq2] setloop (split_tac [expand_if])));
+by (blast_tac (claset() addDs [real_le_mult_order]) 1);
+by (auto_tac (claset() addSDs [not_real_leE],simpset()));
+by (EVERY1[dtac real_mult_le_zero, assume_tac, dtac real_le_anti_sym]);
+by (EVERY[dtac real_mult_le_zero 3, assume_tac 3, dtac real_le_anti_sym 3]);
+by (dtac real_mult_less_zero1 5 THEN assume_tac 5);
+by (auto_tac (claset() addDs [real_less_asym,sym],
+    simpset() addsimps [real_minus_mult_eq2 RS sym] @real_mult_ac));
+qed "rabs_mult";
+
+Goalw [rabs_def] "!!x. x~= 0r ==> rabs(rinv(x)) = rinv(rabs(x))";
+by (auto_tac (claset(),simpset() addsimps [real_minus_rinv] 
+    setloop (split_tac [expand_if])));
+by (ALLGOALS(dtac not_real_leE));
+by (etac real_less_asym 1);
+by (blast_tac (claset() addDs [real_le_imp_less_or_eq,
+          real_rinv_gt_zero]) 1);
+by (dtac (rinv_not_zero RS not_sym) 1);
+by (rtac (real_rinv_less_zero RSN (2,real_less_asym)) 1);
+by (assume_tac 2);
+by (blast_tac (claset() addSDs [real_le_imp_less_or_eq]) 1);
+qed "rabs_rinv";
+
+val [prem] = goal thy "y ~= 0r ==> rabs(x*rinv(y)) = rabs(x)*rinv(rabs(y))";
+by (res_inst_tac [("c1","rabs y")] (real_mult_left_cancel RS subst) 1);
+by (simp_tac (simpset() addsimps [(rabs_not_zero_iff RS sym), prem]) 1);
+by (simp_tac (simpset() addsimps [(rabs_mult RS sym) ,real_mult_inv_right, 
+    prem,rabs_not_zero_iff RS sym] @ real_mult_ac) 1);
+qed "rabs_mult_rinv";
+
+Goal "rabs(x+y) <= rabs x + rabs y";
+by (EVERY1 [res_inst_tac [("Q1","0r<=x+y")] (expand_if RS ssubst), rtac conjI]);
+by (asm_simp_tac (simpset() addsimps [rabs_eqI1,real_add_le_mono,rabs_ge_self]) 1);
+by (asm_simp_tac (simpset() addsimps [not_real_leE,rabs_minus_eqI2,real_add_le_mono,
+                                     rabs_ge_minus_self,real_minus_add_eq]) 1);
+qed "rabs_triangle_ineq";
+
+Goal "rabs(w + x + y + z) <= rabs(w) + rabs(x) + rabs(y) + rabs(z)";
+by (full_simp_tac (simpset() addsimps [real_add_assoc]) 1);
+by (blast_tac (claset() addSIs [(rabs_triangle_ineq RS real_le_trans),
+                real_add_left_le_mono1,real_le_refl]) 1);
+qed "rabs_triangle_ineq_four";
+
+Goalw [rabs_def] "rabs(%~x)=rabs(x)";
+by (auto_tac (claset() addSDs [not_real_leE,real_less_asym] addIs [real_le_anti_sym],
+   simpset() addsimps [real_ge_zero_iff] setloop (split_tac [expand_if])));
+qed "rabs_minus_cancel";
+
+Goal "rabs(x + %~y) <= rabs x + rabs y";
+by (res_inst_tac [("x1","y")] (rabs_minus_cancel RS subst) 1);
+by (rtac rabs_triangle_ineq 1);
+qed "rabs_triangle_minus_ineq";
+
+Goal "rabs (x + y + (%~l + %~m)) <= rabs(x + %~l) + rabs(y + %~m)";
+by (full_simp_tac (simpset() addsimps [real_add_assoc]) 1);
+by (res_inst_tac [("x1","y")] (real_add_left_commute RS ssubst) 1);
+by (rtac (real_add_assoc RS subst) 1);
+by (rtac rabs_triangle_ineq 1);
+qed "rabs_sum_triangle_ineq";
+
+Goal "[| rabs x < r; rabs y < s |] ==> rabs(x+y) < r+s";
+by (rtac real_le_less_trans 1);
+by (rtac rabs_triangle_ineq 1);
+by (REPEAT (ares_tac [real_add_less_mono] 1));
+qed "rabs_add_less";
+
+Goal "!!x y. [| rabs x < r; rabs y < s |] ==> rabs(x+ %~y) < r+s";
+by (rotate_tac 1 1);
+by (dtac (rabs_minus_cancel RS ssubst) 1);
+by (asm_simp_tac (simpset() addsimps [rabs_add_less]) 1);
+qed "rabs_add_minus_less";
+
+(* lemmas manipulating terms *)
+Goal "(0r*x<r)=(0r<r)";
+by (Simp_tac 1);
+qed "real_mult_0_less";
+
+Goal "[| 0r<y; x<r; y*r<t*s |] ==> y*x<t*s";
+(*why PROOF FAILED for this*)
+by (best_tac (claset() addIs [real_mult_less_mono2, real_less_trans]) 1);
+qed "real_mult_less_trans";
+
+Goal "!!(x::real) y.[| 0r<=y; x<r; y*r<t*s; 0r<t*s|] ==> y*x<t*s";
+by (dtac real_le_imp_less_or_eq 1);
+by (fast_tac (HOL_cs addEs [(real_mult_0_less RS iffD2),real_mult_less_trans]) 1);
+qed "real_mult_le_less_trans";
+
+(* proofs lifted from previous older version *)
+Goal "[| rabs x<r; rabs y<s |] ==> rabs(x*y)<r*s";
+by (simp_tac (simpset() addsimps [rabs_mult]) 1);
+by (rtac real_mult_le_less_trans 1);
+by (rtac rabs_ge_zero 1);
+by (assume_tac 1);
+by (blast_tac (HOL_cs addIs [rabs_ge_zero, real_mult_less_mono1, 
+			     real_le_less_trans]) 1);
+by (blast_tac (HOL_cs addIs [rabs_ge_zero, real_mult_order, 
+			     real_le_less_trans]) 1);
+qed "rabs_mult_less";
+
+Goal "!!x. [| rabs x < r; rabs y < s |] \
+\          ==> rabs(x)*rabs(y)<r*s";
+by (auto_tac (claset() addIs [rabs_mult_less],
+              simpset() addsimps [rabs_mult RS sym]));
+qed "rabs_mult_less2";
+
+Goal "!! x y r. 1r < rabs x ==> rabs y <= rabs(x*y)";
+by (cut_inst_tac [("x1","y")] (rabs_ge_zero RS real_le_imp_less_or_eq) 1);
+by (EVERY1[etac disjE,rtac real_less_imp_le]);
+by (dres_inst_tac [("W","1r")]  real_less_sum_gt_zero 1);
+by (forw_inst_tac [("y","rabs x + %~1r")] real_mult_order 1);
+by (assume_tac 1);
+by (rtac real_sum_gt_zero_less 1);
+by (asm_full_simp_tac (simpset() addsimps [real_add_mult_distrib2,
+    rabs_mult, real_mult_commute,real_minus_mult_eq1 RS sym]) 1);
+by (dtac sym 1);
+by (asm_full_simp_tac (simpset() addsimps [real_le_refl,rabs_mult]) 1);
+qed "rabs_mult_le";
+
+Goal "!!x. [| 1r < rabs x; r < rabs y|] ==> r < rabs(x*y)";
+by (fast_tac (HOL_cs addIs [rabs_mult_le, real_less_le_trans]) 1);
+qed "rabs_mult_gt";
+
+Goal "!!r. rabs(x)<r ==> 0r<r";
+by (blast_tac (claset() addSIs [real_le_less_trans,rabs_ge_zero]) 1);
+qed "rabs_less_gt_zero";
+
+Goalw [rabs_def] "rabs 1r = 1r";
+by (auto_tac (claset() addSDs [not_real_leE RS real_less_asym],
+   simpset() addsimps [real_zero_less_one] setloop (split_tac [expand_if])));
+qed "rabs_one";
+
+Goal "[| 0r < x ; x < r |] ==> rabs x < r";
+by (asm_simp_tac (simpset() addsimps [rabs_eqI2]) 1);
+qed "rabs_lessI";
+
+Goal "rabs x =x | rabs x = %~x";
+by (cut_inst_tac [("R1.0","0r"),("R2.0","x")] real_linear 1);
+by (fast_tac (claset() addIs [rabs_eqI2,rabs_minus_eqI2,
+                            rabs_zero,rabs_minus_zero]) 1);
+qed "rabs_disj";
+
+Goal "!!x. rabs x = y ==> x = y | %~x = y";
+by (dtac sym 1);
+by (hyp_subst_tac 1);
+by (res_inst_tac [("x1","x")] (rabs_disj RS disjE) 1);
+by (REPEAT(Asm_simp_tac 1));
+qed "rabs_eq_disj";
+
+Goal "(rabs x < r) = (%~r<x & x<r)";
+by (Step_tac 1);
+by (rtac (real_less_swap_iff RS iffD2) 1);
+by (asm_simp_tac (simpset() addsimps [(rabs_ge_minus_self 
+    RS real_le_less_trans)]) 1);
+by (asm_simp_tac (simpset() addsimps [(rabs_ge_self 
+    RS real_le_less_trans)]) 1);
+by (EVERY1 [dtac (real_less_swap_iff RS iffD1), rotate_tac 1, 
+            dtac (real_minus_minus RS subst), 
+            cut_inst_tac [("x","x")] rabs_disj, dtac disjE ]);
+by (assume_tac 3 THEN Auto_tac);
+qed "rabs_interval_iff";
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real/RealAbs.thy	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,13 @@
+(*  Title       : RealAbs.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Absolute value function for the reals
+*) 
+
+RealAbs = Real +
+
+constdefs
+   rabs   :: real => real
+   "rabs r      == if 0r<=r then r else %~r" 
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Group.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,222 @@
+(*  Title:      HOL/Integ/Group.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1997 TU Muenchen
+*)
+
+(*** Groups ***)
+
+(* Derives the well-known convergent set of equations for groups
+   based on the unary inverse zero-x.
+*)
+
+Goal "!!x::'a::add_group. (zero-x)+(x+y) = y";
+by (rtac trans 1);
+by (rtac (plus_assoc RS sym) 1);
+by (stac left_inv 1);
+by (rtac zeroL 1);
+qed "left_inv2";
+
+Goal "!!x::'a::add_group. (zero-(zero-x)) = x";
+by (rtac trans 1);
+by (res_inst_tac [("x","zero-x")] left_inv2 2);
+by (stac left_inv 1);
+by (rtac (zeroR RS sym) 1);
+qed "inv_inv";
+
+Goal "zero-zero = (zero::'a::add_group)";
+by (rtac trans 1);
+by (rtac (zeroR RS sym) 1);
+by (rtac trans 1);
+by (res_inst_tac [("x","zero")] left_inv2 2);
+by (simp_tac (simpset() addsimps [zeroR]) 1);
+qed "inv_zero";
+
+Goal "!!x::'a::add_group. x+(zero-x) = zero";
+by (rtac trans 1);
+by (res_inst_tac [("x","zero-x")] left_inv 2);
+by (stac inv_inv 1);
+by (rtac refl 1);
+qed "right_inv";
+
+Goal "!!x::'a::add_group. x+((zero-x)+y) = y";
+by (rtac trans 1);
+by (res_inst_tac [("x","zero-x")] left_inv2 2);
+by (stac inv_inv 1);
+by (rtac refl 1);
+qed "right_inv2";
+
+val plus_cong = read_instantiate [("f1","op +")] (arg_cong RS cong);
+
+Goal "!!x::'a::add_group. zero-(x+y) = (zero-y)+(zero-x)";
+by (rtac trans 1);
+ by (rtac zeroR 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (rtac refl 2);
+ by (res_inst_tac [("x","x+y")] right_inv 2);
+by (rtac trans 1);
+ by (rtac plus_assoc 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (simp_tac (simpset() addsimps
+        [plus_assoc,left_inv,left_inv2,right_inv,right_inv2]) 2);
+ by (rtac refl 2);
+by (rtac (zeroL RS sym) 1);
+qed "inv_plus";
+
+(*** convergent TRS for groups with unary inverse zero-x ***)
+val group1_simps =
+  [zeroL,zeroR,plus_assoc,left_inv,left_inv2,right_inv,right_inv2,inv_inv,
+   inv_zero,inv_plus];
+
+val group1_tac =
+  let val ss = HOL_basic_ss addsimps group1_simps
+  in simp_tac ss end;
+
+(* I believe there is no convergent TRS for groups with binary `-',
+   unless you have an extra unary `-' and simply define x-y = x+(-y).
+   This does not work with only a binary `-' because x-y = x+(zero-y) does
+   not terminate. Hence we have a special tactic for converting all
+   occurrences of x-y into x+(zero-y):
+*)
+
+local
+fun find(Const("op -",Type("fun",[T,_]))$s$t) = [(T,s,t)] @ find s @ find t
+  | find(s$t) = find s @ find t
+  | find _ = [];
+
+fun subst_tac sg (tacf,(T,s,t)) = 
+  let val typinst = [(("'a",0),ctyp_of sg T)];
+      val terminst = [(cterm_of sg (Var(("x",0),T)),cterm_of sg s),
+                      (cterm_of sg (Var(("y",0),T)),cterm_of sg t)];
+  in tacf THEN' rtac ((instantiate(typinst,terminst) minus_inv) RS ssubst) end;
+
+in
+val mk_group1_tac = SUBGOAL(fn (t,i) => fn st =>
+      let val sg = #sign(rep_thm st)
+      in foldl (subst_tac sg) (K all_tac,find t) i st
+      end)
+end;
+
+(* The following two equations are not used in any of the decision procedures,
+   but are still very useful. They also demonstrate mk_group1_tac.
+*)
+Goal "x-x = (zero::'a::add_group)";
+by (mk_group1_tac 1);
+by (group1_tac 1);
+qed "minus_self_zero";
+
+Goal "x-zero = (x::'a::add_group)";
+by (mk_group1_tac 1);
+by (group1_tac 1);
+qed "minus_zero";
+
+(*** Abelian Groups ***)
+
+Goal "x+(y+z)=y+(x+z::'a::add_agroup)";
+by (rtac trans 1);
+by (rtac plus_commute 1);
+by (rtac trans 1);
+by (rtac plus_assoc 1);
+by (simp_tac (simpset() addsimps [plus_commute]) 1);
+qed "plus_commuteL";
+
+(* Convergent TRS for Abelian groups with unary inverse zero-x.
+   Requires ordered rewriting
+*)
+
+val agroup1_simps = plus_commute::plus_commuteL::group1_simps;
+
+val agroup1_tac =
+  let val ss = HOL_basic_ss addsimps agroup1_simps
+  in simp_tac ss end;
+
+(* Again, I do not believe there is a convergent TRS for Abelian Groups with
+   binary `-'. However, we can still decide the word problem using additional
+   rules for 
+   1. floating `-' to the top:
+      "x + (y - z) = (x + y) - (z::'a::add_group)"
+      "(x - y) + z = (x + z) - (y::'a::add_agroup)"
+      "(x - y) - z = x - (y + (z::'a::add_agroup))"
+      "x - (y - z) = (x + z) - (y::'a::add_agroup)"
+   2. and for moving `-' over to the other side:
+      (x-y = z) = (x = z+y) and (x = y-z) = (x+z = y)
+*)
+Goal "x + (y - z) = (x + y) - (z::'a::add_group)";
+by (mk_group1_tac 1);
+by (group1_tac 1);
+qed "plus_minusR";
+
+Goal "(x - y) + z = (x + z) - (y::'a::add_agroup)";
+by (mk_group1_tac 1);
+by (agroup1_tac 1);
+qed "plus_minusL";
+
+Goal "(x - y) - z = x - (y + (z::'a::add_agroup))";
+by (mk_group1_tac 1);
+by (agroup1_tac 1);
+qed "minus_minusL";
+
+Goal "x - (y - z) = (x + z) - (y::'a::add_agroup)";
+by (mk_group1_tac 1);
+by (agroup1_tac 1);
+qed "minus_minusR";
+
+Goal "!!x::'a::add_group. (x-y = z) = (x = z+y)";
+by (stac minus_inv 1);
+by (fast_tac (claset() addss (HOL_basic_ss addsimps group1_simps)) 1);
+qed "minusL_iff";
+
+Goal "!!x::'a::add_group. (x = y-z) = (x+z = y)";
+by (stac minus_inv 1);
+by (fast_tac (claset() addss (HOL_basic_ss addsimps group1_simps)) 1);
+qed "minusR_iff";
+
+val agroup2_simps =
+  [zeroL,zeroR,plus_assoc,plus_commute,plus_commuteL,
+   plus_minusR,plus_minusL,minus_minusL,minus_minusR,minusL_iff,minusR_iff];
+
+(* This two-phase ordered rewriting tactic works, but agroup_simps are
+   simpler. However, agroup_simps is not confluent for arbitrary terms,
+   it merely decides equality.
+(* Phase 1 *)
+
+Goal "!!x::'a::add_agroup. (x+(zero-y))+z = (x+z)+(zero-y)";
+by (Simp_tac 1);
+val lemma = result();
+bind_thm("plus_minusL",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma);
+
+Goal "!!x::'a::add_agroup. x+(zero-(y+z)) = (x+(zero-y))+(zero-z)";
+by (Simp_tac 1);
+val lemma = result();
+bind_thm("minus_plusR",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma);
+
+Goal "!!x::'a::add_agroup. x+(zero-(y+(zero-z))) = (x+z)+(zero-y)";
+by (Simp_tac 1);
+val lemma = result();
+bind_thm("minus_minusR",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma);
+
+Goal "!!x::'a::add_agroup. x+(y+(zero-z)) = (x+y)+(zero-z)";
+by (Simp_tac 1);
+val lemma = result();
+bind_thm("plus_minusR",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma);
+
+(* Phase 2 *)
+
+Goal "!!x::'a::add_agroup. (x+y)+(zero-z) = x+(y+(zero-z))";
+by (Simp_tac 1);
+val lemma = result();
+bind_thm("minus_plusL2",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma);
+
+Goal "!!x::'a::add_agroup. (x+y)+(zero-x) = y";
+by (rtac (plus_assoc RS trans) 1);
+by (rtac trans 1);
+ by (rtac plus_cong 1);
+  by (rtac refl 1);
+ by (rtac right_inv2 2);
+by (rtac plus_commute 1);
+val lemma = result();
+bind_thm("minus_plusL3",rewrite_rule[minus_inv RS sym RS eq_reflection]lemma);
+
+*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Group.thy	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,44 @@
+(*  Title:      HOL/Integ/Group.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1996 TU Muenchen
+
+A little bit of group theory leading up to rings. Hence groups are additive.
+*)
+
+Group = Set +
+
+(* 0 already used in Nat *)
+axclass  zero < term
+consts   zero :: "'a::zero"
+
+(* additive semigroups *)
+
+axclass  add_semigroup < plus
+  plus_assoc   "(x + y) + z = x + (y + z)"
+
+
+(* additive monoids *)
+
+axclass  add_monoid < add_semigroup, zero
+  zeroL    "zero + x = x"
+  zeroR    "x + zero = x"
+
+(* additive groups *)
+(*
+The inverse is the binary `-'. Groups with unary and binary inverse are
+interdefinable: x-y := x+(zero-y) and -x := zero-x. The law left_inv is
+simply the translation of (-x)+x = zero. This characterizes groups already,
+provided we only allow (zero-x). Law minus_inv `defines' the general x-y in
+terms of the specific zero-y.
+*)
+axclass  add_group < add_monoid, minus
+  left_inv  "(zero-x)+x = zero"
+  minus_inv "x-y = x + (zero-y)"
+
+(* additive abelian groups *)
+
+axclass  add_agroup < add_group
+  plus_commute  "x + y = y + x"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/IntRing.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,18 @@
+(*  Title:      HOL/Integ/IntRing.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow and Markus Wenzel
+    Copyright   1996 TU Muenchen
+
+The instantiation of Lagrange's lemma for int.
+*)
+
+open IntRing;
+
+Goal "!!i1::int. \
+\  (sq i1 + sq i2 + sq i3 + sq i4) * (sq j1 + sq j2 + sq j3 + sq j4) = \
+\  sq(i1*j1 - i2*j2 - i3*j3 - i4*j4)  + \
+\  sq(i1*j2 + i2*j1 + i3*j4 - i4*j3)  + \
+\  sq(i1*j3 - i2*j4 + i3*j1 + i4*j2)  + \
+\  sq(i1*j4 + i2*j3 - i3*j2 + i4*j1)";
+by (rtac Lagrange_lemma 1);
+qed "Lagrange_lemma_int";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/IntRing.thy	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,19 @@
+(*  Title:      HOL/Integ/IntRing.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow and Markus Wenzel
+    Copyright   1996 TU Muenchen
+
+The integers form a commutative ring.
+With an application of Lagrange's lemma.
+*)
+
+IntRing = IntRingDefs + Lagrange +
+
+instance int :: add_semigroup (zadd_assoc)
+instance int :: add_monoid (zero_int_def,zadd_0,zadd_0_right)
+instance int :: add_group (left_inv_int,minus_inv_int)
+instance int :: add_agroup (zadd_commute)
+instance int :: ring (zmult_assoc,zadd_zmult_distrib2,zadd_zmult_distrib)
+instance int :: cring (zmult_commute)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/IntRingDefs.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,16 @@
+(*  Title:      HOL/Integ/IntRingDefs.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow and Markus Wenzel
+    Copyright   1996 TU Muenchen
+
+*)
+
+open IntRingDefs;
+
+Goalw [zero_int_def,zdiff_def] "(zero-x)+(x::int) = zero";
+by (Simp_tac 1);
+qed "left_inv_int";
+
+Goalw [zero_int_def,zdiff_def] "x-y = (x::int) + (zero-y)";
+by (Simp_tac 1);
+qed "minus_inv_int";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/IntRingDefs.thy	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,15 @@
+(*  Title:      HOL/Integ/IntRingDefs.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow and Markus Wenzel
+    Copyright   1996 TU Muenchen
+
+Provides the basic defs and thms for showing that int is a commutative ring.
+Most of it has been defined and shown already.
+*)
+
+IntRingDefs = Integ + Ring +
+
+instance int :: zero
+defs zero_int_def "zero::int == $# 0"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Lagrange.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,37 @@
+(*  Title:      HOL/Integ/Lagrange.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1996 TU Muenchen
+
+
+The following lemma essentially shows that all composite natural numbers are
+sums of fours squares, provided all prime numbers are. However, this is an
+abstract thm about commutative rings and has a priori nothing to do with nat.
+*)
+
+Goalw [Lagrange.sq_def] "!!x1::'a::cring. \
+\  (sq x1 + sq x2 + sq x3 + sq x4) * (sq y1 + sq y2 + sq y3 + sq y4) = \
+\  sq(x1*y1 - x2*y2 - x3*y3 - x4*y4)  + \
+\  sq(x1*y2 + x2*y1 + x3*y4 - x4*y3)  + \
+\  sq(x1*y3 - x2*y4 + x3*y1 + x4*y2)  + \
+\  sq(x1*y4 + x2*y3 - x3*y2 + x4*y1)";
+(*Takes up to three minutes...*)
+by (cring_tac 1);
+qed "Lagrange_lemma";
+
+(* A challenge by John Harrison.
+   Takes forever because of the naive bottom-up strategy of the rewriter.
+
+Goalw [Lagrange.sq_def] "!!p1::'a::cring.\
+\ (sq p1 + sq q1 + sq r1 + sq s1 + sq t1 + sq u1 + sq v1 + sq w1) * \
+\ (sq p2 + sq q2 + sq r2 + sq s2 + sq t2 + sq u2 + sq v2 + sq w2) \
+\  = sq (p1*p2 - q1*q2 - r1*r2 - s1*s2 - t1*t2 - u1*u2 - v1*v2 - w1*w2) + \
+\    sq (p1*q2 + q1*p2 + r1*s2 - s1*r2 + t1*u2 - u1*t2 - v1*w2 + w1*v2) +\
+\    sq (p1*r2 - q1*s2 + r1*p2 + s1*q2 + t1*v2 + u1*w2 - v1*t2 - w1*u2) +\
+\    sq (p1*s2 + q1*r2 - r1*q2 + s1*p2 + t1*w2 - u1*v2 + v1*u2 - w1*t2) +\
+\    sq (p1*t2 - q1*u2 - r1*v2 - s1*w2 + t1*p2 + u1*q2 + v1*r2 + w1*s2) +\
+\    sq (p1*u2 + q1*t2 - r1*w2 + s1*v2 - t1*q2 + u1*p2 - v1*s2 + w1*r2) +\
+\    sq (p1*v2 + q1*w2 + r1*t2 - s1*u2 - t1*r2 + u1*s2 + v1*p2 - w1*q2) +\
+\    sq (p1*w2 - q1*v2 + r1*u2 + s1*t2 - t1*s2 - u1*r2 + v1*q2 + w1*p2)";
+
+*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Lagrange.thy	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,18 @@
+(*  Title:      HOL/Integ/Lagrange.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1996 TU Muenchen
+
+
+This theory only contains a single thm, which is a lemma in Lagrange's proof
+that every natural number is the sum of 4 squares.  It's sole purpose is to
+demonstrate ordered rewriting for commutative rings.
+
+The enterprising reader might consider proving all of Lagrange's thm.
+*)
+Lagrange = Ring +
+
+constdefs sq :: 'a::times => 'a
+         "sq x == x*x"
+
+end
--- a/src/HOL/ex/Primrec.ML	Wed Jun 24 13:59:45 1998 +0200
+++ b/src/HOL/ex/Primrec.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -128,7 +128,7 @@
 val lemma = result();
 
 Goal "!!i j k. i<j ==> ack(i,k) < ack(j,k)";
-by (etac less_natE 1);
+by (dtac less_eq_Suc_add 1);
 by (blast_tac (claset() addSIs [lemma]) 1);
 qed "ack_less_mono1";
 
--- a/src/HOL/ex/ROOT.ML	Wed Jun 24 13:59:45 1998 +0200
+++ b/src/HOL/ex/ROOT.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -28,6 +28,8 @@
 time_use_thy "Qsort";
 time_use_thy "Puzzle";
 
+time_use_thy "IntRing";
+
 time_use     "set.ML";
 time_use_thy "MT";
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Ring.ML	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,139 @@
+(*  Title:      HOL/Integ/Ring.ML
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1996 TU Muenchen
+
+Derives a few equational consequences about rings
+and defines cring_simpl, a simplification tactic for commutative rings.
+*)
+
+Goal "!!x::'a::cring. x*(y*z)=y*(x*z)";
+by (rtac trans 1);
+by (rtac times_commute 1);
+by (rtac trans 1);
+by (rtac times_assoc 1);
+by (simp_tac (HOL_basic_ss addsimps [times_commute]) 1);
+qed "times_commuteL";
+
+val times_cong = read_instantiate [("f1","op *")] (arg_cong RS cong);
+
+Goal "!!x::'a::ring. zero*x = zero";
+by (rtac trans 1);
+ by (rtac right_inv 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (rtac refl 3);
+ by (rtac trans 2);
+  by (rtac times_cong 3);
+   by (rtac zeroL 3);
+  by (rtac refl 3);
+ by (rtac (distribR RS sym) 2);
+by (rtac trans 1);
+ by (rtac (plus_assoc RS sym) 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (rtac refl 2);
+ by (rtac (right_inv RS sym) 2);
+by (rtac (zeroR RS sym) 1);
+qed "mult_zeroL";
+
+Goal "!!x::'a::ring. x*zero = zero";
+by (rtac trans 1);
+ by (rtac right_inv 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (rtac refl 3);
+ by (rtac trans 2);
+  by (rtac times_cong 3);
+   by (rtac zeroL 4);
+  by (rtac refl 3);
+ by (rtac (distribL RS sym) 2);
+by (rtac trans 1);
+ by (rtac (plus_assoc RS sym) 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (rtac refl 2);
+ by (rtac (right_inv RS sym) 2);
+by (rtac (zeroR RS sym) 1);
+qed "mult_zeroR";
+
+Goal "!!x::'a::ring. (zero-x)*y = zero-(x*y)";
+by (rtac trans 1);
+ by (rtac zeroL 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (rtac refl 3);
+ by (rtac mult_zeroL 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (rtac refl 3);
+ by (rtac times_cong 2);
+  by (rtac left_inv 2);
+ by (rtac refl 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (rtac refl 3);
+ by (rtac (distribR RS sym) 2);
+by (rtac trans 1);
+ by (rtac (plus_assoc RS sym) 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (rtac refl 2);
+ by (rtac (right_inv RS sym) 2);
+by (rtac (zeroR RS sym) 1);
+qed "mult_invL";
+
+Goal "!!x::'a::ring. x*(zero-y) = zero-(x*y)";
+by (rtac trans 1);
+ by (rtac zeroL 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (rtac refl 3);
+ by (rtac mult_zeroR 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (rtac refl 3);
+ by (rtac times_cong 2);
+  by (rtac refl 2);
+ by (rtac left_inv 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (rtac refl 3);
+ by (rtac (distribL RS sym) 2);
+by (rtac trans 1);
+ by (rtac (plus_assoc RS sym) 2);
+by (rtac trans 1);
+ by (rtac plus_cong 2);
+  by (rtac refl 2);
+ by (rtac (right_inv RS sym) 2);
+by (rtac (zeroR RS sym) 1);
+qed "mult_invR";
+
+Goal "x*(y-z) = (x*y - x*z::'a::ring)";
+by (mk_group1_tac 1);
+by (simp_tac (HOL_basic_ss addsimps [distribL,mult_invR]) 1);
+qed "minus_distribL";
+
+Goal "(x-y)*z = (x*z - y*z::'a::ring)";
+by (mk_group1_tac 1);
+by (simp_tac (HOL_basic_ss addsimps [distribR,mult_invL]) 1);
+qed "minus_distribR";
+
+val cring_simps = [times_assoc,times_commute,times_commuteL,
+                   distribL,distribR,minus_distribL,minus_distribR]
+                  @ agroup2_simps;
+
+val cring_tac =
+  let val ss = HOL_basic_ss addsimps cring_simps
+  in simp_tac ss end;
+
+
+(*** The order [minus_plusL3,minus_plusL2] is important because minus_plusL3
+     MUST be tried first
+val cring_simp =
+  let val phase1 = simpset() addsimps
+                   [plus_minusL,minus_plusR,minus_minusR,plus_minusR]
+      val phase2 = HOL_ss addsimps [minus_plusL3,minus_plusL2,
+                                    zeroL,zeroR,mult_zeroL,mult_zeroR]
+  in simp_tac phase1 THEN' simp_tac phase2 end;
+***)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Ring.thy	Thu Jun 25 13:57:34 1998 +0200
@@ -0,0 +1,24 @@
+(*  Title:      HOL/Integ/Ring.thy
+    ID:         $Id$
+    Author:     Tobias Nipkow
+    Copyright   1996 TU Muenchen
+
+Bits of rings.
+Main output: a simplification tactic for commutative rings.
+*)
+
+Ring = Group +
+
+(* Ring *)
+
+axclass  ring < add_agroup, times
+  times_assoc "(x*y)*z = x*(y*z)"
+  distribL    "x*(y+z) = x*y + x*z"
+  distribR    "(x+y)*z = x*z + y*z"
+
+(* Commutative ring *)
+
+axclass cring < ring
+  times_commute "x*y = y*x"
+
+end