| author | wenzelm | 
| Fri, 05 May 2023 12:01:09 +0200 | |
| changeset 77967 | 6bb2f9b32804 | 
| parent 70586 | 57df8a85317a | 
| child 80809 | 4a64fc4d1cde | 
| permissions | -rw-r--r-- | 
| 63197 
af562e976038
rat.ML is now part of Pure to allow tigther integration with Isabelle/ML;
 wenzelm parents: 
30161diff
changeset | 1 | (* Title: Pure/General/rat.ML | 
| 23520 | 2 | Author: Tobias Nipkow, Florian Haftmann, TU Muenchen | 
| 63207 | 3 | Author: Makarius | 
| 23251 | 4 | |
| 5 | Canonical implementation of exact rational numbers. | |
| 6 | *) | |
| 7 | ||
| 8 | signature RAT = | |
| 9 | sig | |
| 23297 | 10 | eqtype rat | 
| 63201 | 11 | val of_int: int -> rat | 
| 12 | val make: int * int -> rat | |
| 13 | val dest: rat -> int * int | |
| 23251 | 14 | val string_of_rat: rat -> string | 
| 63208 | 15 | val signed_string_of_rat: rat -> string | 
| 70586 | 16 | val ord: rat ord | 
| 23251 | 17 | val le: rat -> rat -> bool | 
| 18 | val lt: rat -> rat -> bool | |
| 23520 | 19 | val sign: rat -> order | 
| 20 | val abs: rat -> rat | |
| 23251 | 21 | val add: rat -> rat -> rat | 
| 22 | val mult: rat -> rat -> rat | |
| 23 | val neg: rat -> rat | |
| 24 | val inv: rat -> rat | |
| 63206 | 25 | val floor: rat -> int | 
| 26 | val ceil: rat -> int | |
| 23251 | 27 | end; | 
| 28 | ||
| 29 | structure Rat : RAT = | |
| 30 | struct | |
| 31 | ||
| 63202 | 32 | datatype rat = Rat of int * int; (*numerator, positive (!) denominator*) | 
| 63199 
da38571dd5bd
prefer more efficient Poly/ML operations, taking care of sign;
 wenzelm parents: 
63198diff
changeset | 33 | |
| 63217 | 34 | fun of_int i = Rat (i, 1); | 
| 35 | ||
| 23520 | 36 | fun common (p1, q1) (p2, q2) = | 
| 63227 
d3ed7f00e818
Integer.lcm normalizes the sign as in HOL/GCD.thy;
 wenzelm parents: 
63218diff
changeset | 37 | let val m = Integer.lcm q1 q2 | 
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 38 | in ((p1 * (m div q1), p2 * (m div q2)), m) end; | 
| 23520 | 39 | |
| 63211 | 40 | fun make (_, 0) = raise Div | 
| 63210 
a0685d2b420b
clarified exception -- actually reject denominator = 0;
 wenzelm parents: 
63209diff
changeset | 41 | | make (p, q) = | 
| 
a0685d2b420b
clarified exception -- actually reject denominator = 0;
 wenzelm parents: 
63209diff
changeset | 42 | let | 
| 63227 
d3ed7f00e818
Integer.lcm normalizes the sign as in HOL/GCD.thy;
 wenzelm parents: 
63218diff
changeset | 43 | val m = Integer.gcd p q; | 
| 63210 
a0685d2b420b
clarified exception -- actually reject denominator = 0;
 wenzelm parents: 
63209diff
changeset | 44 | val (p', q') = (p div m, q div m); | 
| 
a0685d2b420b
clarified exception -- actually reject denominator = 0;
 wenzelm parents: 
