src/Tools/rat.ML
changeset 23297 06f108974fa1
parent 23261 85f27f79232f
child 23520 483fe92f00c1
--- a/src/Tools/rat.ML	Fri Jun 08 18:13:58 2007 +0200
+++ b/src/Tools/rat.ML	Sat Jun 09 00:28:46 2007 +0200
@@ -7,7 +7,7 @@
 
 signature RAT =
 sig
-  type rat
+  eqtype rat
   exception DIVZERO
   val zero: rat
   val one: rat
@@ -36,37 +36,37 @@
 
 exception DIVZERO;
 
-val zero = Rat (true, Integer.zero, Integer.one);
-val one = Rat (true, Integer.one, Integer.one);
-val two = Rat (true, Integer.two, Integer.one);
+val zero = Rat (true, 0, 1);
+val one = Rat (true, 1, 1);
+val two = Rat (true, 2, 1);
 
 fun rat_of_int i =
   let
     val (a, p) = Integer.signabs i
-  in Rat (a, p, Integer.one) end;
+  in Rat (a, p, 1) end;
 
 fun norm (a, p, q) =
-  if Integer.cmp_zero p = EQUAL then Rat (true, Integer.zero, Integer.one)
+  if p = 0 then Rat (true, 0, 1)
   else
     let
       val (b, absp) = Integer.signabs p;
       val m = Integer.gcd absp q;
-    in Rat (a = b, Integer.div absp m, Integer.div q m) end;
+    in Rat (a = b, absp div m, q div m) end;
 
 fun common (p1, q1, p2, q2) =
   let
     val q' = Integer.lcm q1 q2;
-  in (p1 *% (Integer.div q' q1), p2 *% (Integer.div q' q2), q') end
+  in (p1 * (q' div q1), p2 * (q' div q2), q') end
 
 fun rat_of_quotient (p, q) =
   let
     val (a, absq) = Integer.signabs q;
   in
-    if Integer.cmp_zero absq = EQUAL then raise DIVZERO
+    if absq = 0 then raise DIVZERO
     else norm (a, p, absq)
   end;
 
-fun quotient_of_rat (Rat (a, p, q)) = (if a then p else Integer.neg p, q);
+fun quotient_of_rat (Rat (a, p, q)) = (if a then p else ~ p, q);
 
 fun string_of_rat r =
   let
@@ -75,7 +75,7 @@
 
 fun eq (Rat (false, _, _), Rat (true, _, _)) = false
   | eq (Rat (true, _, _), Rat (false, _, _)) = false
-  | eq (Rat (_, p1, q1), Rat (_, p2, q2)) = p1 =% p2 andalso q1 =% q2;
+  | eq (Rat (_, p1, q1), Rat (_, p2, q2)) = p1 = p2 andalso q1 = q2;
 
 fun cmp (Rat (false, _, _), Rat (true, _, _)) = LESS
   | cmp (Rat (true, _, _), Rat (false, _, _)) = GREATER
@@ -93,48 +93,48 @@
 fun add (Rat (a1, p1, q1)) (Rat(a2, p2, q2)) =
   let
     val (r1, r2, den) = common (p1, q1, p2, q2);
-    val num = (if a1 then r1 else Integer.neg r1)
-      +% (if a2 then r2 else Integer.neg r2);
+    val num = (if a1 then r1 else ~ r1)
+      + (if a2 then r2 else ~ r2);
   in norm (true, num, den) end;
 
 fun mult (Rat (a1, p1, q1)) (Rat (a2, p2, q2)) =
-  norm (a1 = a2, p1 *% p2, q1 *% q2);
+  norm (a1 = a2, p1 * p2, q1 * q2);
 
 fun neg (r as Rat (b, p, q)) =
-  if Integer.cmp_zero p = EQUAL then r
+  if p = 0 then r
   else Rat (not b, p, q);
 
 fun inv (Rat (a, p, q)) =
-  if Integer.cmp_zero q = EQUAL then raise DIVZERO
+  if q = 0 then raise DIVZERO
   else Rat (a, q, p);
 
 fun roundup (r as Rat (a, p, q)) =
-  if q = Integer.one then r
+  if q = 1 then r
   else
     let
-      fun round true q = Rat (true, Integer.inc q, Integer.one)
+      fun round true q = Rat (true, q + 1, 1)
         | round false q =
-            Rat (Integer.cmp_zero q = EQUAL, Integer.int 0, Integer.int 1);
-    in round a (Integer.div p q) end;
+            Rat (q = 0, 0, 1);
+    in round a (p div q) end;
 
 fun rounddown (r as Rat (a, p, q)) =
-  if q = Integer.one then r
+  if q = 1 then r
   else
     let
-      fun round true q = Rat (true, q, Integer.one)
-        | round false q = Rat (false, Integer.inc q, Integer.one)
-    in round a (Integer.div p q) end;
+      fun round true q = Rat (true, q, 1)
+        | round false q = Rat (false, q + 1, 1)
+    in round a (p div q) end;
 
 end;
 
 infix 7 */ //;
-infix 6 +/ -/; 
+infix 6 +/ -/;
 infix 4 =/ </ <=/ >/ >=/ <>/;
 
 fun a +/ b = Rat.add a b;
 fun a -/ b = a +/ Rat.neg b;
 fun a */ b = Rat.mult a b;
-fun a // b = a */ Rat.inv b; 
+fun a // b = a */ Rat.inv b;
 fun a =/ b = Rat.eq (a, b);
 fun a </ b = Rat.lt a b;
 fun a <=/ b = Rat.le a b;