merged
authorwenzelm
Tue, 10 Mar 2009 17:54:49 +0100
changeset 30414 aa8c5ab72cc6
parent 30413 c41afa5607be (diff)
parent 30412 7f5b0a020ccd (current diff)
child 30417 09a757ca128f
merged
--- a/src/HOL/Decision_Procs/Approximation.thy	Tue Mar 10 16:51:08 2009 +0100
+++ b/src/HOL/Decision_Procs/Approximation.thy	Tue Mar 10 17:54:49 2009 +0100
@@ -2422,10 +2422,6 @@
 lemma "0 < (d :: int) \<Longrightarrow> 0 \<le> i mod d \<and> i mod d < d" by auto
 lemma "(d :: int) < 0 \<Longrightarrow> d < i mod d \<and> i mod d \<le> 0" by auto
 
-code_const "op div :: int \<Rightarrow> int \<Rightarrow> int" (SML "(fn i => fn d => if d = 0 then 0 else i div d)")
-code_const "op mod :: int \<Rightarrow> int \<Rightarrow> int" (SML "(fn i => fn d => if d = 0 then i else i mod d)")
-code_const "divmod :: int \<Rightarrow> int \<Rightarrow> (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"}
--- a/src/HOL/ex/ApproximationEx.thy	Tue Mar 10 16:51:08 2009 +0100
+++ b/src/HOL/ex/ApproximationEx.thy	Tue Mar 10 17:54:49 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 \<le> x \<and> x \<le> 4.5 \<Longrightarrow> \<bar> arctan x - 0.91 \<bar> < 0.455"
   by (approximation 10)
 
+lemma "0.49 \<le> x \<and> x \<le> 4.49 \<Longrightarrow> \<bar> arctan x - 0.91 \<bar> < 0.455"
+  by (approximation 20)
+
+lemma "1 * pow2 -1 \<le> x \<and> x \<le> 9 * pow2 -1 \<Longrightarrow> \<bar> arctan x - 0.91 \<bar> < 0.455"
+  by (approximation 10)
+
 lemma "0 \<le> x \<and> x \<le> 1 \<Longrightarrow> 0 \<le> sin x"
   by (approximation 10)
 
-
 end