# HG changeset patch # User haftmann # Date 1179139976 -7200 # Node ID 2284e0d02e7f951062a48a36b2fa4472b565eb2d # Parent 509b1da3cee18d0f021773ba34a957dd1b12f5d7 reorganized float arithmetic diff -r 509b1da3cee1 -r 2284e0d02e7f src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon May 14 12:52:54 2007 +0200 +++ b/src/HOL/IsaMakefile Mon May 14 12:52:56 2007 +0200 @@ -153,11 +153,12 @@ HOL-Complex: HOL $(OUT)/HOL-Complex -$(OUT)/HOL-Complex: $(OUT)/HOL Complex/ROOT.ML Library/Zorn.thy \ - Real/ContNotDenum.thy Real/Ferrante_Rackoff.thy Real/float.ML \ +$(OUT)/HOL-Complex: $(OUT)/HOL Complex/ROOT.ML $(SRC)/Pure/General/float.ML \ + Library/Zorn.thy \ + Real/ContNotDenum.thy Real/Ferrante_Rackoff.thy Real/float_arith.ML \ Real/Float.thy Real/Lubs.thy Real/PReal.thy Real/RComplete.thy \ Real/ROOT.ML Real/Rational.thy Real/Real.thy Real/RealDef.thy \ - Real/RealPow.thy Real/RealVector.thy Real/ferrante_rackoff_proof.ML \ + Real/RealPow.thy Real/RealVector.thy Real/ferrante_rackoff_proof.ML \ Real/ferrante_rackoff.ML Real/rat_arith.ML Real/real_arith.ML \ Hyperreal/StarDef.thy Hyperreal/StarClasses.thy \ Hyperreal/EvenOdd.thy Hyperreal/Fact.thy Hyperreal/HLog.thy \ diff -r 509b1da3cee1 -r 2284e0d02e7f src/HOL/Matrix/cplex/FloatSparseMatrixBuilder.ML --- a/src/HOL/Matrix/cplex/FloatSparseMatrixBuilder.ML Mon May 14 12:52:54 2007 +0200 +++ b/src/HOL/Matrix/cplex/FloatSparseMatrixBuilder.ML Mon May 14 12:52:56 2007 +0200 @@ -3,182 +3,96 @@ Author: Steven Obua *) -structure FloatSparseMatrixBuilder : +signature FLOAT_SPARSE_MATIRX_BUILDER = sig - include MATRIX_BUILDER - - structure cplex : CPLEX + include MATRIX_BUILDER - type float = IntInf.int*IntInf.int - type floatfunc = float -> float - + structure cplex : CPLEX - val float2cterm : IntInf.int * IntInf.int -> cterm - - val approx_value : int -> floatfunc -> string -> cterm * cterm - val approx_vector : int -> floatfunc -> vector -> cterm * cterm - val approx_matrix : int -> floatfunc -> matrix -> cterm * cterm + type float = Float.float + val approx_value : int -> (float -> float) -> string -> term * term + val approx_vector : int -> (float -> float) -> vector -> term * term + val approx_matrix : int -> (float -> float) -> matrix -> term * term - val mk_spvec_entry : int -> float -> term - val empty_spvec : term - val cons_spvec : term -> term -> term - val empty_spmat : term - val mk_spmat_entry : int -> term -> term - val cons_spmat : term -> term -> term - val sign_term : term -> cterm - - val v_elem_at : vector -> int -> string option - val m_elem_at : matrix -> int -> vector option - val v_only_elem : vector -> int option - val v_fold : (int * string -> 'a -> 'a) -> vector -> 'a -> 'a - val m_fold : (int * vector -> 'a -> 'a) -> matrix -> 'a -> 'a + val mk_spvec_entry : Intt.int -> float -> term + val mk_spmat_entry : Intt.int -> term -> term + val spvecT: typ + val spmatT: typ + + val v_elem_at : vector -> int -> string option + val m_elem_at : matrix -> int -> vector option + val v_only_elem : vector -> int option + val v_fold : (int * string -> 'a -> 'a) -> vector -> 'a -> 'a + val m_fold : (int * vector -> 'a -> 'a) -> matrix -> 'a -> 'a - val transpose_matrix : matrix -> matrix + val transpose_matrix : matrix -> matrix - val cut_vector : int -> vector -> vector - val cut_matrix : vector -> (int option) -> matrix -> matrix + val cut_vector : int -> vector -> vector + val cut_matrix : vector -> int option -> matrix -> matrix - (* cplexProg c A b *) - val cplexProg : vector -> matrix -> vector -> (cplex.cplexProg * (string -> int)) - (* dual_cplexProg c A b *) - val dual_cplexProg : vector -> matrix -> vector -> (cplex.cplexProg * (string -> int)) + (* cplexProg c A b *) + val cplexProg : vector -> matrix -> vector -> cplex.cplexProg * (string -> int) + (* dual_cplexProg c A b *) + val dual_cplexProg : vector -> matrix -> vector -> cplex.cplexProg * (string -> int) +end; - val real_spmatT : typ - val real_spvecT : typ -end -= +structure FloatSparseMatrixBuilder : FLOAT_SPARSE_MATIRX_BUILDER = struct - -structure Inttab = TableFun(type key = int val ord = (rev_order o int_ord)); +type float = Float.float +structure Inttab = TableFun(type key = int val ord = rev_order o int_ord); type vector = string Inttab.table type matrix = vector Inttab.table -type float = IntInf.int*IntInf.int -type floatfunc = float -> float - -val th = theory "FloatSparseMatrix" - -fun readtype s = Sign.intern_type th s -fun readterm s = Sign.intern_const th s - -val ty_list = readtype "list" -val term_Nil = readterm "Nil" -val term_Cons = readterm "Cons" - -val spvec_elemT = HOLogic.mk_prodT (HOLogic.natT, HOLogic.realT) -val spvecT = Type (ty_list, [spvec_elemT]) -val spmat_elemT = HOLogic.mk_prodT (HOLogic.natT, spvecT) -val spmatT = Type (ty_list, [spmat_elemT]) - -val real_spmatT = spmatT -val real_spvecT = spvecT - -val empty_matrix_const = Const (term_Nil, spmatT) -val empty_vector_const = Const (term_Nil, spvecT) - -val Cons_spvec_const = Const (term_Cons, spvec_elemT --> spvecT --> spvecT) -val Cons_spmat_const = Const (term_Cons, spmat_elemT --> spmatT --> spmatT) -val float_const = Float.float_const - -val zero = IntInf.fromInt 0 -val minus_one = IntInf.fromInt ~1 -val two = IntInf.fromInt 2 - -val mk_intinf = Float.mk_intinf - -val mk_float = Float.mk_float - -fun float2cterm (a,b) = cterm_of th (mk_float (a,b)) +val spvec_elemT = HOLogic.mk_prodT (HOLogic.natT, HOLogic.realT); +val spvecT = HOLogic.listT spvec_elemT; +val spmat_elemT = HOLogic.mk_prodT (HOLogic.natT, spvecT); +val spmatT = HOLogic.listT spmat_elemT; -fun approx_value_term prec f = Float.approx_float prec (fn (x, y) => (f x, f y)) - -fun approx_value prec pprt value = - let - val (flower, fupper) = approx_value_term prec pprt value - in - (cterm_of th flower, cterm_of th fupper) - end - -fun sign_term t = cterm_of th t - -val empty_spvec = empty_vector_const - -val empty_spmat = empty_matrix_const +fun approx_value prec f = + FloatArith.approx_float prec (fn (x, y) => (f x, f y)); fun mk_spvec_entry i f = - let - val term_i = mk_intinf HOLogic.natT (IntInf.fromInt i) - val term_f = mk_float f - in - HOLogic.mk_prod (term_i, term_f) - end + HOLogic.mk_prod (HOLogic.mk_number HOLogic.natT i, FloatArith.mk_float f); fun mk_spmat_entry i e = - let - val term_i = mk_intinf HOLogic.natT (IntInf.fromInt i) - in - HOLogic.mk_prod (term_i, e) - end - -fun cons_spvec h t = Cons_spvec_const $ h $ t - -fun cons_spmat h t = Cons_spmat_const $ h $ t - -fun approx_vector_term prec pprt vector = - let - fun app (index, s) (vlower, vupper) = - let - val (flower, fupper) = approx_value_term prec pprt s - val index = mk_intinf HOLogic.natT (IntInf.fromInt index) - val elower = HOLogic.mk_prod (index, flower) - val eupper = HOLogic.mk_prod (index, fupper) - in - (Cons_spvec_const $ elower $ vlower, - Cons_spvec_const $ eupper $ vupper) - end - in - Inttab.fold app vector (empty_vector_const, empty_vector_const) - end - -fun approx_matrix_term prec pprt matrix = - let - fun app (index, vector) (mlower, mupper) = - let - val (vlower, vupper) = approx_vector_term prec pprt vector - val index = mk_intinf HOLogic.natT (IntInf.fromInt index) - val elower = HOLogic.mk_prod (index, vlower) - val eupper = HOLogic.mk_prod (index, vupper) - in - (Cons_spmat_const $ elower $ mlower, Cons_spmat_const $ eupper $ mupper) - end - val (mlower, mupper) = Inttab.fold app matrix (empty_matrix_const, empty_matrix_const) - in Inttab.fold app matrix (empty_matrix_const, empty_matrix_const) end; + HOLogic.mk_prod (HOLogic.mk_number HOLogic.natT i, e); fun approx_vector prec pprt vector = - let - val (l, u) = approx_vector_term prec pprt vector - in - (cterm_of th l, cterm_of th u) - end + let + fun app (index, s) (lower, upper) = + let + val (flower, fupper) = approx_value prec pprt s + val index = HOLogic.mk_number HOLogic.natT (Intt.int index) + val elower = HOLogic.mk_prod (index, flower) + val eupper = HOLogic.mk_prod (index, fupper) + in (elower :: lower, eupper :: upper) end; + in + pairself (HOLogic.mk_list spvecT) (Inttab.fold app vector ([], [])) + end; -fun approx_matrix prec pprt matrix = - let - val (l, u) = approx_matrix_term prec pprt matrix - in - (cterm_of th l, cterm_of th u) - end - +fun approx_matrix prec pprt vector = + let + fun app (index, v) (lower, upper) = + let + val (flower, fupper) = approx_vector prec pprt v + val index = HOLogic.mk_number HOLogic.natT (Intt.int index) + val elower = HOLogic.mk_prod (index, flower) + val eupper = HOLogic.mk_prod (index, fupper) + in (elower :: lower, eupper :: upper) end; + in + pairself (HOLogic.mk_list spmatT) (Inttab.fold app vector ([], [])) + end; exception Nat_expected of int; -val zero_interval = approx_value_term 1 I "0" +val zero_interval = approx_value 1 I "0" fun set_elem vector index str = if index < 0 then raise (Nat_expected index) - else if (approx_value_term 1 I str) = zero_interval then + else if (approx_value 1 I str) = zero_interval then vector else Inttab.update (index, str) vector diff -r 509b1da3cee1 -r 2284e0d02e7f src/HOL/Matrix/cplex/fspmlp.ML --- a/src/HOL/Matrix/cplex/fspmlp.ML Mon May 14 12:52:54 2007 +0200 +++ b/src/HOL/Matrix/cplex/fspmlp.ML Mon May 14 12:52:56 2007 +0200 @@ -9,25 +9,25 @@ type vector = FloatSparseMatrixBuilder.vector type matrix = FloatSparseMatrixBuilder.matrix - val y : linprog -> cterm - val A : linprog -> cterm * cterm - val b : linprog -> cterm - val c : linprog -> cterm * cterm - val r : linprog -> cterm - val r12 : linprog -> cterm * cterm + val y : linprog -> term + val A : linprog -> term * term + val b : linprog -> term + val c : linprog -> term * term + val r : linprog -> term + val r12 : linprog -> term * term exception Load of string val load : string -> int -> bool -> linprog end -structure fspmlp : FSPMLP = +structure Fspmlp : FSPMLP = struct type vector = FloatSparseMatrixBuilder.vector type matrix = FloatSparseMatrixBuilder.matrix -type linprog = cterm * (cterm * cterm) * cterm * (cterm * cterm) * cterm * (cterm * cterm) +type linprog = term * (term * term) * term * (term * term) * term * (term * term) fun y (c1, c2, c3, c4, c5, _) = c1 fun A (c1, c2, c3, c4, c5, _) = c2 @@ -77,8 +77,8 @@ | SOME (NONE, f) => (SOME bound, f) | SOME (SOME old_bound, f) => (SOME ((case btype of - UPPER => FloatArith.min - | LOWER => FloatArith.max) + UPPER => Float.min + | LOWER => Float.max) old_bound bound), f) in VarGraph.update (key, x) g @@ -125,16 +125,16 @@ fun calc_sure_bound_from_sources g (key as (_, btype)) = let fun mult_upper x (lower, upper) = - if FloatArith.is_negative x then - FloatArith.mul x lower + if Float.cmp_zero x = LESS then + Float.mult x lower else - FloatArith.mul x upper + Float.mult x upper fun mult_lower x (lower, upper) = - if FloatArith.is_negative x then - FloatArith.mul x upper + if Float.cmp_zero x = LESS then + Float.mult x upper else - FloatArith.mul x lower + Float.mult x lower val mult_btype = case btype of UPPER => mult_upper | LOWER => mult_lower @@ -146,7 +146,7 @@ | SOME x => (case get_sure_bound g src_key of NONE => NONE - | SOME src_sure_bound => SOME (FloatArith.add x (mult_btype src_sure_bound coeff))) + | SOME src_sure_bound => SOME (Float.add x (mult_btype src_sure_bound coeff))) in case fold add_src_bound sources (SOME row_bound) of NONE => sure_bound @@ -155,8 +155,8 @@ NONE => new_sure_bound | SOME old_bound => SOME (case btype of - UPPER => FloatArith.min old_bound new_bound - | LOWER => FloatArith.max old_bound new_bound)) + UPPER => Float.min old_bound new_bound + | LOWER => Float.max old_bound new_bound)) end in case VarGraph.lookup g key of @@ -187,6 +187,15 @@ exception Load of string; +(*val empty_spvec = @ {term "Nil :: (nat * real) list"}; +fun cons_spvec x xs = @ {term "Cons :: nat * real => nat * real => (nat * real) list"} $ x $ xs; +val empty_spmat = @ {term "Nil :: (nat * (nat * real) list) list"}; +fun cons_spmat x xs = @ {term "Cons :: nat * (nat * real) list => (nat * (nat * real) list) list => (nat * (nat * real) list) list"} $ x $ xs;*) +val empty_spvec = Bound 0 +fun cons_spvec x xs = Bound 0 +val empty_spmat = Bound 0 +fun cons_spmat x xs = Bound 0 + fun calcr safe_propagation xlen names prec A b = let val empty = Inttab.empty @@ -198,17 +207,17 @@ fun test_1 (lower, upper) = if lower = upper then - (if FloatArith.is_equal lower (IntInf.fromInt ~1, FloatArith.izero) then ~1 - else if FloatArith.is_equal lower (IntInf.fromInt 1, FloatArith.izero) then 1 + (if Float.eq (lower, (Intt.int ~1, Intt.zero)) then ~1 + else if Float.eq (lower, (Intt.int 1, Intt.zero)) then 1 else 0) else 0 fun calcr (row_index, a) g = let val b = FloatSparseMatrixBuilder.v_elem_at b row_index - val (_, b2) = ExactFloatingPoint.approx_decstr_by_bin prec (case b of NONE => "0" | SOME b => b) + val (_, b2) = FloatArith.approx_decstr_by_bin prec (case b of NONE => "0" | SOME b => b) val approx_a = FloatSparseMatrixBuilder.v_fold (fn (i, s) => fn l => - (i, ExactFloatingPoint.approx_decstr_by_bin prec s)::l) a [] + (i, FloatArith.approx_decstr_by_bin prec s)::l) a [] fun fold_dest_nodes (dest_index, dest_value) g = let @@ -219,7 +228,7 @@ else let val (dest_key as (_, dest_btype), row_bound) = if dest_test = ~1 then - ((dest_index, LOWER), FloatArith.neg b2) + ((dest_index, LOWER), Float.neg b2) else ((dest_index, UPPER), b2) @@ -228,10 +237,10 @@ else let val coeff = case dest_btype of - UPPER => (FloatArith.neg src_upper, FloatArith.neg src_lower) + UPPER => (Float.neg src_upper, Float.neg src_lower) | LOWER => src_value in - if FloatArith.is_negative src_lower then + if Float.cmp_zero src_lower = LESS then add_edge g (src_index, UPPER) dest_key row_index coeff else add_edge g (src_index, LOWER) dest_key row_index coeff @@ -248,7 +257,7 @@ val atest = test_1 a in if atest = ~1 then - update_sure_bound g (u, LOWER) (FloatArith.neg b2) + update_sure_bound g (u, LOWER) (Float.neg b2) else if atest = 1 then update_sure_bound g (u, UPPER) b2 else @@ -264,28 +273,28 @@ fun add_row_entry m index value = let - val vec = FloatSparseMatrixBuilder.cons_spvec (FloatSparseMatrixBuilder.mk_spvec_entry 0 value) FloatSparseMatrixBuilder.empty_spvec + val vec = cons_spvec (FloatSparseMatrixBuilder.mk_spvec_entry 0 value) empty_spvec in - FloatSparseMatrixBuilder.cons_spmat (FloatSparseMatrixBuilder.mk_spmat_entry index vec) m + cons_spmat (FloatSparseMatrixBuilder.mk_spmat_entry index vec) m end fun abs_estimate i r1 r2 = if i = 0 then - let val e = FloatSparseMatrixBuilder.empty_spmat in (e, (e, e)) end + let val e = empty_spmat in (e, (e, e)) end else let val index = xlen-i val (r, (r12_1, r12_2)) = abs_estimate (i-1) r1 r2 val b1 = case Inttab.lookup r1 index of NONE => raise (Load ("x-value not bounded from below: "^(names index))) | SOME x => x val b2 = case Inttab.lookup r2 index of NONE => raise (Load ("x-value not bounded from above: "^(names index))) | SOME x => x - val abs_max = FloatArith.max (FloatArith.neg (FloatArith.negative_part b1)) (FloatArith.positive_part b2) + val abs_max = Float.max (Float.neg (Float.negative_part b1)) (Float.positive_part b2) + val i' = Intt.int index in - (add_row_entry r index abs_max, (add_row_entry r12_1 index b1, add_row_entry r12_2 index b2)) + (add_row_entry r i' abs_max, (add_row_entry r12_1 i' b1, add_row_entry r12_2 i' b2)) end - val sign = FloatSparseMatrixBuilder.sign_term val (r, (r1, r2)) = abs_estimate xlen r1 r2 in - (sign r, (sign r1, sign r2)) + (r, (r1, r2)) end fun load filename prec safe_propagation = @@ -304,7 +313,7 @@ val v = FloatSparseMatrixBuilder.set_vector FloatSparseMatrixBuilder.empty_matrix 0 v val b = FloatSparseMatrixBuilder.transpose_matrix (FloatSparseMatrixBuilder.set_vector FloatSparseMatrixBuilder.empty_matrix 0 b) val c = FloatSparseMatrixBuilder.set_vector FloatSparseMatrixBuilder.empty_matrix 0 c - val (y1, _) = FloatSparseMatrixBuilder.approx_matrix prec FloatArith.positive_part v + val (y1, _) = FloatSparseMatrixBuilder.approx_matrix prec Float.positive_part v val A = FloatSparseMatrixBuilder.approx_matrix prec id A val (_,b2) = FloatSparseMatrixBuilder.approx_matrix prec id b val c = FloatSparseMatrixBuilder.approx_matrix prec id c diff -r 509b1da3cee1 -r 2284e0d02e7f src/HOL/Matrix/cplex/matrixlp.ML --- a/src/HOL/Matrix/cplex/matrixlp.ML Mon May 14 12:52:54 2007 +0200 +++ b/src/HOL/Matrix/cplex/matrixlp.ML Mon May 14 12:52:56 2007 +0200 @@ -23,13 +23,13 @@ ([(certT (TVar (hd (term_tvars (prop_of thm)))), certT HOLogic.realT)], []) thm) end -val spm_mult_le_dual_prts_no_let = thm "SparseMatrix.spm_mult_le_dual_prts_no_let" +val spm_mult_le_dual_prts_no_let = thm "SparseMatrix.spm_mult_le_dual_prts_no_let"; +val cert = cterm_of (Thm.theory_of_thm spm_mult_le_dual_prts_no_let); fun lp_dual_estimate_prt_primitive (y, (A1, A2), (c1, c2), b, (r1, r2)) = let - val cert = cterm_of (Thm.theory_of_thm spm_mult_le_dual_prts_no_let) val th = inst_real spm_mult_le_dual_prts_no_let - fun var s x = (cert (Var ((s,0), FloatSparseMatrixBuilder.real_spmatT)), x) + fun var s x = (cert (Var ((s,0), FloatSparseMatrixBuilder.spmatT)), x) val th = Thm.instantiate ([], [var "A1" A1, var "A2" A2, var "y" y, var "c1" c1, var "c2" c2, var "r1" r1, var "r2" r2, var "b" b]) th in th end @@ -38,10 +38,10 @@ let val certificate = let - open fspmlp + open Fspmlp val l = load lptfile prec false in - (y l, A l, c l, b l, r12 l) + (y l |> cert, A l |> pairself cert, c l |> pairself cert, b l |> cert, r12 l |> pairself cert) end in lp_dual_estimate_prt_primitive certificate @@ -55,7 +55,7 @@ check (concl_of th) end -fun prep ths = (Library.filter is_meta_eq ths) @ (map (standard o mk_meta_eq) (Library.filter (not o is_meta_eq) ths)) +fun prep ths = filter is_meta_eq ths @ map (standard o mk_meta_eq) (filter_out is_meta_eq ths) fun inst_tvar ty thm = let @@ -71,8 +71,8 @@ val matrix_compute = let - val spvecT = FloatSparseMatrixBuilder.real_spvecT - val spmatT = FloatSparseMatrixBuilder.real_spmatT + val spvecT = FloatSparseMatrixBuilder.spvecT + val spmatT = FloatSparseMatrixBuilder.spmatT val spvecT_elem = HOLogic.mk_prodT (HOLogic.natT, HOLogic.realT) val spmatT_elem = HOLogic.mk_prodT (HOLogic.natT, spvecT) val case_compute = map thm ["list_case_compute", "list_case_compute_empty", "list_case_compute_cons"] @@ -86,7 +86,7 @@ @ (map (inst_tvar HOLogic.realT) (thms "MatrixLP.sparse_row_matrix_arith_simps")) @ (thms "MatrixLP.boolarith") @ (inst_tvars [HOLogic.natT, HOLogic.realT] [thm "fst_compute", thm "snd_compute"]) - @ (inst_tvars [HOLogic.natT, FloatSparseMatrixBuilder.real_spvecT] [thm "fst_compute", thm "snd_compute"]) + @ (inst_tvars [HOLogic.natT, FloatSparseMatrixBuilder.spvecT] [thm "fst_compute", thm "snd_compute"]) @ (inst_tvars [HOLogic.boolT, spmatT_elem] case_compute) @ (inst_tvars [HOLogic.boolT, spvecT_elem] case_compute) @ (inst_tvars [HOLogic.boolT, HOLogic.realT] case_compute) diff -r 509b1da3cee1 -r 2284e0d02e7f src/HOL/Real/Float.thy --- a/src/HOL/Real/Float.thy Mon May 14 12:52:54 2007 +0200 +++ b/src/HOL/Real/Float.thy Mon May 14 12:52:56 2007 +0200 @@ -7,7 +7,7 @@ theory Float imports Real Parity -uses ("float.ML") +uses "~~/src/Pure/General/float.ML" ("float_arith.ML") begin definition @@ -526,6 +526,6 @@ (* for use with the compute oracle *) lemmas arith = binarith intarith intarithrel natarith powerarith floatarith not_false_eq_true not_true_eq_false -use "float.ML"; +use "float_arith.ML"; end diff -r 509b1da3cee1 -r 2284e0d02e7f src/HOL/Real/float.ML --- a/src/HOL/Real/float.ML Mon May 14 12:52:54 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,373 +0,0 @@ -(* Title: HOL/Real/Float.ML - ID: $Id$ - Author: Steven Obua -*) - -structure ExactFloatingPoint : -sig - exception Destruct_floatstr of string - val destruct_floatstr : (char -> bool) -> (char -> bool) -> string -> bool * string * string * bool * string - - exception Floating_point of string - - type floatrep = IntInf.int * IntInf.int - val approx_dec_by_bin : IntInf.int -> floatrep -> floatrep * floatrep - val approx_decstr_by_bin : int -> string -> floatrep * floatrep -end -= -struct - -exception Destruct_floatstr of string; - -fun destruct_floatstr isDigit isExp number = - let - val numlist = filter (not o Char.isSpace) (String.explode number) - - fun countsigns ((#"+")::cs) = countsigns cs - | countsigns ((#"-")::cs) = - let - val (positive, rest) = countsigns cs - in - (not positive, rest) - end - | countsigns cs = (true, cs) - - fun readdigits [] = ([], []) - | readdigits (q as c::cs) = - if (isDigit c) then - let - val (digits, rest) = readdigits cs - in - (c::digits, rest) - end - else - ([], q) - - fun readfromexp_helper cs = - let - val (positive, rest) = countsigns cs - val (digits, rest') = readdigits rest - in - case rest' of - [] => (positive, digits) - | _ => raise (Destruct_floatstr number) - end - - fun readfromexp [] = (true, []) - | readfromexp (c::cs) = - if isExp c then - readfromexp_helper cs - else - raise (Destruct_floatstr number) - - fun readfromdot [] = ([], readfromexp []) - | readfromdot ((#".")::cs) = - let - val (digits, rest) = readdigits cs - val exp = readfromexp rest - in - (digits, exp) - end - | readfromdot cs = readfromdot ((#".")::cs) - - val (positive, numlist) = countsigns numlist - val (digits1, numlist) = readdigits numlist - val (digits2, exp) = readfromdot numlist - in - (positive, String.implode digits1, String.implode digits2, fst exp, String.implode (snd exp)) - end - -type floatrep = IntInf.int * IntInf.int - -exception Floating_point of string; - -val ln2_10 = (Math.ln 10.0)/(Math.ln 2.0) - -fun intmul a b = IntInf.* (a,b) -fun intsub a b = IntInf.- (a,b) -fun intadd a b = IntInf.+ (a,b) -fun intpow a b = IntInf.pow (a, IntInf.toInt b); -fun intle a b = IntInf.<= (a, b); -fun intless a b = IntInf.< (a, b); -fun intneg a = IntInf.~ a; -val zero = IntInf.fromInt 0; -val one = IntInf.fromInt 1; -val two = IntInf.fromInt 2; -val ten = IntInf.fromInt 10; -val five = IntInf.fromInt 5; - -fun find_most_significant q r = - let - fun int2real i = - case Real.fromString (IntInf.toString i) of - SOME r => r - | NONE => raise (Floating_point "int2real") - fun subtract (q, r) (q', r') = - if intle r r' then - (intsub q (intmul q' (intpow ten (intsub r' r))), r) - else - (intsub (intmul q (intpow ten (intsub r r'))) q', r') - fun bin2dec d = - if intle zero d then - (intpow two d, zero) - else - (intpow five (intneg d), d) - - val L = IntInf.fromInt (Real.floor (int2real (IntInf.fromInt (IntInf.log2 q)) + (int2real r) * ln2_10)) - val L1 = intadd L one - - val (q1, r1) = subtract (q, r) (bin2dec L1) - in - if intle zero q1 then - let - val (q2, r2) = subtract (q, r) (bin2dec (intadd L1 one)) - in - if intle zero q2 then - raise (Floating_point "find_most_significant") - else - (L1, (q1, r1)) - end - else - let - val (q0, r0) = subtract (q, r) (bin2dec L) - in - if intle zero q0 then - (L, (q0, r0)) - else - raise (Floating_point "find_most_significant") - end - end - -fun approx_dec_by_bin n (q,r) = - let - fun addseq acc d' [] = acc - | addseq acc d' (d::ds) = addseq (intadd acc (intpow two (intsub d d'))) d' ds - - fun seq2bin [] = (zero, zero) - | seq2bin (d::ds) = (intadd (addseq zero d ds) one, d) - - fun approx d_seq d0 precision (q,r) = - if q = zero then - let val x = seq2bin d_seq in - (x, x) - end - else - let - val (d, (q', r')) = find_most_significant q r - in - if intless precision (intsub d0 d) then - let - val d' = intsub d0 precision - val x1 = seq2bin (d_seq) - val x2 = (intadd (intmul (fst x1) (intpow two (intsub (snd x1) d'))) one, d') (* = seq2bin (d'::d_seq) *) - in - (x1, x2) - end - else - approx (d::d_seq) d0 precision (q', r') - end - - fun approx_start precision (q, r) = - if q = zero then - ((zero, zero), (zero, zero)) - else - let - val (d, (q', r')) = find_most_significant q r - in - if intle precision zero then - let - val x1 = seq2bin [d] - in - if q' = zero then - (x1, x1) - else - (x1, seq2bin [intadd d one]) - end - else - approx [d] d precision (q', r') - end - in - if intle zero q then - approx_start n (q,r) - else - let - val ((a1,b1), (a2, b2)) = approx_start n (intneg q, r) - in - ((intneg a2, b2), (intneg a1, b1)) - end - end - -fun approx_decstr_by_bin n decstr = - let - fun str2int s = case IntInf.fromString s of SOME x => x | NONE => zero - fun signint p x = if p then x else intneg x - - val (p, d1, d2, ep, e) = destruct_floatstr Char.isDigit (fn e => e = #"e" orelse e = #"E") decstr - val s = IntInf.fromInt (size d2) - - val q = signint p (intadd (intmul (str2int d1) (intpow ten s)) (str2int d2)) - val r = intsub (signint ep (str2int e)) s - in - approx_dec_by_bin (IntInf.fromInt n) (q,r) - end - -end; - -structure FloatArith = -struct - -type float = IntInf.int * IntInf.int - -val izero = IntInf.fromInt 0 -val ione = IntInf.fromInt 1 -val imone = IntInf.fromInt ~1 -val itwo = IntInf.fromInt 2 -fun imul a b = IntInf.* (a,b) -fun isub a b = IntInf.- (a,b) -fun iadd a b = IntInf.+ (a,b) - -val floatzero = (izero, izero) - -fun positive_part (a,b) = - (if IntInf.< (a,izero) then izero else a, b) - -fun negative_part (a,b) = - (if IntInf.< (a,izero) then a else izero, b) - -fun is_negative (a,b) = - if IntInf.< (a, izero) then true else false - -fun is_positive (a,b) = - if IntInf.< (izero, a) then true else false - -fun is_zero (a,b) = - if a = izero then true else false - -fun ipow2 a = IntInf.pow ((IntInf.fromInt 2), IntInf.toInt a) - -fun add (a1, b1) (a2, b2) = - if IntInf.< (b1, b2) then - (iadd a1 (imul a2 (ipow2 (isub b2 b1))), b1) - else - (iadd (imul a1 (ipow2 (isub b1 b2))) a2, b2) - -fun sub (a1, b1) (a2, b2) = - if IntInf.< (b1, b2) then - (isub a1 (imul a2 (ipow2 (isub b2 b1))), b1) - else - (isub (imul a1 (ipow2 (isub b1 b2))) a2, b2) - -fun neg (a, b) = (IntInf.~ a, b) - -fun is_equal a b = is_zero (sub a b) - -fun is_less a b = is_negative (sub a b) - -fun max a b = if is_less a b then b else a - -fun min a b = if is_less a b then a else b - -fun abs a = if is_negative a then neg a else a - -fun mul (a1, b1) (a2, b2) = (imul a1 a2, iadd b1 b2) - -end; - - -structure Float: -sig - type float = FloatArith.float - type floatfunc = float * float -> float * float - - val mk_intinf : typ -> IntInf.int -> term - val mk_float : float -> term - - exception Dest_intinf; - val dest_intinf : term -> IntInf.int - val dest_nat : term -> IntInf.int - - exception Dest_float; - val dest_float : term -> float - - val float_const : term - - val float_add_const : term - val float_diff_const : term - val float_uminus_const : term - val float_pprt_const : term - val float_nprt_const : term - val float_abs_const : term - val float_mult_const : term - val float_le_const : term - - val nat_le_const : term - val nat_less_const : term - val nat_eq_const : term - - val approx_float : int -> floatfunc -> string -> term * term - -end -= -struct - -structure Inttab = TableFun(type key = int val ord = (rev_order o int_ord)); - -type float = IntInf.int*IntInf.int -type floatfunc = float*float -> float*float - -val float_const = Const ("Float.float", HOLogic.mk_prodT (HOLogic.intT, HOLogic.intT) --> HOLogic.realT) - -val float_add_const = Const ("HOL.plus", HOLogic.realT --> HOLogic.realT --> HOLogic.realT) -val float_diff_const = Const ("HOL.minus", HOLogic.realT --> HOLogic.realT --> HOLogic.realT) -val float_mult_const = Const ("HOL.times", HOLogic.realT --> HOLogic.realT --> HOLogic.realT) -val float_uminus_const = Const ("HOL.uminus", HOLogic.realT --> HOLogic.realT) -val float_abs_const = Const ("HOL.abs", HOLogic.realT --> HOLogic.realT) -val float_le_const = Const ("Orderings.less_eq", HOLogic.realT --> HOLogic.realT --> HOLogic.boolT) -val float_pprt_const = Const ("OrderedGroup.pprt", HOLogic.realT --> HOLogic.realT) -val float_nprt_const = Const ("OrderedGroup.nprt", HOLogic.realT --> HOLogic.realT) - -val nat_le_const = Const ("Orderings.less_eq", HOLogic.natT --> HOLogic.natT --> HOLogic.boolT) -val nat_less_const = Const ("Orderings.less", HOLogic.natT --> HOLogic.natT --> HOLogic.boolT) -val nat_eq_const = Const ("op =", HOLogic.natT --> HOLogic.natT --> HOLogic.boolT) - -val zero = FloatArith.izero -val minus_one = FloatArith.imone -val two = FloatArith.itwo - -exception Dest_intinf; -exception Dest_float; - -fun mk_intinf ty n = HOLogic.number_of_const ty $ HOLogic.mk_numeral n; - -val dest_intinf = snd o HOLogic.dest_number - -fun mk_float (a,b) = - float_const $ (HOLogic.mk_prod ((mk_intinf HOLogic.intT a), (mk_intinf HOLogic.intT b))) - -fun dest_float f = - case f of - (Const ("Float.float", _) $ (Const ("Pair", _) $ a $ b)) => (dest_intinf a, dest_intinf b) - | Const ("Numeral.number_of",_) $ a => (dest_intinf f, 0) - | Const ("Numeral0", _) => (FloatArith.izero, FloatArith.izero) - | Const ("Numeral1", _) => (FloatArith.ione, FloatArith.izero) - | _ => raise Dest_float - -fun dest_nat n = - let - val v = dest_intinf n - in - if IntInf.< (v, FloatArith.izero) then - FloatArith.izero - else - v - end - -fun approx_float prec f value = - let - val interval = ExactFloatingPoint.approx_decstr_by_bin prec value - val (flower, fupper) = f interval - in - (mk_float flower, mk_float fupper) - end - -end; diff -r 509b1da3cee1 -r 2284e0d02e7f src/HOL/Real/float_arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Real/float_arith.ML Mon May 14 12:52:56 2007 +0200 @@ -0,0 +1,226 @@ +(* Title: HOL/Real/Float.ML + ID: $Id$ + Author: Steven Obua +*) + +signature FLOAT_ARITH = +sig + exception Destruct_floatstr of string + 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_decstr_by_bin: int -> string -> Float.float * Float.float + + val mk_float: Float.float -> term + val dest_float: term -> Float.float + + val approx_float: int -> (Float.float * Float.float -> Float.float * Float.float) + -> string -> term * term +end; + +structure FloatArith : FLOAT_ARITH = +struct + +exception Destruct_floatstr of string; + +fun destruct_floatstr isDigit isExp number = + let + val numlist = filter (not o Char.isSpace) (String.explode number) + + fun countsigns ((#"+")::cs) = countsigns cs + | countsigns ((#"-")::cs) = + let + val (positive, rest) = countsigns cs + in + (not positive, rest) + end + | countsigns cs = (true, cs) + + fun readdigits [] = ([], []) + | readdigits (q as c::cs) = + if (isDigit c) then + let + val (digits, rest) = readdigits cs + in + (c::digits, rest) + end + else + ([], q) + + fun readfromexp_helper cs = + let + val (positive, rest) = countsigns cs + val (digits, rest') = readdigits rest + in + case rest' of + [] => (positive, digits) + | _ => raise (Destruct_floatstr number) + end + + fun readfromexp [] = (true, []) + | readfromexp (c::cs) = + if isExp c then + readfromexp_helper cs + else + raise (Destruct_floatstr number) + + fun readfromdot [] = ([], readfromexp []) + | readfromdot ((#".")::cs) = + let + val (digits, rest) = readdigits cs + val exp = readfromexp rest + in + (digits, exp) + end + | readfromdot cs = readfromdot ((#".")::cs) + + val (positive, numlist) = countsigns numlist + val (digits1, numlist) = readdigits numlist + val (digits2, exp) = readfromdot numlist + in + (positive, String.implode digits1, String.implode digits2, fst exp, String.implode (snd exp)) + end + +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); + +fun intle a b = not (Intt.cmp (a, b) = GREATER); +fun intless a b = Intt.cmp (a, b) = LESS; + +fun find_most_significant q r = + let + fun int2real i = + case (Real.fromString o Intt.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) + else + (Intt.sub (Intt.mult q (exp10 (Intt.sub r r'))) q', r') + fun bin2dec d = + if intle Intt.zero d then + (Intt.exp d, Intt.zero) + else + (exp5 (Intt.neg d), d) + + val L = Intt.int (Real.floor (int2real (Intt.log q) + int2real r * ln2_10)) + val L1 = Intt.inc L + + val (q1, r1) = subtract (q, r) (bin2dec L1) + in + if intle Intt.zero q1 then + let + val (q2, r2) = subtract (q, r) (bin2dec (Intt.inc L1)) + in + if intle Intt.zero q2 then + raise (Floating_point "find_most_significant") + else + (L1, (q1, r1)) + end + else + let + val (q0, r0) = subtract (q, r) (bin2dec L) + in + if intle Intt.zero q0 then + (L, (q0, r0)) + else + raise (Floating_point "find_most_significant") + end + end + +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 + + fun seq2bin [] = (Intt.zero, Intt.zero) + | seq2bin (d::ds) = (Intt.inc (addseq Intt.zero d ds), d) + + fun approx d_seq d0 precision (q,r) = + if q = Intt.zero then + let val x = seq2bin d_seq in + (x, x) + end + else + let + val (d, (q', r')) = find_most_significant q r + in + if intless precision (Intt.sub d0 d) then + let + val d' = Intt.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) *) + in + (x1, x2) + end + else + approx (d::d_seq) d0 precision (q', r') + end + + fun approx_start precision (q, r) = + if q = Intt.zero then + ((Intt.zero, Intt.zero), (Intt.zero, Intt.zero)) + else + let + val (d, (q', r')) = find_most_significant q r + in + if intle precision Intt.zero then + let + val x1 = seq2bin [d] + in + if q' = Intt.zero then + (x1, x1) + else + (x1, seq2bin [Intt.inc d]) + end + else + approx [d] d precision (q', r') + end + in + if intle Intt.zero q then + approx_start n (q,r) + else + let + val ((a1,b1), (a2, b2)) = approx_start n (Intt.neg q, r) + in + ((Intt.neg a2, b2), (Intt.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 + + 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 q = signint p (Intt.add (Intt.mult (str2int d1) (exp10 s)) (str2int d2)) + val r = Intt.sub (signint ep (str2int e)) s + in + approx_dec_by_bin (Intt.int n) (q,r) + end + +fun mk_float (a, b) = @{term "float"} $ + HOLogic.mk_prod (pairself (HOLogic.mk_number HOLogic.intT) (a, b)); + +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); + +fun approx_float prec f value = + let + val interval = approx_decstr_by_bin prec value + val (flower, fupper) = f interval + in + (mk_float flower, mk_float fupper) + end; + +end; diff -r 509b1da3cee1 -r 2284e0d02e7f src/Pure/General/float.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/General/float.ML Mon May 14 12:52:56 2007 +0200 @@ -0,0 +1,60 @@ +(* Title: Pure/General/float.ML + ID: $Id$ + Author: Steven Obua, Florian Haftmann, TU Muenchen + +Implementation of real numbers as mantisse-exponent pairs. +*) + +signature FLOAT = +sig + type float = Intt.int * Intt.int + val zero: float + val eq: float * float -> bool + val cmp: float * float -> order + val cmp_zero: float -> order + val min: float -> float -> float + val max: float -> float -> float + val add: float -> float -> float + val sub: float -> float -> float + val neg: float -> float + val mult: float -> float -> float + val positive_part: float -> float + val negative_part: float -> float +end; + +structure Float : FLOAT = +struct + +type float = Intt.int * Intt.int; + +val zero = (Intt.zero, Intt.zero); + +fun add (a1, b1) (a2, b2) = + if Intt.cmp (b1, b2) = LESS then + (Intt.add a1 (Intt.mult a2 (Intt.exp (Intt.sub b2 b1))), b1) + else + (Intt.add (Intt.mult a1 (Intt.exp (Intt.sub b1 b2))) a2, b2); + +fun sub (a1, b1) (a2, b2) = + if Intt.cmp (b1, b2) = LESS then + (Intt.sub a1 (Intt.mult a2 (Intt.exp (Intt.sub b2 b1))), b1) + else + (Intt.sub (Intt.mult a1 (Intt.exp (Intt.sub b1 b2))) a2, b2); + +fun neg (a, b) = (Intt.neg a, b); + +fun mult (a1, b1) (a2, b2) = (Intt.mult a1 a2, Intt.add b1 b2); + +fun cmp_zero (a, b) = Intt.cmp_zero a; + +fun cmp (r, s) = cmp_zero (sub r s); + +fun eq (r, s) = cmp (r, s) = EQUAL; + +fun min r s = case cmp (r, s) of LESS => r | _ => s; +fun max r s = case cmp (r, s) of LESS => s | _ => r; + +fun positive_part (a, b) = (Intt.max Intt.zero a, b); +fun negative_part (a, b) = (Intt.min Intt.zero a, b); + +end; diff -r 509b1da3cee1 -r 2284e0d02e7f src/Pure/General/int.ML --- a/src/Pure/General/int.ML Mon May 14 12:52:54 2007 +0200 +++ b/src/Pure/General/int.ML Mon May 14 12:52:56 2007 +0200 @@ -19,6 +19,9 @@ val cmp: int * int -> order val le: int -> int -> bool val cmp_zero: int -> order + val min: int -> int -> int + val max: int -> int -> int + val inc: int -> int val add: int -> int -> int val sub: int -> int -> int val mult: int -> int -> int @@ -27,6 +30,8 @@ val mod: int -> int -> int val neg: int -> int val exp: int -> int + val log: int -> int + val pow: int -> int -> int (* exponent -> base -> result *) end; structure Intt: INTT = @@ -34,29 +39,35 @@ open IntInf; -val int = IntInf.fromInt; +val int = fromInt; val zero = int 0; val one = int 1; val two = int 2; -val machine_int = IntInf.toInt; -val string_of_int = IntInf.toString; -val int_of_string = IntInf.fromString; +val machine_int = toInt; +val string_of_int = toString; +val int_of_string = fromString; -val eq = op =; +val eq = op = : int * int -> bool; val cmp = compare; val le = curry (op <=); val cmp_zero = curry cmp zero; + +val min = curry min; +val max = curry max; + +val inc = curry (op +) one; + val add = curry (op +); val sub = curry (op -); val mult = curry ( op * ); val divmod = curry divMod; -nonfix div -val div = curry div; -nonfix mod -val mod = curry mod; -val neg = IntInf.~; -fun exp n = pow (2, IntInf.toInt n); +nonfix div val div = curry div; +nonfix mod val mod = curry mod; +val neg = ~; +val pow = fn k => fn l => pow (l, toInt k); +fun exp k = pow k two; +val log = int o log2; end;