src/Pure/General/float.ML
author haftmann
Mon, 14 May 2007 12:52:56 +0200
changeset 22964 2284e0d02e7f
permissions -rw-r--r--
reorganized float arithmetic
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22964
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
     1
(*  Title:      Pure/General/float.ML
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
     2
    ID:         $Id$
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
     3
    Author:     Steven Obua, Florian Haftmann, TU Muenchen
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
     4
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
     5
Implementation of real numbers as mantisse-exponent pairs.
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
     6
*)
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
     7
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
     8
signature FLOAT =
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
     9
sig
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    10
  type float = Intt.int * Intt.int
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    11
  val zero: float
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    12
  val eq: float * float -> bool
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    13
  val cmp: float * float -> order
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    14
  val cmp_zero: float -> order
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    15
  val min: float -> float -> float
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    16
  val max: float -> float -> float
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    17
  val add: float -> float -> float
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    18
  val sub: float -> float -> float
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    19
  val neg: float -> float
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    20
  val mult: float -> float -> float
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    21
  val positive_part: float -> float
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    22
  val negative_part: float -> float
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    23
end;
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    24
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    25
structure Float : FLOAT =
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    26
struct
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    27
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    28
type float = Intt.int * Intt.int;
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    29
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    30
val zero = (Intt.zero, Intt.zero);
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    31
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    32
fun add (a1, b1) (a2, b2) =
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    33
  if Intt.cmp (b1, b2) = LESS then
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    34
    (Intt.add a1 (Intt.mult a2 (Intt.exp (Intt.sub b2 b1))), b1)
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    35
  else
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    36
    (Intt.add (Intt.mult a1 (Intt.exp (Intt.sub b1 b2))) a2, b2);
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    37
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    38
fun sub (a1, b1) (a2, b2) =
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    39
  if Intt.cmp (b1, b2) = LESS then
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    40
    (Intt.sub a1 (Intt.mult a2 (Intt.exp (Intt.sub b2 b1))), b1)
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    41
  else
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    42
    (Intt.sub (Intt.mult a1 (Intt.exp (Intt.sub b1 b2))) a2, b2);
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    43
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    44
fun neg (a, b) = (Intt.neg a, b);
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    45
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    46
fun mult (a1, b1) (a2, b2) = (Intt.mult a1 a2, Intt.add b1 b2);
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    47
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    48
fun cmp_zero (a, b) = Intt.cmp_zero a;
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    49
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    50
fun cmp (r, s) = cmp_zero (sub r s);
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    51
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    52
fun eq (r, s) = cmp (r, s) = EQUAL;
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    53
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    54
fun min r s = case cmp (r, s) of LESS => r | _ => s;
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    55
fun max r s = case cmp (r, s) of LESS => s | _ => r;
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    56
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    57
fun positive_part (a, b) = (Intt.max Intt.zero a, b);
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    58
fun negative_part (a, b) = (Intt.min Intt.zero a, b);
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    59
2284e0d02e7f reorganized float arithmetic
haftmann
parents:
diff changeset
    60
end;