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