src/HOL/Real/float.ML
changeset 22964 2284e0d02e7f
parent 22963 509b1da3cee1
child 22965 b81bbe298406
equal deleted inserted replaced
22963:509b1da3cee1 22964:2284e0d02e7f
     1 (*  Title: HOL/Real/Float.ML
       
     2     ID:    $Id$
       
     3     Author: Steven Obua
       
     4 *)
       
     5 
       
     6 structure ExactFloatingPoint :
       
     7 sig
       
     8     exception Destruct_floatstr of string
       
     9     val destruct_floatstr : (char -> bool) -> (char -> bool) -> string -> bool * string * string * bool * string
       
    10 
       
    11     exception Floating_point of string
       
    12 
       
    13     type floatrep = IntInf.int * IntInf.int
       
    14     val approx_dec_by_bin : IntInf.int -> floatrep -> floatrep * floatrep
       
    15     val approx_decstr_by_bin : int -> string -> floatrep * floatrep
       
    16 end
       
    17 =
       
    18 struct
       
    19 
       
    20 exception Destruct_floatstr of string;
       
    21 
       
    22 fun destruct_floatstr isDigit isExp number =
       
    23     let
       
    24         val numlist = filter (not o Char.isSpace) (String.explode number)
       
    25 
       
    26         fun countsigns ((#"+")::cs) = countsigns cs
       
    27           | countsigns ((#"-")::cs) =
       
    28             let
       
    29                 val (positive, rest) = countsigns cs
       
    30             in
       
    31                 (not positive, rest)
       
    32             end
       
    33           | countsigns cs = (true, cs)
       
    34 
       
    35         fun readdigits [] = ([], [])
       
    36           | readdigits (q as c::cs) =
       
    37             if (isDigit c) then
       
    38                 let
       
    39                     val (digits, rest) = readdigits cs
       
    40                 in
       
    41                     (c::digits, rest)
       
    42                 end
       
    43             else
       
    44                 ([], q)
       
    45 
       
    46         fun readfromexp_helper cs =
       
    47             let
       
    48                 val (positive, rest) = countsigns cs
       
    49                 val (digits, rest') = readdigits rest
       
    50             in
       
    51                 case rest' of
       
    52                     [] => (positive, digits)
       
    53                   | _ => raise (Destruct_floatstr number)
       
    54             end
       
    55 
       
    56         fun readfromexp [] = (true, [])
       
    57           | readfromexp (c::cs) =
       
    58             if isExp c then
       
    59                 readfromexp_helper cs
       
    60             else
       
    61                 raise (Destruct_floatstr number)
       
    62 
       
    63         fun readfromdot [] = ([], readfromexp [])
       
    64           | readfromdot ((#".")::cs) =
       
    65             let
       
    66                 val (digits, rest) = readdigits cs
       
    67                 val exp = readfromexp rest
       
    68             in
       
    69                 (digits, exp)
       
    70             end
       
    71           | readfromdot cs = readfromdot ((#".")::cs)
       
    72 
       
    73         val (positive, numlist) = countsigns numlist
       
    74         val (digits1, numlist) = readdigits numlist
       
    75          val (digits2, exp) = readfromdot numlist
       
    76     in
       
    77         (positive, String.implode digits1, String.implode digits2, fst exp, String.implode (snd exp))
       
    78     end
       
    79 
       
    80 type floatrep = IntInf.int * IntInf.int
       
    81 
       
    82 exception Floating_point of string;
       
    83 
       
    84 val ln2_10 = (Math.ln 10.0)/(Math.ln 2.0)
       
    85 
       
    86 fun intmul a b = IntInf.* (a,b)
       
    87 fun intsub a b = IntInf.- (a,b)
       
    88 fun intadd a b = IntInf.+ (a,b) 
       
    89 fun intpow a b = IntInf.pow (a, IntInf.toInt b);
       
    90 fun intle a b = IntInf.<= (a, b);
       
    91 fun intless a b = IntInf.< (a, b);
       
    92 fun intneg a = IntInf.~ a;
       
    93 val zero = IntInf.fromInt 0;
       
    94 val one = IntInf.fromInt 1;
       
    95 val two = IntInf.fromInt 2;
       
    96 val ten = IntInf.fromInt 10;
       
    97 val five = IntInf.fromInt 5;
       
    98 
       
    99 fun find_most_significant q r =
       
   100     let
       
   101         fun int2real i =
       
   102             case Real.fromString (IntInf.toString i) of
       
   103                 SOME r => r
       
   104               | NONE => raise (Floating_point "int2real")
       
   105         fun subtract (q, r) (q', r') =
       
   106             if intle r r' then
       
   107                 (intsub q (intmul q' (intpow ten (intsub r' r))), r)
       
   108             else
       
   109                 (intsub (intmul q (intpow ten (intsub r r'))) q', r')
       
   110         fun bin2dec d =
       
   111             if intle zero d then
       
   112                 (intpow two d, zero)
       
   113             else
       
   114                 (intpow five (intneg d), d)
       
   115 
       
   116         val L = IntInf.fromInt (Real.floor (int2real (IntInf.fromInt (IntInf.log2 q)) + (int2real r) * ln2_10))
       
   117         val L1 = intadd L one
       
   118 
       
   119         val (q1, r1) = subtract (q, r) (bin2dec L1) 
       
   120     in
       
   121         if intle zero q1 then
       
   122             let
       
   123                 val (q2, r2) = subtract (q, r) (bin2dec (intadd L1 one))
       
   124             in
       
   125                 if intle zero q2 then
       
   126                     raise (Floating_point "find_most_significant")
       
   127                 else
       
   128                     (L1, (q1, r1))
       
   129             end
       
   130         else
       
   131             let
       
   132                 val (q0, r0) = subtract (q, r) (bin2dec L)
       
   133             in
       
   134                 if intle zero q0 then
       
   135                     (L, (q0, r0))
       
   136                 else
       
   137                     raise (Floating_point "find_most_significant")
       
   138             end
       
   139     end
       
   140 
       
   141 fun approx_dec_by_bin n (q,r) =
       
   142     let
       
   143         fun addseq acc d' [] = acc
       
   144           | addseq acc d' (d::ds) = addseq (intadd acc (intpow two (intsub d d'))) d' ds
       
   145 
       
   146         fun seq2bin [] = (zero, zero)
       
   147           | seq2bin (d::ds) = (intadd (addseq zero d ds) one, d)
       
   148 
       
   149         fun approx d_seq d0 precision (q,r) =
       
   150             if q = zero then
       
   151                 let val x = seq2bin d_seq in
       
   152                     (x, x)
       
   153                 end
       
   154             else
       
   155                 let
       
   156                     val (d, (q', r')) = find_most_significant q r
       
   157                 in
       
   158                     if intless precision (intsub d0 d) then
       
   159                         let
       
   160                             val d' = intsub d0 precision
       
   161                             val x1 = seq2bin (d_seq)
       
   162                             val x2 = (intadd (intmul (fst x1) (intpow two (intsub (snd x1) d'))) one,  d') (* = seq2bin (d'::d_seq) *)
       
   163                         in
       
   164                             (x1, x2)
       
   165                         end
       
   166                     else
       
   167                         approx (d::d_seq) d0 precision (q', r')
       
   168                 end
       
   169 
       
   170         fun approx_start precision (q, r) =
       
   171             if q = zero then
       
   172                 ((zero, zero), (zero, zero))
       
   173             else
       
   174                 let
       
   175                     val (d, (q', r')) = find_most_significant q r
       
   176                 in
       
   177                     if intle precision zero then
       
   178                         let
       
   179                             val x1 = seq2bin [d]
       
   180                         in
       
   181                             if q' = zero then
       
   182                                 (x1, x1)
       
   183                             else
       
   184                                 (x1, seq2bin [intadd d one])
       
   185                         end
       
   186                     else
       
   187                         approx [d] d precision (q', r')
       
   188                 end
       
   189     in
       
   190         if intle zero q then
       
   191             approx_start n (q,r)
       
   192         else
       
   193             let
       
   194                 val ((a1,b1), (a2, b2)) = approx_start n (intneg q, r)
       
   195             in
       
   196                 ((intneg a2, b2), (intneg a1, b1))
       
   197             end
       
   198     end
       
   199 
       
   200 fun approx_decstr_by_bin n decstr =
       
   201     let
       
   202         fun str2int s = case IntInf.fromString s of SOME x => x | NONE => zero
       
   203         fun signint p x = if p then x else intneg x
       
   204 
       
   205         val (p, d1, d2, ep, e) = destruct_floatstr Char.isDigit (fn e => e = #"e" orelse e = #"E") decstr
       
   206         val s = IntInf.fromInt (size d2)
       
   207 
       
   208         val q = signint p (intadd (intmul (str2int d1) (intpow ten s)) (str2int d2))
       
   209         val r = intsub (signint ep (str2int e)) s
       
   210     in
       
   211         approx_dec_by_bin (IntInf.fromInt n) (q,r)
       
   212     end
       
   213 
       
   214 end;
       
   215 
       
   216 structure FloatArith =
       
   217 struct
       
   218 
       
   219 type float = IntInf.int * IntInf.int
       
   220 
       
   221 val izero = IntInf.fromInt 0
       
   222 val ione = IntInf.fromInt 1
       
   223 val imone = IntInf.fromInt ~1
       
   224 val itwo = IntInf.fromInt 2
       
   225 fun imul a b = IntInf.* (a,b)
       
   226 fun isub a b = IntInf.- (a,b)
       
   227 fun iadd a b = IntInf.+ (a,b)
       
   228 
       
   229 val floatzero = (izero, izero)
       
   230 
       
   231 fun positive_part (a,b) =
       
   232     (if IntInf.< (a,izero) then izero else a, b)
       
   233 
       
   234 fun negative_part (a,b) =
       
   235     (if IntInf.< (a,izero) then a else izero, b)
       
   236 
       
   237 fun is_negative (a,b) =
       
   238     if IntInf.< (a, izero) then true else false
       
   239 
       
   240 fun is_positive (a,b) =
       
   241     if IntInf.< (izero, a) then true else false
       
   242 
       
   243 fun is_zero (a,b) =
       
   244     if a = izero then true else false
       
   245 
       
   246 fun ipow2 a = IntInf.pow ((IntInf.fromInt 2), IntInf.toInt a)
       
   247 
       
   248 fun add (a1, b1) (a2, b2) =
       
   249     if IntInf.< (b1, b2) then
       
   250         (iadd a1 (imul a2 (ipow2 (isub b2 b1))), b1)
       
   251     else
       
   252         (iadd (imul a1 (ipow2 (isub b1 b2))) a2, b2)
       
   253 
       
   254 fun sub (a1, b1) (a2, b2) =
       
   255     if IntInf.< (b1, b2) then
       
   256         (isub a1 (imul a2 (ipow2 (isub b2 b1))), b1)
       
   257     else
       
   258         (isub (imul a1 (ipow2 (isub b1 b2))) a2, b2)
       
   259 
       
   260 fun neg (a, b) = (IntInf.~ a, b)
       
   261 
       
   262 fun is_equal a b = is_zero (sub a b)
       
   263 
       
   264 fun is_less a b = is_negative (sub a b)
       
   265 
       
   266 fun max a b = if is_less a b then b else a
       
   267 
       
   268 fun min a b = if is_less a b then a else b
       
   269 
       
   270 fun abs a = if is_negative a then neg a else a
       
   271 
       
   272 fun mul (a1, b1) (a2, b2) = (imul a1 a2, iadd b1 b2)
       
   273 
       
   274 end;
       
   275 
       
   276 
       
   277 structure Float:
       
   278 sig
       
   279     type float = FloatArith.float
       
   280     type floatfunc = float * float -> float * float
       
   281 
       
   282     val mk_intinf : typ -> IntInf.int -> term
       
   283     val mk_float : float -> term
       
   284 
       
   285     exception Dest_intinf;
       
   286     val dest_intinf : term -> IntInf.int
       
   287     val dest_nat : term -> IntInf.int
       
   288 
       
   289     exception Dest_float;
       
   290     val dest_float : term -> float
       
   291 
       
   292     val float_const : term
       
   293 
       
   294     val float_add_const : term
       
   295     val float_diff_const : term
       
   296     val float_uminus_const : term
       
   297     val float_pprt_const : term
       
   298     val float_nprt_const : term
       
   299     val float_abs_const : term
       
   300     val float_mult_const : term
       
   301     val float_le_const : term
       
   302 
       
   303     val nat_le_const : term
       
   304     val nat_less_const : term
       
   305     val nat_eq_const : term
       
   306 
       
   307     val approx_float : int -> floatfunc -> string -> term * term
       
   308 
       
   309 end
       
   310 =
       
   311 struct
       
   312 
       
   313 structure Inttab = TableFun(type key = int val ord = (rev_order o int_ord));
       
   314 
       
   315 type float = IntInf.int*IntInf.int
       
   316 type floatfunc = float*float -> float*float
       
   317 
       
   318 val float_const = Const ("Float.float", HOLogic.mk_prodT (HOLogic.intT, HOLogic.intT) --> HOLogic.realT)
       
   319 
       
   320 val float_add_const = Const ("HOL.plus", HOLogic.realT --> HOLogic.realT --> HOLogic.realT)
       
   321 val float_diff_const = Const ("HOL.minus", HOLogic.realT --> HOLogic.realT --> HOLogic.realT)
       
   322 val float_mult_const = Const ("HOL.times", HOLogic.realT --> HOLogic.realT --> HOLogic.realT)
       
   323 val float_uminus_const = Const ("HOL.uminus", HOLogic.realT --> HOLogic.realT)
       
   324 val float_abs_const = Const ("HOL.abs", HOLogic.realT --> HOLogic.realT)
       
   325 val float_le_const = Const ("Orderings.less_eq", HOLogic.realT --> HOLogic.realT --> HOLogic.boolT)
       
   326 val float_pprt_const = Const ("OrderedGroup.pprt", HOLogic.realT --> HOLogic.realT)
       
   327 val float_nprt_const = Const ("OrderedGroup.nprt", HOLogic.realT --> HOLogic.realT)
       
   328 
       
   329 val nat_le_const = Const ("Orderings.less_eq", HOLogic.natT --> HOLogic.natT --> HOLogic.boolT)
       
   330 val nat_less_const = Const ("Orderings.less", HOLogic.natT --> HOLogic.natT --> HOLogic.boolT)
       
   331 val nat_eq_const = Const ("op =", HOLogic.natT --> HOLogic.natT --> HOLogic.boolT)
       
   332  
       
   333 val zero = FloatArith.izero
       
   334 val minus_one = FloatArith.imone
       
   335 val two = FloatArith.itwo
       
   336 
       
   337 exception Dest_intinf;
       
   338 exception Dest_float;
       
   339 
       
   340 fun mk_intinf ty n = HOLogic.number_of_const ty $ HOLogic.mk_numeral n;
       
   341 
       
   342 val dest_intinf = snd o HOLogic.dest_number
       
   343 
       
   344 fun mk_float (a,b) =
       
   345     float_const $ (HOLogic.mk_prod ((mk_intinf HOLogic.intT a), (mk_intinf HOLogic.intT b)))
       
   346 
       
   347 fun dest_float f =
       
   348     case f of
       
   349         (Const ("Float.float", _) $ (Const ("Pair", _) $ a $ b)) => (dest_intinf a, dest_intinf b)
       
   350       | Const ("Numeral.number_of",_) $ a => (dest_intinf f, 0)
       
   351       | Const ("Numeral0", _) => (FloatArith.izero, FloatArith.izero)
       
   352       | Const ("Numeral1", _) => (FloatArith.ione, FloatArith.izero)
       
   353       | _ => raise Dest_float
       
   354 
       
   355 fun dest_nat n =
       
   356     let
       
   357         val v = dest_intinf n
       
   358     in
       
   359         if IntInf.< (v, FloatArith.izero) then
       
   360             FloatArith.izero
       
   361         else
       
   362             v
       
   363     end
       
   364 
       
   365 fun approx_float prec f value =
       
   366     let
       
   367         val interval = ExactFloatingPoint.approx_decstr_by_bin prec value
       
   368         val (flower, fupper) = f interval
       
   369     in
       
   370         (mk_float flower, mk_float fupper)
       
   371     end
       
   372 
       
   373 end;