# HG changeset patch # User hoelzl # Date 1236699382 -3600 # Node ID c41afa5607be35f22250b26f657942d535de10b9 # Parent 81218f70997fc80bc67b47c3c51d314a8c18d5fe Fixed type error which appeared when Approximation bounds where specified as floating point numbers diff -r 81218f70997f -r c41afa5607be src/HOL/Decision_Procs/Approximation.thy --- a/src/HOL/Decision_Procs/Approximation.thy Tue Mar 10 11:01:28 2009 +0100 +++ b/src/HOL/Decision_Procs/Approximation.thy Tue Mar 10 16:36:22 2009 +0100 @@ -2422,10 +2422,6 @@ lemma "0 < (d :: int) \ 0 \ i mod d \ i mod d < d" by auto lemma "(d :: int) < 0 \ d < i mod d \ i mod d \ 0" by auto -code_const "op div :: int \ int \ int" (SML "(fn i => fn d => if d = 0 then 0 else i div d)") -code_const "op mod :: int \ int \ int" (SML "(fn i => fn d => if d = 0 then i else i mod d)") -code_const "divmod :: int \ int \ (int * int)" (SML "(fn i => fn d => if d = 0 then (0, i) else IntInf.divMod (i, d))") - ML {* val uneq_equations = PureThy.get_thms @{theory} "uneq_equations"; val bounded_by_equations = PureThy.get_thms @{theory} "bounded_by_equations"; @@ -2448,7 +2444,7 @@ fun bot_float (Const (@{const_name "times"}, _) $ mantisse $ (Const (@{const_name "pow2"}, _) $ exp)) = @{term "Float"} $ to_int mantisse $ to_int exp | bot_float (Const (@{const_name "divide"}, _) $ mantisse $ (Const (@{const_name "power"}, _) $ ten $ exp)) - = @{term "float_divl"} $ prec' $ (@{term "Float"} $ to_int mantisse $ @{term "0 :: int"}) $ (@{term "power_float (Float 5 1)"} $ to_nat exp) + = @{term "float_divl"} $ prec' $ (@{term "Float"} $ to_int mantisse $ @{term "0 :: int"}) $ (@{term "power (Float 5 1)"} $ to_nat exp) | bot_float (Const (@{const_name "divide"}, _) $ mantisse $ ten) = @{term "float_divl"} $ prec' $ (@{term "Float"} $ to_int mantisse $ @{term "0 :: int"}) $ @{term "Float 5 1"} | bot_float mantisse = @{term "Float"} $ to_int mantisse $ @{term "0 :: int"} @@ -2456,7 +2452,7 @@ fun top_float (Const (@{const_name "times"}, _) $ mantisse $ (Const (@{const_name "pow2"}, _) $ exp)) = @{term "Float"} $ to_int mantisse $ to_int exp | top_float (Const (@{const_name "divide"}, _) $ mantisse $ (Const (@{const_name "power"}, _) $ ten $ exp)) - = @{term "float_divr"} $ prec' $ (@{term "Float"} $ to_int mantisse $ @{term "0 :: int"}) $ (@{term "power_float (Float 5 1)"} $ to_nat exp) + = @{term "float_divr"} $ prec' $ (@{term "Float"} $ to_int mantisse $ @{term "0 :: int"}) $ (@{term "power (Float 5 1)"} $ to_nat exp) | top_float (Const (@{const_name "divide"}, _) $ mantisse $ ten) = @{term "float_divr"} $ prec' $ (@{term "Float"} $ to_int mantisse $ @{term "0 :: int"}) $ @{term "Float 5 1"} | top_float mantisse = @{term "Float"} $ to_int mantisse $ @{term "0 :: int"} diff -r 81218f70997f -r c41afa5607be src/HOL/ex/ApproximationEx.thy --- a/src/HOL/ex/ApproximationEx.thy Tue Mar 10 11:01:28 2009 +0100 +++ b/src/HOL/ex/ApproximationEx.thy Tue Mar 10 16:36:22 2009 +0100 @@ -3,7 +3,7 @@ *) theory ApproximationEx -imports "~~/src/HOL/Reflection/Approximation" +imports "~~/src/HOL/Decision_Procs/Approximation" begin text {* @@ -39,9 +39,14 @@ lemma "0.5 \ x \ x \ 4.5 \ \ arctan x - 0.91 \ < 0.455" by (approximation 10) +lemma "0.49 \ x \ x \ 4.49 \ \ arctan x - 0.91 \ < 0.455" + by (approximation 20) + +lemma "1 * pow2 -1 \ x \ x \ 9 * pow2 -1 \ \ arctan x - 0.91 \ < 0.455" + by (approximation 10) + lemma "0 \ x \ x \ 1 \ 0 \ sin x" by (approximation 10) - end