src/HOL/Real/float_arith.ML
changeset 23261 85f27f79232f
parent 22964 2284e0d02e7f
child 23297 06f108974fa1
--- a/src/HOL/Real/float_arith.ML	Tue Jun 05 18:36:10 2007 +0200
+++ b/src/HOL/Real/float_arith.ML	Tue Jun 05 19:19:30 2007 +0200
@@ -9,7 +9,7 @@
   val destruct_floatstr: (char -> bool) -> (char -> bool) -> string -> bool * string * string * bool * string
 
   exception Floating_point of string
-  val approx_dec_by_bin: Intt.int -> Float.float -> Float.float * Float.float
+  val approx_dec_by_bin: integer -> Float.float -> Float.float * Float.float
   val approx_decstr_by_bin: int -> string -> Float.float * Float.float
 
   val mk_float: Float.float -> term
@@ -85,39 +85,39 @@
 exception Floating_point of string;
 
 val ln2_10 = Math.ln 10.0 / Math.ln 2.0;
-val exp5 = Intt.pow (Intt.int 5);
-val exp10 = Intt.pow (Intt.int 10);
+val exp5 = Integer.pow (Integer.int 5);
+val exp10 = Integer.pow (Integer.int 10);
 
-fun intle a b = not (Intt.cmp (a, b) = GREATER);
-fun intless a b = Intt.cmp (a, b) = LESS;
+fun intle a b = not (Integer.cmp (a, b) = GREATER);
+fun intless a b = Integer.cmp (a, b) = LESS;
 
 fun find_most_significant q r =
   let
     fun int2real i =
-      case (Real.fromString o Intt.string_of_int) i of
+      case (Real.fromString o Integer.string_of_int) i of
         SOME r => r
         | NONE => raise (Floating_point "int2real")
     fun subtract (q, r) (q', r') =
       if intle r r' then
-        (Intt.sub q (Intt.mult q' (exp10 (Intt.sub r' r))), r)
+        (Integer.sub q (Integer.mult q' (exp10 (Integer.sub r' r))), r)
       else
