src/Pure/General/rat.ML
author chaieb
Mon, 04 Jun 2007 11:38:34 +0200
changeset 23231 aef7b4e5c8fe
parent 23063 b4ee6ec4f9c6
child 23232 861ab9c18e18
permissions -rw-r--r--
opaque-constraint removed
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
     1
(*  Title:      Pure/General/rat.ML
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
     2
    ID:         $Id$
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
     3
    Author:     Tobias Nipkow, TU Muenchen
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
     4
17940
3706c254d296 added rounding functions
haftmann
parents: 17848
diff changeset
     5
Canonical implementation of exact rational numbers.
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
     6
*)
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
     7
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
     8
signature RAT =
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
     9
sig
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    10
  type rat
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    11
  exception DIVZERO
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    12
  val zero: rat
19557
4866ebb16ba8 Rat.one added
webertj
parents: 17950
diff changeset
    13
  val one: rat
22950
haftmann
parents: 22574
diff changeset
    14
  val two: rat
haftmann
parents: 22574
diff changeset
    15
  val rat_of_int: Intt.int -> rat
haftmann
parents: 22574
diff changeset
    16
  val rat_of_quotient: Intt.int * Intt.int -> rat
haftmann
parents: 22574
diff changeset
    17
  val quotient_of_rat: rat -> Intt.int * Intt.int
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    18
  val string_of_rat: rat -> string
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    19
  val eq: rat * rat -> bool
22950
haftmann
parents: 22574
diff changeset
    20
  val cmp: rat * rat -> order
haftmann
parents: 22574
diff changeset
    21
  val le: rat -> rat -> bool
23036
65b4f545a76f added lt and some other infix operation analogous to Ocaml's num library
chaieb
parents: 22950
diff changeset
    22
  val lt: rat -> rat -> bool
22950
haftmann
parents: 22574
diff changeset
    23
  val cmp_zero: rat -> order
haftmann
parents: 22574
diff changeset
    24
  val add: rat -> rat -> rat
haftmann
parents: 22574
diff changeset
    25
  val mult: rat -> rat -> rat
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    26
  val neg: rat -> rat
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    27
  val inv: rat -> rat
17940
3706c254d296 added rounding functions
haftmann
parents: 17848
diff changeset
    28
  val roundup: rat -> rat
3706c254d296 added rounding functions
haftmann
parents: 17848
diff changeset
    29
  val rounddown: rat -> rat
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    30
end;
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    31
23231
aef7b4e5c8fe opaque-constraint removed
chaieb
parents: 23063
diff changeset
    32
structure Rat : RAT =
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    33
struct
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    34
22950
haftmann
parents: 22574
diff changeset
    35
datatype rat = Rat of bool * Intt.int * Intt.int;
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    36
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    37
exception DIVZERO;
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    38
22950
haftmann
parents: 22574
diff changeset
    39
val zero = Rat (true, Intt.int 0, Intt.int 1);
haftmann
parents: 22574
diff changeset
    40
val one = Rat (true, Intt.int 1, Intt.int 1);
haftmann
parents: 22574
diff changeset
    41
val two = Rat (true, Intt.int 2, Intt.int 1);
19557
4866ebb16ba8 Rat.one added
webertj
parents: 17950
diff changeset
    42
22950
haftmann
parents: 22574
diff changeset
    43
fun rat_of_int i =
haftmann
parents: 22574
diff changeset
    44
  if i < Intt.int 0
haftmann
parents: 22574
diff changeset
    45
  then Rat (false, ~i, Intt.int 1)
haftmann
parents: 22574
diff changeset
    46
  else Rat (true, i, Intt.int 1);
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    47
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
    48
fun norm (a, p, q) =
22950
haftmann
parents: 22574
diff changeset
    49
  if p = Intt.int 0 then Rat (true, Intt.int 0, Intt.int 1)
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
    50
  else
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
    51
    let
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
    52
      val absp = abs p
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
    53
      val m = gcd (absp, q)