63209diff
changeset | 45 | in Rat (if q' < 0 then (~ p', ~ q') else (p', q')) end | 
| 23520 | 46 | |
| 63201 | 47 | fun dest (Rat r) = r; | 
| 23520 | 48 | |
| 63208 | 49 | fun string_of_rat (Rat (p, 1)) = string_of_int p | 
| 50 | | string_of_rat (Rat (p, q)) = string_of_int p ^ "/" ^ string_of_int q; | |
| 51 | ||
| 52 | fun signed_string_of_rat (Rat (p, 1)) = signed_string_of_int p | |
| 53 | | signed_string_of_rat (Rat (p, q)) = signed_string_of_int p ^ "/" ^ string_of_int q; | |
| 23520 | 54 | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 55 | fun ord (Rat (p1, q1), Rat (p2, q2)) = | 
| 63202 | 56 | (case (Integer.sign p1, Integer.sign p2) of | 
| 57 | (LESS, EQUAL) => LESS | |
| 23520 | 58 | | (LESS, GREATER) => LESS | 
| 59 | | (EQUAL, LESS) => GREATER | |
| 60 | | (EQUAL, EQUAL) => EQUAL | |
| 61 | | (EQUAL, GREATER) => LESS | |
| 62 | | (GREATER, LESS) => GREATER | |
| 63 | | (GREATER, EQUAL) => GREATER | |
| 63202 | 64 | | _ => int_ord (fst (common (p1, q1) (p2, q2)))); | 
| 23251 | 65 | |
| 63211 | 66 | fun le a b = ord (a, b) <> GREATER; | 
| 63202 | 67 | fun lt a b = ord (a, b) = LESS; | 
| 23251 | 68 | |
| 23520 | 69 | fun sign (Rat (p, _)) = Integer.sign p; | 
| 23251 | 70 | |
| 63211 | 71 | fun abs (r as Rat (p, q)) = if p < 0 then Rat (~ p, q) else r; | 
| 23251 | 72 | |
| 63217 | 73 | fun add (Rat r1) (Rat r2) = | 
| 74 | let val ((m1, m2), n) = common r1 r2 | |
| 63201 | 75 | in make (m1 + m2, n) end; | 
| 23251 | 76 | |
| 63202 | 77 | fun mult (Rat (p1, q1)) (Rat (p2, q2)) = make (p1 * p2, q1 * q2); | 
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 78 | |
| 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 79 | fun neg (Rat (p, q)) = Rat (~ p, q); | 
| 23251 | 80 | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24584diff
changeset | 81 | fun inv (Rat (p, q)) = | 
| 63202 | 82 | (case Integer.sign p of | 
| 83 | LESS => Rat (~ q, ~ p) | |
| 63210 
a0685d2b420b
clarified exception -- actually reject denominator = 0;
 wenzelm parents: 
63209diff
changeset | 84 | | EQUAL => raise Div | 
| 63202 | 85 | | GREATER => Rat (q, p)); | 
| 23251 | 86 | |
| 63206 | 87 | fun floor (Rat (p, q)) = p div q; | 
| 23520 | 88 | |
| 63206 | 89 | fun ceil (Rat (p, q)) = | 
| 63202 | 90 | (case Integer.div_mod p q of | 
| 63206 | 91 | (m, 0) => m | 
| 63218 | 92 | | (m, _) => m + 1); | 
| 23251 | 93 | |
| 94 | end; | |
| 95 | ||
| 63198 
c583ca33076a
ad-hoc overloading for standard operations on type Rat.rat;
 wenzelm parents: 
63197diff
changeset | 96 | ML_system_overload (uncurry Rat.add) "+"; | 
| 
c583ca33076a
ad-hoc overloading for standard operations on type Rat.rat;
 wenzelm parents: 
63197diff
changeset | 97 | ML_system_overload (fn (a, b) => Rat.add a (Rat.neg b)) "-"; | 
| 
c583ca33076a
ad-hoc overloading for standard operations on type Rat.rat;
 wenzelm parents: 
63197diff
changeset | 98 | ML_system_overload (uncurry Rat.mult) "*"; | 
| 
c583ca33076a
ad-hoc overloading for standard operations on type Rat.rat;
 wenzelm parents: 
63197diff
changeset | 99 | ML_system_overload (fn (a, b) => Rat.mult a (Rat.inv b)) "/"; | 
| 
c583ca33076a
ad-hoc overloading for standard operations on type Rat.rat;
 wenzelm parents: 
63197diff
changeset | 100 | ML_system_overload (uncurry Rat.lt) "<"; | 
| 
c583ca33076a
ad-hoc overloading for standard operations on type Rat.rat;
 wenzelm parents: 
63197diff
changeset | 101 | ML_system_overload (uncurry Rat.le) "<="; | 
| 
c583ca33076a
ad-hoc overloading for standard operations on type Rat.rat;
 wenzelm parents: 
63197diff
changeset | 102 | ML_system_overload (fn (a, b) => Rat.lt b a) ">"; | 
| 
c583ca33076a
ad-hoc overloading for standard operations on type Rat.rat;
 wenzelm parents: 
63197diff
changeset | 103 | ML_system_overload (fn (a, b) => Rat.le b a) ">="; | 
| 63211 | 104 | ML_system_overload Rat.neg "~"; | 
| 105 | ML_system_overload Rat.abs "abs"; | |
| 63209 | 106 | |
| 107 | ML_system_pp (fn _ => fn _ => fn x => Pretty.to_polyml (Pretty.str ("@" ^ Rat.string_of_rat x)));
 |