-        (Intt.sub (Intt.mult q (exp10 (Intt.sub r r'))) q', r')
+        (Integer.sub (Integer.mult q (exp10 (Integer.sub r r'))) q', r')
     fun bin2dec d =
-      if intle Intt.zero d then
-        (Intt.exp d, Intt.zero)
+      if intle Integer.zero d then
+        (Integer.exp d, Integer.zero)
       else
-        (exp5 (Intt.neg d), d)
+        (exp5 (Integer.neg d), d)
 
-    val L = Intt.int (Real.floor (int2real (Intt.log q) + int2real r * ln2_10))
-    val L1 = Intt.inc L
+    val L = Integer.int (Real.floor (int2real (Integer.log q) + int2real r * ln2_10))
+    val L1 = Integer.inc L
 
     val (q1, r1) = subtract (q, r) (bin2dec L1) 
   in
-    if intle Intt.zero q1 then
+    if intle Integer.zero q1 then
       let
-        val (q2, r2) = subtract (q, r) (bin2dec (Intt.inc L1))
+        val (q2, r2) = subtract (q, r) (bin2dec (Integer.inc L1))
       in
-        if intle Intt.zero q2 then
+        if intle Integer.zero q2 then
           raise (Floating_point "find_most_significant")
         else
           (L1, (q1, r1))
@@ -126,7 +126,7 @@
       let
         val (q0, r0) = subtract (q, r) (bin2dec L)
       in
-        if intle Intt.zero q0 then
+        if intle Integer.zero q0 then
           (L, (q0, r0))
         else
           raise (Floating_point "find_most_significant")
@@ -136,13 +136,13 @@
 fun approx_dec_by_bin n (q,r) =
   let
     fun addseq acc d' [] = acc
-      | addseq acc d' (d::ds) = addseq (Intt.add acc (Intt.exp (Intt.sub d d'))) d' ds
+      | addseq acc d' (d::ds) = addseq (Integer.add acc (Integer.exp (Integer.sub d d'))) d' ds
 
-    fun seq2bin [] = (Intt.zero, Intt.zero)
-      | seq2bin (d::ds) = (Intt.inc (addseq Intt.zero d ds), d)
+    fun seq2bin [] = (Integer.zero, Integer.zero)
+      | seq2bin (d::ds) = (Integer.inc (addseq Integer.zero d ds), d)
 
     fun approx d_seq d0 precision (q,r) =
-      if q = Intt.zero then
+      if q = Integer.zero then
         let val x = seq2bin d_seq in
           (x, x)
         end
@@ -150,13 +150,13 @@
         let
           val (d, (q', r')) = find_most_significant q r
         in
-          if intless precision (Intt.sub d0 d) then
+          if intless precision (Integer.sub d0 d) then
             let
-              val d' = Intt.sub d0 precision
+              val d' = Integer.sub d0 precision
               val x1 = seq2bin (d_seq)
-              val x2 = (Intt.inc
-                (Intt.mult (fst x1)
-                (Intt.exp (Intt.sub (snd x1) d'))),  d') (* = seq2bin (d'::d_seq) *)
+              val x2 = (Integer.inc
+                (Integer.mult (fst x1)
+                (Integer.exp (Integer.sub (snd x1) d'))),  d') (* = seq2bin (d'::d_seq) *)
             in
               (x1, x2)
             end
@@ -165,47 +165,47 @@
         end
 
     fun approx_start precision (q, r) =
-      if q = Intt.zero then
-        ((Intt.zero, Intt.zero), (Intt.zero, Intt.zero))
+      if q = Integer.zero then
+        ((Integer.zero, Integer.zero), (Integer.zero, Integer.zero))
       else
         let
           val (d, (q', r')) = find_most_significant q r
         in
-          if intle precision Intt.zero then
+          if intle precision Integer.zero then
             let
               val x1 = seq2bin [d]
             in
-              if q' = Intt.zero then
+              if q' = Integer.zero then
                 (x1, x1)
               else
-                (x1, seq2bin [Intt.inc d])
+                (x1, seq2bin [Integer.inc d])
             end
           else
             approx [d] d precision (q', r')
         end
   in
-    if intle Intt.zero q then
+    if intle Integer.zero q then
       approx_start n (q,r)
     else
       let
-        val ((a1,b1), (a2, b2)) = approx_start n (Intt.neg q, r)
+        val ((a1,b1), (a2, b2)) = approx_start n (Integer.neg q, r)
       in
-        ((Intt.neg a2, b2), (Intt.neg a1, b1))
+        ((Integer.neg a2, b2), (Integer.neg a1, b1))
       end
   end
 
 fun approx_decstr_by_bin n decstr =
   let
-    fun str2int s = the_default Intt.zero (Intt.int_of_string s);
-    fun signint p x = if p then x else Intt.neg x
+    fun str2int s = the_default Integer.zero (Integer.int_of_string s);
+    fun signint p x = if p then x else Integer.neg x
 
     val (p, d1, d2, ep, e) = destruct_floatstr Char.isDigit (fn e => e = #"e" orelse e = #"E") decstr
-    val s = Intt.int (size d2)
+    val s = Integer.int (size d2)
 
-    val q = signint p (Intt.add (Intt.mult (str2int d1) (exp10 s)) (str2int d2))
-    val r = Intt.sub (signint ep (str2int e)) s
+    val q = signint p (Integer.add (Integer.mult (str2int d1) (exp10 s)) (str2int d2))
+    val r = Integer.sub (signint ep (str2int e)) s
   in
-    approx_dec_by_bin (Intt.int n) (q,r)
+    approx_dec_by_bin (Integer.int n) (q,r)
   end
 
 fun mk_float (a, b) = @{term "float"} $
@@ -213,7 +213,7 @@
 
 fun dest_float (Const ("Float.float", _) $ (Const ("Pair", _) $ a $ b)) =
       pairself (snd o HOLogic.dest_number) (a, b)
-  | dest_float t = ((snd o HOLogic.dest_number) t, Intt.zero);
+  | dest_float t = ((snd o HOLogic.dest_number) t, Integer.zero);
 
 fun approx_float prec f value =
   let