22950
haftmann
parents: 22574
diff changeset
    54
    in Rat(a = (Intt.int 0 <= p), absp div m, q div m) end;
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    55
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    56
fun common (p1, q1, p2, q2) =
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    57
  let val q' = lcm (q1, q2)
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    58
  in (p1 * (q' div q1), p2 * (q' div q2), q') end
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    59
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
    60
fun rat_of_quotient (p, q) =
22950
haftmann
parents: 22574
diff changeset
    61
  if q = Intt.int 0 then raise DIVZERO
haftmann
parents: 22574
diff changeset
    62
  else norm (Intt.int 0 <= q, p, abs q);
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    63
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    64
fun quotient_of_rat (Rat (a, p, q)) = (if a then p else ~p, q);
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    65
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    66
fun string_of_rat r =
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    67
  let val (p, q) = quotient_of_rat r
22950
haftmann
parents: 22574
diff changeset
    68
  in Intt.string_of_int p ^ "/" ^ Intt.string_of_int q end;
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    69
17940
3706c254d296 added rounding functions
haftmann
parents: 17848
diff changeset
    70
fun eq (Rat (false, _, _), Rat (true, _, _)) = false
3706c254d296 added rounding functions
haftmann
parents: 17848
diff changeset
    71
  | eq (Rat (true, _, _), Rat (false, _, _)) = false
3706c254d296 added rounding functions
haftmann
parents: 17848
diff changeset
    72
  | eq (Rat (_, p1, q1), Rat (_, p2, q2)) = p1 = p2 andalso q1 = q2
3706c254d296 added rounding functions
haftmann
parents: 17848
diff changeset
    73
22950
haftmann
parents: 22574
diff changeset
    74
fun cmp (Rat (false, _, _), Rat (true, _, _)) = LESS
haftmann
parents: 22574
diff changeset
    75
  | cmp (Rat (true, _, _), Rat (false, _, _)) = GREATER
haftmann
parents: 22574
diff changeset
    76
  | cmp (Rat (a, p1, q1), Rat (_, p2, q2)) =
17940
3706c254d296 added rounding functions
haftmann
parents: 17848
diff changeset
    77
      let val (r1, r2, _) = common (p1, q1, p2, q2)
22950
haftmann
parents: 22574
diff changeset
    78
      in if a then Intt.cmp (r1, r2) else Intt.cmp (r2, r1) end;
17940
3706c254d296 added rounding functions
haftmann
parents: 17848
diff changeset
    79
23063
haftmann
parents: 23036
diff changeset
    80
fun le a b = let val order = cmp (a, b) in order = LESS orelse order = EQUAL end;
haftmann
parents: 23036
diff changeset
    81
fun lt a b = cmp (a, b) = LESS;
haftmann
parents: 23036
diff changeset
    82
22950
haftmann
parents: 22574
diff changeset
    83
fun cmp_zero (Rat (false, _, _)) = LESS
haftmann
parents: 22574
diff changeset
    84
  | cmp_zero (Rat (_, 0, _)) = EQUAL
haftmann
parents: 22574
diff changeset
    85
  | cmp_zero (Rat (_, _, _)) = GREATER;
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    86
22950
haftmann
parents: 22574
diff changeset
    87
fun add (Rat (a1, p1, q1)) (Rat(a2, p2, q2)) =
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    88
  let
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    89
    val (r1, r2, den) = common (p1, q1, p2, q2)
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    90
    val num = (if a1 then r1 else ~r1) + (if a2 then r2 else ~r2)
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    91
  in norm (true, num, den) end;
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    92
22950
haftmann
parents: 22574
diff changeset
    93
fun mult (Rat (a1, p1, q1)) (Rat (a2, p2, q2)) =
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    94
  norm (a1=a2, p1*p2, q1*q2);
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    95
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
    96
fun neg (r as Rat (b, p, q)) =
22950
haftmann
parents: 22574
diff changeset
    97
  if p = Intt.int 0 then r
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
    98
  else Rat (not b, p, q);
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
    99
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
   100
fun inv (Rat (a, p, q)) =
22950
haftmann
parents: 22574
diff changeset
   101
  if p = Intt.int 0 then raise DIVZERO
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
   102
  else Rat (a, q, p);
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
   103
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
   104
fun roundup (r as Rat (a, p, q)) =
22950
haftmann
parents: 22574
diff changeset
   105
  if q = Intt.int 1 then r
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
   106
  else
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
   107
    let
22950
haftmann
parents: 22574
diff changeset
   108
      fun round true q = Rat (true, q + Intt.int 1, Intt.int 1)
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
   109
        | round false q =
22950
haftmann
parents: 22574
diff changeset
   110
            if q = Intt.int 0
haftmann
parents: 22574
diff changeset
   111
            then Rat (true, Intt.int 0, Intt.int 1)
haftmann
parents: 22574
diff changeset
   112
            else Rat (false, q, Intt.int 1);
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
   113
    in round a (p div q) end;
17940
3706c254d296 added rounding functions
haftmann
parents: 17848
diff changeset
   114
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
   115
fun rounddown (r as Rat (a, p, q)) =
22950
haftmann
parents: 22574
diff changeset
   116
  if q = Intt.int 1 then r
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
   117
  else
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
   118
    let
22950
haftmann
parents: 22574
diff changeset
   119
      fun round true q = Rat (true, q, Intt.int 1)
haftmann
parents: 22574
diff changeset
   120
        | round false q = Rat (false, q + Intt.int 1, Intt.int 1)
22574
e6c25fd3de2a avoid overloaded integer constants (accomodate Alice);
wenzelm
parents: 22189
diff changeset
   121
    in round a (p div q) end;
17940
3706c254d296 added rounding functions
haftmann
parents: 17848
diff changeset
   122
17848
de5d9d5e99f5 added module rat.ML for rational numbers
haftmann
parents:
diff changeset
   123
end;
23036
65b4f545a76f added lt and some other infix operation analogous to Ocaml's num library
chaieb
parents: 22950
diff changeset
   124
23231
aef7b4e5c8fe opaque-constraint removed
chaieb
parents: 23063
diff changeset
   125
infix 5 +/; (*FIXME infix 5?*)
23063
haftmann
parents: 23036
diff changeset
   126
infix 5 -/;
haftmann
parents: 23036
diff changeset
   127
infix 7 */;
23231
aef7b4e5c8fe opaque-constraint removed
chaieb
parents: 23063
diff changeset
   128
infix 7 //; (*FIXME infix 7?*)
23063
haftmann
parents: 23036
diff changeset
   129
infix 4 =/ </ <=/ >/ >=/ <>/;
23036
65b4f545a76f added lt and some other infix operation analogous to Ocaml's num library
chaieb
parents: 22950
diff changeset
   130
65b4f545a76f added lt and some other infix operation analogous to Ocaml's num library
chaieb
parents: 22950
diff changeset
   131
fun a +/ b = Rat.add a b;
23063
haftmann
parents: 23036
diff changeset
   132
fun a -/ b = a +/ Rat.neg b;
haftmann
parents: 23036
diff changeset
   133
fun a */ b = Rat.mult a b;
haftmann
parents: 23036
diff changeset
   134
fun a // b = a */ Rat.inv b; 
23036
65b4f545a76f added lt and some other infix operation analogous to Ocaml's num library
chaieb
parents: 22950
diff changeset
   135
fun a =/ b = Rat.eq (a,b);
65b4f545a76f added lt and some other infix operation analogous to Ocaml's num library
chaieb
parents: 22950
diff changeset
   136
fun a </ b = Rat.lt a b;
65b4f545a76f added lt and some other infix operation analogous to Ocaml's num library
chaieb
parents: 22950
diff changeset
   137
fun a <=/ b = Rat.le a b;
65b4f545a76f added lt and some other infix operation analogous to Ocaml's num library
chaieb
parents: 22950
diff changeset
   138
fun a >/ b = b </ a;
65b4f545a76f added lt and some other infix operation analogous to Ocaml's num library
chaieb
parents: 22950
diff changeset
   139
fun a >=/ b = b <=/ a;
65b4f545a76f added lt and some other infix operation analogous to Ocaml's num library
chaieb
parents: 22950
diff changeset
   140
fun a <>/ b = not (a =/ b);