| author | huffman | 
| Tue, 12 Oct 2010 06:20:05 -0700 | |
| changeset 40006 | 116e94f9543b | 
| parent 39161 | 75849a560c09 | 
| child 41024 | ba961a606c67 | 
| permissions | -rw-r--r-- | 
| 30122 | 1 | (* Title: HOL/Library/Float.thy | 
| 2 | Author: Steven Obua 2008 | |
| 3 | Author: Johannes Hoelzl, TU Muenchen <hoelzl@in.tum.de> 2008 / 2009 | |
| 4 | *) | |
| 29988 | 5 | |
| 6 | header {* Floating-Point Numbers *}
 | |
| 7 | ||
| 20485 | 8 | theory Float | 
| 35032 
7efe662e41b4
separate library theory for type classes combining lattices with various algebraic structures
 haftmann parents: 
33555diff
changeset | 9 | imports Complex_Main Lattice_Algebras | 
| 20485 | 10 | begin | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 11 | |
| 19765 | 12 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21256diff
changeset | 13 | pow2 :: "int \<Rightarrow> real" where | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 14 | [simp]: "pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 15 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 16 | datatype float = Float int int | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 17 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 18 | primrec of_float :: "float \<Rightarrow> real" where | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 19 | "of_float (Float a b) = real a * pow2 b" | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 20 | |
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 21 | defs (overloaded) | 
| 31998 
2c7a24f74db9
code attributes use common underscore convention
 haftmann parents: 
31863diff
changeset | 22 | real_of_float_def [code_unfold]: "real == of_float" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 23 | |
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 24 | primrec mantissa :: "float \<Rightarrow> int" where | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 25 | "mantissa (Float a b) = a" | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 26 | |
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 27 | primrec scale :: "float \<Rightarrow> int" where | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 28 | "scale (Float a b) = b" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21256diff
changeset | 29 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 30 | instantiation float :: zero begin | 
| 31467 
f7d2aa438bee
Approximation: Implemented argument reduction for cosine. Sinus is now implemented in terms of cosine. Sqrt computes on the entire real numbers
 hoelzl parents: 
31098diff
changeset | 31 | definition zero_float where "0 = Float 0 0" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 32 | instance .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 33 | end | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 34 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 35 | instantiation float :: one begin | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 36 | definition one_float where "1 = Float 1 0" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 37 | instance .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 38 | end | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 39 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 40 | instantiation float :: number begin | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 41 | definition number_of_float where "number_of n = Float n 0" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 42 | instance .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 43 | end | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 44 | |
| 32069 
6d28bbd33e2c
prefer code_inline over code_unfold; use code_unfold_post where appropriate
 haftmann parents: 
31998diff
changeset | 45 | lemma number_of_float_Float [code_unfold_post]: | 
| 31467 
f7d2aa438bee
Approximation: Implemented argument reduction for cosine. Sinus is now implemented in terms of cosine. Sqrt computes on the entire real numbers
 hoelzl parents: 
31098diff
changeset | 46 | "number_of k = Float (number_of k) 0" | 
| 
f7d2aa438bee
Approximation: Implemented argument reduction for cosine. Sinus is now implemented in terms of cosine. Sqrt computes on the entire real numbers
 hoelzl parents: 
31098diff
changeset | 47 | by (simp add: number_of_float_def number_of_is_id) | 
| 
f7d2aa438bee
Approximation: Implemented argument reduction for cosine. Sinus is now implemented in terms of cosine. Sqrt computes on the entire real numbers
 hoelzl parents: 
31098diff
changeset | 48 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 49 | lemma real_of_float_simp[simp]: "real (Float a b) = real a * pow2 b" | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 50 | unfolding real_of_float_def using of_float.simps . | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 51 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 52 | lemma real_of_float_neg_exp: "e < 0 \<Longrightarrow> real (Float m e) = real m * inverse (2^nat (-e))" by auto | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 53 | lemma real_of_float_nge0_exp: "\<not> 0 \<le> e \<Longrightarrow> real (Float m e) = real m * inverse (2^nat (-e))" by auto | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 54 | lemma real_of_float_ge0_exp: "0 \<le> e \<Longrightarrow> real (Float m e) = real m * (2^nat e)" by auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 55 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 56 | lemma Float_num[simp]: shows | 
| 31467 
f7d2aa438bee
Approximation: Implemented argument reduction for cosine. Sinus is now implemented in terms of cosine. Sqrt computes on the entire real numbers
 hoelzl parents: 
31098diff
changeset | 57 | "real (Float 1 0) = 1" and "real (Float 1 1) = 2" and "real (Float 1 2) = 4" and | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 58 | "real (Float 1 -1) = 1/2" and "real (Float 1 -2) = 1/4" and "real (Float 1 -3) = 1/8" and | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 59 | "real (Float -1 0) = -1" and "real (Float (number_of n) 0) = number_of n" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 60 | by auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 61 | |
| 31863 
e391eee8bf14
Implemented taylor series expansion for approximation
 hoelzl parents: 
31467diff
changeset | 62 | lemma float_number_of[simp]: "real (number_of x :: float) = number_of x" | 
| 
e391eee8bf14
Implemented taylor series expansion for approximation
 hoelzl parents: 
31467diff
changeset | 63 | by (simp only:number_of_float_def Float_num[unfolded number_of_is_id]) | 
| 
e391eee8bf14
Implemented taylor series expansion for approximation
 hoelzl parents: 
31467diff
changeset | 64 | |
| 
e391eee8bf14
Implemented taylor series expansion for approximation
 hoelzl parents: 
31467diff
changeset | 65 | lemma float_number_of_int[simp]: "real (Float n 0) = real n" | 
| 
e391eee8bf14
Implemented taylor series expansion for approximation
 hoelzl parents: 
31467diff
changeset | 66 | by (simp add: Float_num[unfolded number_of_is_id] real_of_float_simp pow2_def) | 
| 
e391eee8bf14
Implemented taylor series expansion for approximation
 hoelzl parents: 
31467diff
changeset | 67 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 68 | lemma pow2_0[simp]: "pow2 0 = 1" by simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 69 | lemma pow2_1[simp]: "pow2 1 = 2" by simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 70 | lemma pow2_neg: "pow2 x = inverse (pow2 (-x))" by simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 71 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 72 | declare pow2_def[simp del] | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 73 | |
| 19765 | 74 | lemma pow2_add1: "pow2 (1 + a) = 2 * (pow2 a)" | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 75 | proof - | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 76 | have h: "! n. nat (2 + int n) - Suc 0 = nat (1 + int n)" by arith | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 77 | have g: "! a b. a - -1 = a + (1::int)" by arith | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 78 | have pos: "! n. pow2 (int n + 1) = 2 * pow2 (int n)" | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 79 | apply (auto, induct_tac n) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 80 | apply (simp_all add: pow2_def) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 81 | apply (rule_tac m1="2" and n1="nat (2 + int na)" in ssubst[OF realpow_num_eq_if]) | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23365diff
changeset | 82 | by (auto simp add: h) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 83 | show ?thesis | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 84 | proof (induct a) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 85 | case (1 n) | 
| 29667 | 86 | from pos show ?case by (simp add: algebra_simps) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 87 | next | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 88 | case (2 n) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 89 | show ?case | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 90 | apply (auto) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 91 | apply (subst pow2_neg[of "- int n"]) | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23365diff
changeset | 92 | apply (subst pow2_neg[of "-1 - int n"]) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 93 | apply (auto simp add: g pos) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 94 | done | 
| 19765 | 95 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 96 | qed | 
| 19765 | 97 | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 98 | lemma pow2_add: "pow2 (a+b) = (pow2 a) * (pow2 b)" | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 99 | proof (induct b) | 
| 19765 | 100 | case (1 n) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 101 | show ?case | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 102 | proof (induct n) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 103 | case 0 | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 104 | show ?case by simp | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 105 | next | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 106 | case (Suc m) | 
| 29667 | 107 | show ?case by (auto simp add: algebra_simps pow2_add1 prems) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 108 | qed | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 109 | next | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 110 | case (2 n) | 
| 19765 | 111 | show ?case | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 112 | proof (induct n) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 113 | case 0 | 
| 19765 | 114 | show ?case | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 115 | apply (auto) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 116 | apply (subst pow2_neg[of "a + -1"]) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 117 | apply (subst pow2_neg[of "-1"]) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 118 | apply (simp) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 119 | apply (insert pow2_add1[of "-a"]) | 
| 29667 | 120 | apply (simp add: algebra_simps) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 121 | apply (subst pow2_neg[of "-a"]) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 122 | apply (simp) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 123 | done | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 124 | case (Suc m) | 
| 19765 | 125 | have a: "int m - (a + -2) = 1 + (int m - a + 1)" by arith | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 126 | have b: "int m - -2 = 1 + (int m + 1)" by arith | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 127 | show ?case | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 128 | apply (auto) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 129 | apply (subst pow2_neg[of "a + (-2 - int m)"]) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 130 | apply (subst pow2_neg[of "-2 - int m"]) | 
| 29667 | 131 | apply (auto simp add: algebra_simps) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 132 | apply (subst a) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 133 | apply (subst b) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 134 | apply (simp only: pow2_add1) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 135 | apply (subst pow2_neg[of "int m - a + 1"]) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 136 | apply (subst pow2_neg[of "int m + 1"]) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 137 | apply auto | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 138 | apply (insert prems) | 
| 29667 | 139 | apply (auto simp add: algebra_simps) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 140 | done | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 141 | qed | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 142 | qed | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 143 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 144 | lemma float_components[simp]: "Float (mantissa f) (scale f) = f" by (cases f, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 145 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 146 | lemma float_split: "\<exists> a b. x = Float a b" by (cases x, auto) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 147 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 148 | lemma float_split2: "(\<forall> a b. x \<noteq> Float a b) = False" by (auto simp add: float_split) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 149 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 150 | lemma float_zero[simp]: "real (Float 0 e) = 0" by simp | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 151 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 152 | lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> abs((a::int) div 2) < abs a" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 153 | by arith | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21256diff
changeset | 154 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 155 | function normfloat :: "float \<Rightarrow> float" where | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 156 | "normfloat (Float a b) = (if a \<noteq> 0 \<and> even a then normfloat (Float (a div 2) (b+1)) else if a=0 then Float 0 0 else Float a b)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 157 | by pat_completeness auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 158 | termination by (relation "measure (nat o abs o mantissa)") (auto intro: abs_div_2_less) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 159 | declare normfloat.simps[simp del] | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 160 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 161 | theorem normfloat[symmetric, simp]: "real f = real (normfloat f)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 162 | proof (induct f rule: normfloat.induct) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 163 | case (1 a b) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 164 | have real2: "2 = real (2::int)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 165 | by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 166 | show ?case | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 167 | apply (subst normfloat.simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 168 | apply (auto simp add: float_zero) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 169 | apply (subst 1[symmetric]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 170 | apply (auto simp add: pow2_add even_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 171 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 172 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 173 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 174 | lemma pow2_neq_zero[simp]: "pow2 x \<noteq> 0" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 175 | by (auto simp add: pow2_def) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 176 | |
| 26313 | 177 | lemma pow2_int: "pow2 (int c) = 2^c" | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 178 | by (simp add: pow2_def) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 179 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 180 | lemma zero_less_pow2[simp]: | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 181 | "0 < pow2 x" | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 182 | proof - | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 183 |   {
 | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 184 | fix y | 
| 19765 | 185 | have "0 <= y \<Longrightarrow> 0 < pow2 y" | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 186 | by (induct y, induct_tac n, simp_all add: pow2_add) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 187 | } | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 188 | note helper=this | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 189 | show ?thesis | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 190 | apply (case_tac "0 <= x") | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 191 | apply (simp add: helper) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 192 | apply (subst pow2_neg) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 193 | apply (simp add: helper) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 194 | done | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 195 | qed | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 196 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 197 | lemma normfloat_imp_odd_or_zero: "normfloat f = Float a b \<Longrightarrow> odd a \<or> (a = 0 \<and> b = 0)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 198 | proof (induct f rule: normfloat.induct) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 199 | case (1 u v) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 200 | from 1 have ab: "normfloat (Float u v) = Float a b" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 201 |   {
 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 202 | assume eu: "even u" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 203 | assume z: "u \<noteq> 0" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 204 | have "normfloat (Float u v) = normfloat (Float (u div 2) (v + 1))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 205 | apply (subst normfloat.simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 206 | by (simp add: eu z) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 207 | with ab have "normfloat (Float (u div 2) (v + 1)) = Float a b" by simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 208 | with 1 eu z have ?case by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 209 | } | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 210 | note case1 = this | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 211 |   {
 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 212 | assume "odd u \<or> u = 0" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 213 | then have ou: "\<not> (u \<noteq> 0 \<and> even u)" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 214 | have "normfloat (Float u v) = (if u = 0 then Float 0 0 else Float u v)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 215 | apply (subst normfloat.simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 216 | apply (simp add: ou) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 217 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 218 | with ab have "Float a b = (if u = 0 then Float 0 0 else Float u v)" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 219 | then have ?case | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 220 | apply (case_tac "u=0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 221 | apply (auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 222 | by (insert ou, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 223 | } | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 224 | note case2 = this | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 225 | show ?case | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 226 | apply (case_tac "odd u \<or> u = 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 227 | apply (rule case2) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 228 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 229 | apply (rule case1) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 230 | apply auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 231 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 232 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 233 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 234 | lemma float_eq_odd_helper: | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 235 | assumes odd: "odd a'" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 236 | and floateq: "real (Float a b) = real (Float a' b')" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 237 | shows "b \<le> b'" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 238 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 239 |   {
 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 240 | assume bcmp: "b > b'" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 241 | from floateq have eq: "real a * pow2 b = real a' * pow2 b'" by simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 242 |     {
 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 243 | fix x y z :: real | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 244 | assume "y \<noteq> 0" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 245 | then have "(x * inverse y = z) = (x = z * y)" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 246 | by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 247 | } | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 248 | note inverse = this | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 249 | have eq': "real a * (pow2 (b - b')) = real a'" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 250 | apply (subst diff_int_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 251 | apply (subst pow2_add) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 252 | apply (subst pow2_neg[where x = "-b'"]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 253 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 254 | apply (subst mult_assoc[symmetric]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 255 | apply (subst inverse) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 256 | apply (simp_all add: eq) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 257 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 258 | have "\<exists> z > 0. pow2 (b-b') = 2^z" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 259 | apply (rule exI[where x="nat (b - b')"]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 260 | apply (auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 261 | apply (insert bcmp) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 262 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 263 | apply (subst pow2_int[symmetric]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 264 | apply auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 265 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 266 | then obtain z where z: "z > 0 \<and> pow2 (b-b') = 2^z" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 267 | with eq' have "real a * 2^z = real a'" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 268 | by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 269 | then have "real a * real ((2::int)^z) = real a'" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 270 | by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 271 | then have "real (a * 2^z) = real a'" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 272 | apply (subst real_of_int_mult) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 273 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 274 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 275 | then have a'_rep: "a * 2^z = a'" by arith | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 276 | then have "a' = a*2^z" by simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 277 | with z have "even a'" by simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 278 | with odd have False by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 279 | } | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 280 | then show ?thesis by arith | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 281 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 282 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 283 | lemma float_eq_odd: | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 284 | assumes odd1: "odd a" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 285 | and odd2: "odd a'" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 286 | and floateq: "real (Float a b) = real (Float a' b')" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 287 | shows "a = a' \<and> b = b'" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 288 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 289 | from | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 290 | float_eq_odd_helper[OF odd2 floateq] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 291 | float_eq_odd_helper[OF odd1 floateq[symmetric]] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 292 | have beq: "b = b'" by arith | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 293 | with floateq show ?thesis by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 294 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 295 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 296 | theorem normfloat_unique: | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 297 | assumes real_of_float_eq: "real f = real g" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 298 | shows "normfloat f = normfloat g" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 299 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 300 | from float_split[of "normfloat f"] obtain a b where normf:"normfloat f = Float a b" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 301 | from float_split[of "normfloat g"] obtain a' b' where normg:"normfloat g = Float a' b'" by auto | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 302 | have "real (normfloat f) = real (normfloat g)" | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 303 | by (simp add: real_of_float_eq) | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 304 | then have float_eq: "real (Float a b) = real (Float a' b')" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 305 | by (simp add: normf normg) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 306 | have ab: "odd a \<or> (a = 0 \<and> b = 0)" by (rule normfloat_imp_odd_or_zero[OF normf]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 307 | have ab': "odd a' \<or> (a' = 0 \<and> b' = 0)" by (rule normfloat_imp_odd_or_zero[OF normg]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 308 |   {
 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 309 | assume odd: "odd a" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 310 | then have "a \<noteq> 0" by (simp add: even_def, arith) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 311 | with float_eq have "a' \<noteq> 0" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 312 | with ab' have "odd a'" by simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 313 | from odd this float_eq have "a = a' \<and> b = b'" by (rule float_eq_odd) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 314 | } | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 315 | note odd_case = this | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 316 |   {
 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 317 | assume even: "even a" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 318 | with ab have a0: "a = 0" by simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 319 | with float_eq have a0': "a' = 0" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 320 | from a0 a0' ab ab' have "a = a' \<and> b = b'" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 321 | } | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 322 | note even_case = this | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 323 | from odd_case even_case show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 324 | apply (simp add: normf normg) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 325 | apply (case_tac "even a") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 326 | apply auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 327 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 328 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 329 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 330 | instantiation float :: plus begin | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 331 | fun plus_float where | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 332 | [simp del]: "(Float a_m a_e) + (Float b_m b_e) = | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 333 | (if a_e \<le> b_e then Float (a_m + b_m * 2^(nat(b_e - a_e))) a_e | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 334 | else Float (a_m * 2^(nat (a_e - b_e)) + b_m) b_e)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 335 | instance .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 336 | end | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 337 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 338 | instantiation float :: uminus begin | 
| 30960 | 339 | primrec uminus_float where [simp del]: "uminus_float (Float m e) = Float (-m) e" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 340 | instance .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 341 | end | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 342 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 343 | instantiation float :: minus begin | 
| 30960 | 344 | definition minus_float where [simp del]: "(z::float) - w = z + (- w)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 345 | instance .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 346 | end | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 347 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 348 | instantiation float :: times begin | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 349 | fun times_float where [simp del]: "(Float a_m a_e) * (Float b_m b_e) = Float (a_m * b_m) (a_e + b_e)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 350 | instance .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 351 | end | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 352 | |
| 30960 | 353 | primrec float_pprt :: "float \<Rightarrow> float" where | 
| 354 | "float_pprt (Float a e) = (if 0 <= a then (Float a e) else 0)" | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 355 | |
| 30960 | 356 | primrec float_nprt :: "float \<Rightarrow> float" where | 
| 357 | "float_nprt (Float a e) = (if 0 <= a then 0 else (Float a e))" | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 358 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 359 | instantiation float :: ord begin | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 360 | definition le_float_def: "z \<le> (w :: float) \<equiv> real z \<le> real w" | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 361 | definition less_float_def: "z < (w :: float) \<equiv> real z < real w" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 362 | instance .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 363 | end | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 364 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 365 | lemma real_of_float_add[simp]: "real (a + b) = real a + real (b :: float)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 366 | by (cases a, cases b, simp add: algebra_simps plus_float.simps, | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 367 | auto simp add: pow2_int[symmetric] pow2_add[symmetric]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 368 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 369 | lemma real_of_float_minus[simp]: "real (- a) = - real (a :: float)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 370 | by (cases a, simp add: uminus_float.simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 371 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 372 | lemma real_of_float_sub[simp]: "real (a - b) = real a - real (b :: float)" | 
| 30960 | 373 | by (cases a, cases b, simp add: minus_float_def) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 374 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 375 | lemma real_of_float_mult[simp]: "real (a*b) = real a * real (b :: float)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 376 | by (cases a, cases b, simp add: times_float.simps pow2_add) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 377 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 378 | lemma real_of_float_0[simp]: "real (0 :: float) = 0" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 379 | by (auto simp add: zero_float_def float_zero) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 380 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 381 | lemma real_of_float_1[simp]: "real (1 :: float) = 1" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 382 | by (auto simp add: one_float_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 383 | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 384 | lemma zero_le_float: | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 385 | "(0 <= real (Float a b)) = (0 <= a)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 386 | apply auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 387 | apply (auto simp add: zero_le_mult_iff) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 388 | apply (insert zero_less_pow2[of b]) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 389 | apply (simp_all) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 390 | done | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 391 | |
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 392 | lemma float_le_zero: | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 393 | "(real (Float a b) <= 0) = (a <= 0)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 394 | apply auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 395 | apply (auto simp add: mult_le_0_iff) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 396 | apply (insert zero_less_pow2[of b]) | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 397 | apply auto | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 398 | done | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 399 | |
| 39161 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 400 | lemma zero_less_float: | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 401 | "(0 < real (Float a b)) = (0 < a)" | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 402 | apply auto | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 403 | apply (auto simp add: zero_less_mult_iff) | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 404 | apply (insert zero_less_pow2[of b]) | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 405 | apply (simp_all) | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 406 | done | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 407 | |
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 408 | lemma float_less_zero: | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 409 | "(real (Float a b) < 0) = (a < 0)" | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 410 | apply auto | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 411 | apply (auto simp add: mult_less_0_iff) | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 412 | apply (insert zero_less_pow2[of b]) | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 413 | apply (simp_all) | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 414 | done | 
| 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 415 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 416 | declare real_of_float_simp[simp del] | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 417 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 418 | lemma real_of_float_pprt[simp]: "real (float_pprt a) = pprt (real a)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 419 | by (cases a, auto simp add: float_pprt.simps zero_le_float float_le_zero float_zero) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 420 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 421 | lemma real_of_float_nprt[simp]: "real (float_nprt a) = nprt (real a)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 422 | by (cases a, auto simp add: float_nprt.simps zero_le_float float_le_zero float_zero) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 423 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 424 | instance float :: ab_semigroup_add | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 425 | proof (intro_classes) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 426 | fix a b c :: float | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 427 | show "a + b + c = a + (b + c)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 428 | by (cases a, cases b, cases c, auto simp add: algebra_simps power_add[symmetric] plus_float.simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 429 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 430 | fix a b :: float | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 431 | show "a + b = b + a" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 432 | by (cases a, cases b, simp add: plus_float.simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 433 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 434 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 435 | instance float :: comm_monoid_mult | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 436 | proof (intro_classes) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 437 | fix a b c :: float | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 438 | show "a * b * c = a * (b * c)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 439 | by (cases a, cases b, cases c, simp add: times_float.simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 440 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 441 | fix a b :: float | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 442 | show "a * b = b * a" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 443 | by (cases a, cases b, simp add: times_float.simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 444 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 445 | fix a :: float | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 446 | show "1 * a = a" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 447 | by (cases a, simp add: times_float.simps one_float_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 448 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 449 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 450 | (* Floats do NOT form a cancel_semigroup_add: *) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 451 | lemma "0 + Float 0 1 = 0 + Float 0 2" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 452 | by (simp add: plus_float.simps zero_float_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 453 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 454 | instance float :: comm_semiring | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 455 | proof (intro_classes) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 456 | fix a b c :: float | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 457 | show "(a + b) * c = a * c + b * c" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 458 | by (cases a, cases b, cases c, simp, simp add: plus_float.simps times_float.simps algebra_simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 459 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 460 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 461 | (* Floats do NOT form an order, because "(x < y) = (x <= y & x <> y)" does NOT hold *) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 462 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 463 | instance float :: zero_neq_one | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 464 | proof (intro_classes) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 465 | show "(0::float) \<noteq> 1" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 466 | by (simp add: zero_float_def one_float_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 467 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 468 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 469 | lemma float_le_simp: "((x::float) \<le> y) = (0 \<le> y - x)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 470 | by (auto simp add: le_float_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 471 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 472 | lemma float_less_simp: "((x::float) < y) = (0 < y - x)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 473 | by (auto simp add: less_float_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 474 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 475 | lemma real_of_float_min: "real (min x y :: float) = min (real x) (real y)" unfolding min_def le_float_def by auto | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 476 | lemma real_of_float_max: "real (max a b :: float) = max (real a) (real b)" unfolding max_def le_float_def by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 477 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 478 | lemma float_power: "real (x ^ n :: float) = real x ^ n" | 
| 30960 | 479 | by (induct n) simp_all | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 480 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 481 | lemma zero_le_pow2[simp]: "0 \<le> pow2 s" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 482 | apply (subgoal_tac "0 < pow2 s") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 483 | apply (auto simp only:) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 484 | apply auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 485 | done | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 486 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 487 | lemma pow2_less_0_eq_False[simp]: "(pow2 s < 0) = False" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 488 | apply auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 489 | apply (subgoal_tac "0 \<le> pow2 s") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 490 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 491 | apply simp | 
| 24301 | 492 | done | 
| 493 | ||
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 494 | lemma pow2_le_0_eq_False[simp]: "(pow2 s \<le> 0) = False" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 495 | apply auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 496 | apply (subgoal_tac "0 < pow2 s") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 497 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 498 | apply simp | 
| 24301 | 499 | done | 
| 500 | ||
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 501 | lemma float_pos_m_pos: "0 < Float m e \<Longrightarrow> 0 < m" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 502 | unfolding less_float_def real_of_float_simp real_of_float_0 zero_less_mult_iff | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 503 | by auto | 
| 19765 | 504 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 505 | lemma float_pos_less1_e_neg: assumes "0 < Float m e" and "Float m e < 1" shows "e < 0" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 506 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 507 | have "0 < m" using float_pos_m_pos `0 < Float m e` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 508 | hence "0 \<le> real m" and "1 \<le> real m" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 509 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 510 | show "e < 0" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 511 | proof (rule ccontr) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 512 | assume "\<not> e < 0" hence "0 \<le> e" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 513 | hence "1 \<le> pow2 e" unfolding pow2_def by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 514 | from mult_mono[OF `1 \<le> real m` this `0 \<le> real m`] | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 515 | have "1 \<le> Float m e" by (simp add: le_float_def real_of_float_simp) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 516 | thus False using `Float m e < 1` unfolding less_float_def le_float_def by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 517 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 518 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 519 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 520 | lemma float_less1_mantissa_bound: assumes "0 < Float m e" "Float m e < 1" shows "m < 2^(nat (-e))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 521 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 522 | have "e < 0" using float_pos_less1_e_neg assms by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 523 | have "\<And>x. (0::real) < 2^x" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 524 | have "real m < 2^(nat (-e))" using `Float m e < 1` | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 525 | unfolding less_float_def real_of_float_neg_exp[OF `e < 0`] real_of_float_1 | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 526 | real_mult_less_iff1[of _ _ 1, OF `0 < 2^(nat (-e))`, symmetric] | 
| 36778 
739a9379e29b
avoid using real-specific versions of generic lemmas
 huffman parents: 
35344diff
changeset | 527 | mult_assoc by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 528 | thus ?thesis unfolding real_of_int_less_iff[symmetric] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 529 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 530 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 531 | function bitlen :: "int \<Rightarrow> int" where | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 532 | "bitlen 0 = 0" | | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 533 | "bitlen -1 = 1" | | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 534 | "0 < x \<Longrightarrow> bitlen x = 1 + (bitlen (x div 2))" | | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 535 | "x < -1 \<Longrightarrow> bitlen x = 1 + (bitlen (x div 2))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 536 | apply (case_tac "x = 0 \<or> x = -1 \<or> x < -1 \<or> x > 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 537 | apply auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 538 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 539 | termination by (relation "measure (nat o abs)", auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 540 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 541 | lemma bitlen_ge0: "0 \<le> bitlen x" by (induct x rule: bitlen.induct, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 542 | lemma bitlen_ge1: "x \<noteq> 0 \<Longrightarrow> 1 \<le> bitlen x" by (induct x rule: bitlen.induct, auto simp add: bitlen_ge0) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 543 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 544 | lemma bitlen_bounds': assumes "0 < x" shows "2^nat (bitlen x - 1) \<le> x \<and> x + 1 \<le> 2^nat (bitlen x)" (is "?P x") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 545 | using `0 < x` | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 546 | proof (induct x rule: bitlen.induct) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 547 | fix x | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 548 | assume "0 < x" and hyp: "0 < x div 2 \<Longrightarrow> ?P (x div 2)" hence "0 \<le> x" and "x \<noteq> 0" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 549 |   { fix x have "0 \<le> 1 + bitlen x" using bitlen_ge0[of x] by auto } note gt0_pls1 = this
 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 550 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 551 | have "0 < (2::int)" by auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 552 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 553 | show "?P x" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 554 | proof (cases "x = 1") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 555 | case True show "?P x" unfolding True by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 556 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 557 | case False hence "2 \<le> x" using `0 < x` `x \<noteq> 1` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 558 | hence "2 div 2 \<le> x div 2" by (rule zdiv_mono1, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 559 | hence "0 < x div 2" and "x div 2 \<noteq> 0" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 560 | hence bitlen_s1_ge0: "0 \<le> bitlen (x div 2) - 1" using bitlen_ge1[OF `x div 2 \<noteq> 0`] by auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 561 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 562 |     { from hyp[OF `0 < x div 2`]
 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 563 | have "2 ^ nat (bitlen (x div 2) - 1) \<le> x div 2" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 564 | hence "2 ^ nat (bitlen (x div 2) - 1) * 2 \<le> x div 2 * 2" by (rule mult_right_mono, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 565 | also have "\<dots> \<le> x" using `0 < x` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 566 | finally have "2^nat (1 + bitlen (x div 2) - 1) \<le> x" unfolding power_Suc2[symmetric] Suc_nat_eq_nat_zadd1[OF bitlen_s1_ge0] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 567 | } moreover | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 568 |     { have "x + 1 \<le> x - x mod 2 + 2"
 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 569 | proof - | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 570 | have "x mod 2 < 2" using `0 < x` by auto | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 571 | hence "x < x - x mod 2 + 2" unfolding algebra_simps by auto | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 572 | thus ?thesis by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 573 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 574 | also have "x - x mod 2 + 2 = (x div 2 + 1) * 2" unfolding algebra_simps using `0 < x` zdiv_zmod_equality2[of x 2 0] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 575 | also have "\<dots> \<le> 2^nat (bitlen (x div 2)) * 2" using hyp[OF `0 < x div 2`, THEN conjunct2] by (rule mult_right_mono, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 576 | also have "\<dots> = 2^(1 + nat (bitlen (x div 2)))" unfolding power_Suc2[symmetric] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 577 | finally have "x + 1 \<le> 2^(1 + nat (bitlen (x div 2)))" . | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 578 | } | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 579 | ultimately show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 580 | unfolding bitlen.simps(3)[OF `0 < x`] nat_add_distrib[OF zero_le_one bitlen_ge0] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 581 | unfolding add_commute nat_add_distrib[OF zero_le_one gt0_pls1] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 582 | by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 583 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 584 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 585 | fix x :: int assume "x < -1" and "0 < x" hence False by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 586 | thus "?P x" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 587 | qed auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 588 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 589 | lemma bitlen_bounds: assumes "0 < x" shows "2^nat (bitlen x - 1) \<le> x \<and> x < 2^nat (bitlen x)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 590 | using bitlen_bounds'[OF `0<x`] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 591 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 592 | lemma bitlen_div: assumes "0 < m" shows "1 \<le> real m / 2^nat (bitlen m - 1)" and "real m / 2^nat (bitlen m - 1) < 2" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 593 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 594 | let ?B = "2^nat(bitlen m - 1)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 595 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 596 | have "?B \<le> m" using bitlen_bounds[OF `0 <m`] .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 597 | hence "1 * ?B \<le> real m" unfolding real_of_int_le_iff[symmetric] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 598 | thus "1 \<le> real m / ?B" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 599 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 600 | have "m \<noteq> 0" using assms by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 601 | have "0 \<le> bitlen m - 1" using bitlen_ge1[OF `m \<noteq> 0`] by auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 602 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 603 | have "m < 2^nat(bitlen m)" using bitlen_bounds[OF `0 <m`] .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 604 | also have "\<dots> = 2^nat(bitlen m - 1 + 1)" using bitlen_ge1[OF `m \<noteq> 0`] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 605 | also have "\<dots> = ?B * 2" unfolding nat_add_distrib[OF `0 \<le> bitlen m - 1` zero_le_one] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 606 | finally have "real m < 2 * ?B" unfolding real_of_int_less_iff[symmetric] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 607 | hence "real m / ?B < 2 * ?B / ?B" by (rule divide_strict_right_mono, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 608 | thus "real m / ?B < 2" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 609 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 610 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 611 | lemma float_gt1_scale: assumes "1 \<le> Float m e" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 612 | shows "0 \<le> e + (bitlen m - 1)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 613 | proof (cases "0 \<le> e") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 614 | have "0 < Float m e" using assms unfolding less_float_def le_float_def by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 615 | hence "0 < m" using float_pos_m_pos by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 616 | hence "m \<noteq> 0" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 617 | case True with bitlen_ge1[OF `m \<noteq> 0`] show ?thesis by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 618 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 619 | have "0 < Float m e" using assms unfolding less_float_def le_float_def by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 620 | hence "0 < m" using float_pos_m_pos by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 621 | hence "m \<noteq> 0" and "1 < (2::int)" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 622 | case False let ?S = "2^(nat (-e))" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 623 | have "1 \<le> real m * inverse ?S" using assms unfolding le_float_def real_of_float_nge0_exp[OF False] by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 624 | hence "1 * ?S \<le> real m * inverse ?S * ?S" by (rule mult_right_mono, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 625 | hence "?S \<le> real m" unfolding mult_assoc by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 626 | hence "?S \<le> m" unfolding real_of_int_le_iff[symmetric] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 627 | from this bitlen_bounds[OF `0 < m`, THEN conjunct2] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 628 | have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF `1 < 2`, symmetric] by (rule order_le_less_trans) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 629 | hence "-e < bitlen m" using False bitlen_ge0 by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 630 | thus ?thesis by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 631 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 632 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 633 | lemma normalized_float: assumes "m \<noteq> 0" shows "real (Float m (- (bitlen m - 1))) = real m / 2^nat (bitlen m - 1)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 634 | proof (cases "- (bitlen m - 1) = 0") | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 635 | case True show ?thesis unfolding real_of_float_simp pow2_def using True by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 636 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 637 | case False hence P: "\<not> 0 \<le> - (bitlen m - 1)" using bitlen_ge1[OF `m \<noteq> 0`] by auto | 
| 36778 
739a9379e29b
avoid using real-specific versions of generic lemmas
 huffman parents: 
35344diff
changeset | 638 | show ?thesis unfolding real_of_float_nge0_exp[OF P] divide_inverse by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 639 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 640 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 641 | lemma bitlen_Pls: "bitlen (Int.Pls) = Int.Pls" by (subst Pls_def, subst Pls_def, simp) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 642 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 643 | lemma bitlen_Min: "bitlen (Int.Min) = Int.Bit1 Int.Pls" by (subst Min_def, simp add: Bit1_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 644 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 645 | lemma bitlen_B0: "bitlen (Int.Bit0 b) = (if iszero b then Int.Pls else Int.succ (bitlen b))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 646 | apply (auto simp add: iszero_def succ_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 647 | apply (simp add: Bit0_def Pls_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 648 | apply (subst Bit0_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 649 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 650 | apply (subgoal_tac "0 < 2 * b \<or> 2 * b < -1") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 651 | apply auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 652 | done | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 653 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 654 | lemma bitlen_B1: "bitlen (Int.Bit1 b) = (if iszero (Int.succ b) then Int.Bit1 Int.Pls else Int.succ (bitlen b))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 655 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 656 | have h: "! x. (2*x + 1) div 2 = (x::int)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 657 | by arith | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 658 | show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 659 | apply (auto simp add: iszero_def succ_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 660 | apply (subst Bit1_def)+ | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 661 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 662 | apply (subgoal_tac "2 * b + 1 = -1") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 663 | apply (simp only:) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 664 | apply simp_all | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 665 | apply (subst Bit1_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 666 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 667 | apply (subgoal_tac "0 < 2 * b + 1 \<or> 2 * b + 1 < -1") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 668 | apply (auto simp add: h) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 669 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 670 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 671 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 672 | lemma bitlen_number_of: "bitlen (number_of w) = number_of (bitlen w)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 673 | by (simp add: number_of_is_id) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 674 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 675 | lemma [code]: "bitlen x = | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 676 | (if x = 0 then 0 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 677 | else if x = -1 then 1 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 678 | else (1 + (bitlen (x div 2))))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 679 | by (cases "x = 0 \<or> x = -1 \<or> 0 < x") auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 680 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 681 | definition lapprox_posrat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 682 | where | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 683 | "lapprox_posrat prec x y = | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 684 | (let | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 685 | l = nat (int prec + bitlen y - bitlen x) ; | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 686 | d = (x * 2^l) div y | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 687 | in normfloat (Float d (- (int l))))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 688 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 689 | lemma pow2_minus: "pow2 (-x) = inverse (pow2 x)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 690 | unfolding pow2_neg[of "-x"] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 691 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 692 | lemma lapprox_posrat: | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 693 | assumes x: "0 \<le> x" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 694 | and y: "0 < y" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 695 | shows "real (lapprox_posrat prec x y) \<le> real x / real y" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 696 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 697 | let ?l = "nat (int prec + bitlen y - bitlen x)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 698 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 699 | have "real (x * 2^?l div y) * inverse (2^?l) \<le> (real (x * 2^?l) / real y) * inverse (2^?l)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 700 | by (rule mult_right_mono, fact real_of_int_div4, simp) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 701 | also have "\<dots> \<le> (real x / real y) * 2^?l * inverse (2^?l)" by auto | 
| 36778 
739a9379e29b
avoid using real-specific versions of generic lemmas
 huffman parents: 
35344diff
changeset | 702 | finally have "real (x * 2^?l div y) * inverse (2^?l) \<le> real x / real y" unfolding mult_assoc by auto | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 703 | thus ?thesis unfolding lapprox_posrat_def Let_def normfloat real_of_float_simp | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 704 | unfolding pow2_minus pow2_int minus_minus . | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 705 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 706 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 707 | lemma real_of_int_div_mult: | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 708 | fixes x y c :: int assumes "0 < y" and "0 < c" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 709 | shows "real (x div y) \<le> real (x * c div y) * inverse (real c)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 710 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 711 | have "c * (x div y) + 0 \<le> c * x div y" unfolding zdiv_zmult1_eq[of c x y] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 712 | by (rule zadd_left_mono, | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 713 | auto intro!: mult_nonneg_nonneg | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 714 | simp add: pos_imp_zdiv_nonneg_iff[OF `0 < y`] `0 < c`[THEN less_imp_le] pos_mod_sign[OF `0 < y`]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 715 | hence "real (x div y) * real c \<le> real (x * c div y)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 716 | unfolding real_of_int_mult[symmetric] real_of_int_le_iff zmult_commute by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 717 | hence "real (x div y) * real c * inverse (real c) \<le> real (x * c div y) * inverse (real c)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 718 | using `0 < c` by auto | 
| 36778 
739a9379e29b
avoid using real-specific versions of generic lemmas
 huffman parents: 
35344diff
changeset | 719 | thus ?thesis unfolding mult_assoc using `0 < c` by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 720 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 721 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 722 | lemma lapprox_posrat_bottom: assumes "0 < y" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 723 | shows "real (x div y) \<le> real (lapprox_posrat n x y)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 724 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 725 | have pow: "\<And>x. (0::int) < 2^x" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 726 | show ?thesis | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 727 | unfolding lapprox_posrat_def Let_def real_of_float_add normfloat real_of_float_simp pow2_minus pow2_int | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 728 | using real_of_int_div_mult[OF `0 < y` pow] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 729 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 730 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 731 | lemma lapprox_posrat_nonneg: assumes "0 \<le> x" and "0 < y" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 732 | shows "0 \<le> real (lapprox_posrat n x y)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 733 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 734 | show ?thesis | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 735 | unfolding lapprox_posrat_def Let_def real_of_float_add normfloat real_of_float_simp pow2_minus pow2_int | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 736 | using pos_imp_zdiv_nonneg_iff[OF `0 < y`] assms by (auto intro!: mult_nonneg_nonneg) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 737 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 738 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 739 | definition rapprox_posrat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 740 | where | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 741 | "rapprox_posrat prec x y = (let | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 742 | l = nat (int prec + bitlen y - bitlen x) ; | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 743 | X = x * 2^l ; | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 744 | d = X div y ; | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 745 | m = X mod y | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 746 | in normfloat (Float (d + (if m = 0 then 0 else 1)) (- (int l))))" | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 747 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 748 | lemma rapprox_posrat: | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 749 | assumes x: "0 \<le> x" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 750 | and y: "0 < y" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 751 | shows "real x / real y \<le> real (rapprox_posrat prec x y)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 752 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 753 | let ?l = "nat (int prec + bitlen y - bitlen x)" let ?X = "x * 2^?l" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 754 | show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 755 | proof (cases "?X mod y = 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 756 | case True hence "y \<noteq> 0" and "y dvd ?X" using `0 < y` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 757 | from real_of_int_div[OF this] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 758 | have "real (?X div y) * inverse (2 ^ ?l) = real ?X / real y * inverse (2 ^ ?l)" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 759 | also have "\<dots> = real x / real y * (2^?l * inverse (2^?l))" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 760 | finally have "real (?X div y) * inverse (2^?l) = real x / real y" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 761 | thus ?thesis unfolding rapprox_posrat_def Let_def normfloat if_P[OF True] | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 762 | unfolding real_of_float_simp pow2_minus pow2_int minus_minus by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 763 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 764 | case False | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 765 | have "0 \<le> real y" and "real y \<noteq> 0" using `0 < y` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 766 | have "0 \<le> real y * 2^?l" by (rule mult_nonneg_nonneg, rule `0 \<le> real y`, auto) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 767 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 768 | have "?X = y * (?X div y) + ?X mod y" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 769 | also have "\<dots> \<le> y * (?X div y) + y" by (rule add_mono, auto simp add: pos_mod_bound[OF `0 < y`, THEN less_imp_le]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 770 | also have "\<dots> = y * (?X div y + 1)" unfolding zadd_zmult_distrib2 by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 771 | finally have "real ?X \<le> real y * real (?X div y + 1)" unfolding real_of_int_le_iff real_of_int_mult[symmetric] . | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 772 | hence "real ?X / (real y * 2^?l) \<le> real y * real (?X div y + 1) / (real y * 2^?l)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 773 | by (rule divide_right_mono, simp only: `0 \<le> real y * 2^?l`) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 774 | also have "\<dots> = real y * real (?X div y + 1) / real y / 2^?l" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 775 | also have "\<dots> = real (?X div y + 1) * inverse (2^?l)" unfolding nonzero_mult_divide_cancel_left[OF `real y \<noteq> 0`] | 
| 36778 
739a9379e29b
avoid using real-specific versions of generic lemmas
 huffman parents: 
35344diff
changeset | 776 | unfolding divide_inverse .. | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 777 | finally show ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_not_P[OF False] | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 778 | unfolding pow2_minus pow2_int minus_minus by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 779 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 780 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 781 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 782 | lemma rapprox_posrat_le1: assumes "0 \<le> x" and "0 < y" and "x \<le> y" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 783 | shows "real (rapprox_posrat n x y) \<le> 1" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 784 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 785 | let ?l = "nat (int n + bitlen y - bitlen x)" let ?X = "x * 2^?l" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 786 | show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 787 | proof (cases "?X mod y = 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 788 | case True hence "y \<noteq> 0" and "y dvd ?X" using `0 < y` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 789 | from real_of_int_div[OF this] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 790 | have "real (?X div y) * inverse (2 ^ ?l) = real ?X / real y * inverse (2 ^ ?l)" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 791 | also have "\<dots> = real x / real y * (2^?l * inverse (2^?l))" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 792 | finally have "real (?X div y) * inverse (2^?l) = real x / real y" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 793 | also have "real x / real y \<le> 1" using `0 \<le> x` and `0 < y` and `x \<le> y` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 794 | finally show ?thesis unfolding rapprox_posrat_def Let_def normfloat if_P[OF True] | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 795 | unfolding real_of_float_simp pow2_minus pow2_int minus_minus by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 796 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 797 | case False | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 798 | have "x \<noteq> y" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 799 | proof (rule ccontr) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 800 | assume "\<not> x \<noteq> y" hence "x = y" by auto | 
| 30034 | 801 | have "?X mod y = 0" unfolding `x = y` using mod_mult_self1_is_0 by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 802 | thus False using False by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 803 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 804 | hence "x < y" using `x \<le> y` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 805 | hence "real x / real y < 1" using `0 < y` and `0 \<le> x` by auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 806 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 807 | from real_of_int_div4[of "?X" y] | 
| 35344 
e0b46cd72414
moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
 huffman parents: 
35032diff
changeset | 808 | have "real (?X div y) \<le> (real x / real y) * 2^?l" unfolding real_of_int_mult times_divide_eq_left real_of_int_power real_number_of . | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 809 | also have "\<dots> < 1 * 2^?l" using `real x / real y < 1` by (rule mult_strict_right_mono, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 810 | finally have "?X div y < 2^?l" unfolding real_of_int_less_iff[of _ "2^?l", symmetric] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 811 | hence "?X div y + 1 \<le> 2^?l" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 812 | hence "real (?X div y + 1) * inverse (2^?l) \<le> 2^?l * inverse (2^?l)" | 
| 35344 
e0b46cd72414
moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
 huffman parents: 
35032diff
changeset | 813 | unfolding real_of_int_le_iff[of _ "2^?l", symmetric] real_of_int_power real_number_of | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 814 | by (rule mult_right_mono, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 815 | hence "real (?X div y + 1) * inverse (2^?l) \<le> 1" by auto | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 816 | thus ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_not_P[OF False] | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 817 | unfolding pow2_minus pow2_int minus_minus by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 818 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 819 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 820 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 821 | lemma zdiv_greater_zero: fixes a b :: int assumes "0 < a" and "a \<le> b" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 822 | shows "0 < b div a" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 823 | proof (rule ccontr) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 824 | have "0 \<le> b" using assms by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 825 | assume "\<not> 0 < b div a" hence "b div a = 0" using `0 \<le> b`[unfolded pos_imp_zdiv_nonneg_iff[OF `0<a`, of b, symmetric]] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 826 | have "b = a * (b div a) + b mod a" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 827 | hence "b = b mod a" unfolding `b div a = 0` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 828 | hence "b < a" using `0 < a`[THEN pos_mod_bound, of b] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 829 | thus False using `a \<le> b` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 830 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 831 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 832 | lemma rapprox_posrat_less1: assumes "0 \<le> x" and "0 < y" and "2 * x < y" and "0 < n" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 833 | shows "real (rapprox_posrat n x y) < 1" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 834 | proof (cases "x = 0") | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 835 | case True thus ?thesis unfolding rapprox_posrat_def True Let_def normfloat real_of_float_simp by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 836 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 837 | case False hence "0 < x" using `0 \<le> x` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 838 | hence "x < y" using assms by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 839 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 840 | let ?l = "nat (int n + bitlen y - bitlen x)" let ?X = "x * 2^?l" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 841 | show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 842 | proof (cases "?X mod y = 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 843 | case True hence "y \<noteq> 0" and "y dvd ?X" using `0 < y` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 844 | from real_of_int_div[OF this] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 845 | have "real (?X div y) * inverse (2 ^ ?l) = real ?X / real y * inverse (2 ^ ?l)" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 846 | also have "\<dots> = real x / real y * (2^?l * inverse (2^?l))" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 847 | finally have "real (?X div y) * inverse (2^?l) = real x / real y" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 848 | also have "real x / real y < 1" using `0 \<le> x` and `0 < y` and `x < y` by auto | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 849 | finally show ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_P[OF True] | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 850 | unfolding pow2_minus pow2_int minus_minus by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 851 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 852 | case False | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 853 | hence "(real x / real y) < 1 / 2" using `0 < y` and `0 \<le> x` `2 * x < y` by auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 854 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 855 | have "0 < ?X div y" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 856 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 857 | have "2^nat (bitlen x - 1) \<le> y" and "y < 2^nat (bitlen y)" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 858 | using bitlen_bounds[OF `0 < x`, THEN conjunct1] bitlen_bounds[OF `0 < y`, THEN conjunct2] `x < y` by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 859 | hence "(2::int)^nat (bitlen x - 1) < 2^nat (bitlen y)" by (rule order_le_less_trans) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 860 | hence "bitlen x \<le> bitlen y" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 861 | hence len_less: "nat (bitlen x - 1) \<le> nat (int (n - 1) + bitlen y)" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 862 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 863 | have "x \<noteq> 0" and "y \<noteq> 0" using `0 < x` `0 < y` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 864 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 865 | have exp_eq: "nat (int (n - 1) + bitlen y) - nat (bitlen x - 1) = ?l" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 866 | using `bitlen x \<le> bitlen y` bitlen_ge1[OF `x \<noteq> 0`] bitlen_ge1[OF `y \<noteq> 0`] `0 < n` by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 867 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 868 | have "y * 2^nat (bitlen x - 1) \<le> y * x" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 869 | using bitlen_bounds[OF `0 < x`, THEN conjunct1] `0 < y`[THEN less_imp_le] by (rule mult_left_mono) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 870 | also have "\<dots> \<le> 2^nat (bitlen y) * x" using bitlen_bounds[OF `0 < y`, THEN conjunct2, THEN less_imp_le] `0 \<le> x` by (rule mult_right_mono) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 871 | also have "\<dots> \<le> x * 2^nat (int (n - 1) + bitlen y)" unfolding mult_commute[of x] by (rule mult_right_mono, auto simp add: `0 \<le> x`) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 872 | finally have "real y * 2^nat (bitlen x - 1) * inverse (2^nat (bitlen x - 1)) \<le> real x * 2^nat (int (n - 1) + bitlen y) * inverse (2^nat (bitlen x - 1))" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 873 | unfolding real_of_int_le_iff[symmetric] by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 874 | hence "real y \<le> real x * (2^nat (int (n - 1) + bitlen y) / (2^nat (bitlen x - 1)))" | 
| 36778 
739a9379e29b
avoid using real-specific versions of generic lemmas
 huffman parents: 
35344diff
changeset | 875 | unfolding mult_assoc divide_inverse by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 876 | also have "\<dots> = real x * (2^(nat (int (n - 1) + bitlen y) - nat (bitlen x - 1)))" using power_diff[of "2::real", OF _ len_less] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 877 | finally have "y \<le> x * 2^?l" unfolding exp_eq unfolding real_of_int_le_iff[symmetric] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 878 | thus ?thesis using zdiv_greater_zero[OF `0 < y`] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 879 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 880 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 881 | from real_of_int_div4[of "?X" y] | 
| 35344 
e0b46cd72414
moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
 huffman parents: 
35032diff
changeset | 882 | have "real (?X div y) \<le> (real x / real y) * 2^?l" unfolding real_of_int_mult times_divide_eq_left real_of_int_power real_number_of . | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 883 | also have "\<dots> < 1/2 * 2^?l" using `real x / real y < 1/2` by (rule mult_strict_right_mono, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 884 | finally have "?X div y * 2 < 2^?l" unfolding real_of_int_less_iff[of _ "2^?l", symmetric] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 885 | hence "?X div y + 1 < 2^?l" using `0 < ?X div y` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 886 | hence "real (?X div y + 1) * inverse (2^?l) < 2^?l * inverse (2^?l)" | 
| 35344 
e0b46cd72414
moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
 huffman parents: 
35032diff
changeset | 887 | unfolding real_of_int_less_iff[of _ "2^?l", symmetric] real_of_int_power real_number_of | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 888 | by (rule mult_strict_right_mono, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 889 | hence "real (?X div y + 1) * inverse (2^?l) < 1" by auto | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 890 | thus ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_not_P[OF False] | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 891 | unfolding pow2_minus pow2_int minus_minus by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 892 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 893 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 894 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 895 | lemma approx_rat_pattern: fixes P and ps :: "nat * int * int" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 896 | assumes Y: "\<And>y prec x. \<lbrakk>y = 0; ps = (prec, x, 0)\<rbrakk> \<Longrightarrow> P" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 897 | and A: "\<And>x y prec. \<lbrakk>0 \<le> x; 0 < y; ps = (prec, x, y)\<rbrakk> \<Longrightarrow> P" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 898 | and B: "\<And>x y prec. \<lbrakk>x < 0; 0 < y; ps = (prec, x, y)\<rbrakk> \<Longrightarrow> P" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 899 | and C: "\<And>x y prec. \<lbrakk>x < 0; y < 0; ps = (prec, x, y)\<rbrakk> \<Longrightarrow> P" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 900 | and D: "\<And>x y prec. \<lbrakk>0 \<le> x; y < 0; ps = (prec, x, y)\<rbrakk> \<Longrightarrow> P" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 901 | shows P | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 902 | proof - | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 903 | obtain prec x y where [simp]: "ps = (prec, x, y)" by (cases ps, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 904 | from Y have "y = 0 \<Longrightarrow> P" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 905 |   moreover { assume "0 < y" have P proof (cases "0 \<le> x") case True with A and `0 < y` show P by auto next case False with B and `0 < y` show P by auto qed } 
 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 906 |   moreover { assume "y < 0" have P proof (cases "0 \<le> x") case True with D and `y < 0` show P by auto next case False with C and `y < 0` show P by auto qed }
 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 907 | ultimately show P by (cases "y = 0 \<or> 0 < y \<or> y < 0", auto) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 908 | qed | 
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 909 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 910 | function lapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 911 | where | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 912 | "y = 0 \<Longrightarrow> lapprox_rat prec x y = 0" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 913 | | "0 \<le> x \<Longrightarrow> 0 < y \<Longrightarrow> lapprox_rat prec x y = lapprox_posrat prec x y" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 914 | | "x < 0 \<Longrightarrow> 0 < y \<Longrightarrow> lapprox_rat prec x y = - (rapprox_posrat prec (-x) y)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 915 | | "x < 0 \<Longrightarrow> y < 0 \<Longrightarrow> lapprox_rat prec x y = lapprox_posrat prec (-x) (-y)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 916 | | "0 \<le> x \<Longrightarrow> y < 0 \<Longrightarrow> lapprox_rat prec x y = - (rapprox_posrat prec x (-y))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 917 | apply simp_all by (rule approx_rat_pattern) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 918 | termination by lexicographic_order | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 919 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 920 | lemma compute_lapprox_rat[code]: | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 921 | "lapprox_rat prec x y = (if y = 0 then 0 else if 0 \<le> x then (if 0 < y then lapprox_posrat prec x y else - (rapprox_posrat prec x (-y))) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 922 | else (if 0 < y then - (rapprox_posrat prec (-x) y) else lapprox_posrat prec (-x) (-y)))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 923 | by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 924 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 925 | lemma lapprox_rat: "real (lapprox_rat prec x y) \<le> real x / real y" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 926 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 927 | have h[rule_format]: "! a b b'. b' \<le> b \<longrightarrow> a \<le> b' \<longrightarrow> a \<le> (b::real)" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 928 | show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 929 | apply (case_tac "y = 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 930 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 931 | apply (case_tac "0 \<le> x \<and> 0 < y") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 932 | apply (simp add: lapprox_posrat) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 933 | apply (case_tac "x < 0 \<and> 0 < y") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 934 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 935 | apply (subst minus_le_iff) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 936 | apply (rule h[OF rapprox_posrat]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 937 | apply (simp_all) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 938 | apply (case_tac "x < 0 \<and> y < 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 939 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 940 | apply (rule h[OF _ lapprox_posrat]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 941 | apply (simp_all) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 942 | apply (case_tac "0 \<le> x \<and> y < 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 943 | apply (simp) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 944 | apply (subst minus_le_iff) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 945 | apply (rule h[OF rapprox_posrat]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 946 | apply simp_all | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 947 | apply arith | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 948 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 949 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 950 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 951 | lemma lapprox_rat_bottom: assumes "0 \<le> x" and "0 < y" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 952 | shows "real (x div y) \<le> real (lapprox_rat n x y)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 953 | unfolding lapprox_rat.simps(2)[OF assms] using lapprox_posrat_bottom[OF `0<y`] . | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 954 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 955 | function rapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 956 | where | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 957 | "y = 0 \<Longrightarrow> rapprox_rat prec x y = 0" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 958 | | "0 \<le> x \<Longrightarrow> 0 < y \<Longrightarrow> rapprox_rat prec x y = rapprox_posrat prec x y" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 959 | | "x < 0 \<Longrightarrow> 0 < y \<Longrightarrow> rapprox_rat prec x y = - (lapprox_posrat prec (-x) y)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 960 | | "x < 0 \<Longrightarrow> y < 0 \<Longrightarrow> rapprox_rat prec x y = rapprox_posrat prec (-x) (-y)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 961 | | "0 \<le> x \<Longrightarrow> y < 0 \<Longrightarrow> rapprox_rat prec x y = - (lapprox_posrat prec x (-y))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 962 | apply simp_all by (rule approx_rat_pattern) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 963 | termination by lexicographic_order | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 964 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 965 | lemma compute_rapprox_rat[code]: | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 966 | "rapprox_rat prec x y = (if y = 0 then 0 else if 0 \<le> x then (if 0 < y then rapprox_posrat prec x y else - (lapprox_posrat prec x (-y))) else | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 967 | (if 0 < y then - (lapprox_posrat prec (-x) y) else rapprox_posrat prec (-x) (-y)))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 968 | by auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 969 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 970 | lemma rapprox_rat: "real x / real y \<le> real (rapprox_rat prec x y)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 971 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 972 | have h[rule_format]: "! a b b'. b' \<le> b \<longrightarrow> a \<le> b' \<longrightarrow> a \<le> (b::real)" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 973 | show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 974 | apply (case_tac "y = 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 975 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 976 | apply (case_tac "0 \<le> x \<and> 0 < y") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 977 | apply (simp add: rapprox_posrat) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 978 | apply (case_tac "x < 0 \<and> 0 < y") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 979 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 980 | apply (subst le_minus_iff) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 981 | apply (rule h[OF _ lapprox_posrat]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 982 | apply (simp_all) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 983 | apply (case_tac "x < 0 \<and> y < 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 984 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 985 | apply (rule h[OF rapprox_posrat]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 986 | apply (simp_all) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 987 | apply (case_tac "0 \<le> x \<and> y < 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 988 | apply (simp) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 989 | apply (subst le_minus_iff) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 990 | apply (rule h[OF _ lapprox_posrat]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 991 | apply simp_all | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 992 | apply arith | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 993 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 994 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 995 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 996 | lemma rapprox_rat_le1: assumes "0 \<le> x" and "0 < y" and "x \<le> y" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 997 | shows "real (rapprox_rat n x y) \<le> 1" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 998 | unfolding rapprox_rat.simps(2)[OF `0 \<le> x` `0 < y`] using rapprox_posrat_le1[OF assms] . | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 999 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1000 | lemma rapprox_rat_neg: assumes "x < 0" and "0 < y" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1001 | shows "real (rapprox_rat n x y) \<le> 0" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1002 | unfolding rapprox_rat.simps(3)[OF assms] using lapprox_posrat_nonneg[of "-x" y n] assms by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1003 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1004 | lemma rapprox_rat_nonneg_neg: assumes "0 \<le> x" and "y < 0" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1005 | shows "real (rapprox_rat n x y) \<le> 0" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1006 | unfolding rapprox_rat.simps(5)[OF assms] using lapprox_posrat_nonneg[of x "-y" n] assms by auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1007 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1008 | lemma rapprox_rat_nonpos_pos: assumes "x \<le> 0" and "0 < y" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1009 | shows "real (rapprox_rat n x y) \<le> 0" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1010 | proof (cases "x = 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1011 | case True hence "0 \<le> x" by auto show ?thesis unfolding rapprox_rat.simps(2)[OF `0 \<le> x` `0 < y`] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1012 | unfolding True rapprox_posrat_def Let_def by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1013 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1014 | case False hence "x < 0" using assms by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1015 | show ?thesis using rapprox_rat_neg[OF `x < 0` `0 < y`] . | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1016 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1017 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1018 | fun float_divl :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1019 | where | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1020 | "float_divl prec (Float m1 s1) (Float m2 s2) = | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1021 | (let | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1022 | l = lapprox_rat prec m1 m2; | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1023 | f = Float 1 (s1 - s2) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1024 | in | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1025 | f * l)" | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1026 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1027 | lemma float_divl: "real (float_divl prec x y) \<le> real x / real y" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1028 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1029 | from float_split[of x] obtain mx sx where x: "x = Float mx sx" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1030 | from float_split[of y] obtain my sy where y: "y = Float my sy" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1031 | have "real mx / real my \<le> (real mx * pow2 sx / (real my * pow2 sy)) / (pow2 (sx - sy))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1032 | apply (case_tac "my = 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1033 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1034 | apply (case_tac "my > 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1035 | apply (subst pos_le_divide_eq) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1036 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1037 | apply (subst pos_le_divide_eq) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1038 | apply (simp add: mult_pos_pos) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1039 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1040 | apply (subst pow2_add[symmetric]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1041 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1042 | apply (subgoal_tac "my < 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1043 | apply auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1044 | apply (simp add: field_simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1045 | apply (subst pow2_add[symmetric]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1046 | apply (simp add: field_simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1047 | done | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1048 | then have "real (lapprox_rat prec mx my) \<le> (real mx * pow2 sx / (real my * pow2 sy)) / (pow2 (sx - sy))" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1049 | by (rule order_trans[OF lapprox_rat]) | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1050 | then have "real (lapprox_rat prec mx my) * pow2 (sx - sy) \<le> real mx * pow2 sx / (real my * pow2 sy)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1051 | apply (subst pos_le_divide_eq[symmetric]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1052 | apply simp_all | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1053 | done | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1054 | then have "pow2 (sx - sy) * real (lapprox_rat prec mx my) \<le> real mx * pow2 sx / (real my * pow2 sy)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1055 | by (simp add: algebra_simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1056 | then show ?thesis | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1057 | by (simp add: x y Let_def real_of_float_simp) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1058 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1059 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1060 | lemma float_divl_lower_bound: assumes "0 \<le> x" and "0 < y" shows "0 \<le> float_divl prec x y" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1061 | proof (cases x, cases y) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1062 | fix xm xe ym ye :: int | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1063 | assume x_eq: "x = Float xm xe" and y_eq: "y = Float ym ye" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1064 | have "0 \<le> xm" using `0 \<le> x`[unfolded x_eq le_float_def real_of_float_simp real_of_float_0 zero_le_mult_iff] by auto | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1065 | have "0 < ym" using `0 < y`[unfolded y_eq less_float_def real_of_float_simp real_of_float_0 zero_less_mult_iff] by auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1066 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1067 | have "\<And>n. 0 \<le> real (Float 1 n)" unfolding real_of_float_simp using zero_le_pow2 by auto | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1068 | moreover have "0 \<le> real (lapprox_rat prec xm ym)" by (rule order_trans[OF _ lapprox_rat_bottom[OF `0 \<le> xm` `0 < ym`]], auto simp add: `0 \<le> xm` pos_imp_zdiv_nonneg_iff[OF `0 < ym`]) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1069 | ultimately show "0 \<le> float_divl prec x y" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1070 | unfolding x_eq y_eq float_divl.simps Let_def le_float_def real_of_float_0 by (auto intro!: mult_nonneg_nonneg) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1071 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1072 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1073 | lemma float_divl_pos_less1_bound: assumes "0 < x" and "x < 1" and "0 < prec" shows "1 \<le> float_divl prec 1 x" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1074 | proof (cases x) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1075 | case (Float m e) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1076 | from `0 < x` `x < 1` have "0 < m" "e < 0" using float_pos_m_pos float_pos_less1_e_neg unfolding Float by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1077 | let ?b = "nat (bitlen m)" and ?e = "nat (-e)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1078 | have "1 \<le> m" and "m \<noteq> 0" using `0 < m` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1079 | with bitlen_bounds[OF `0 < m`] have "m < 2^?b" and "(2::int) \<le> 2^?b" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1080 | hence "1 \<le> bitlen m" using power_le_imp_le_exp[of "2::int" 1 ?b] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1081 | hence pow_split: "nat (int prec + bitlen m - 1) = (prec - 1) + ?b" using `0 < prec` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1082 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1083 | have pow_not0: "\<And>x. (2::real)^x \<noteq> 0" by auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1084 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1085 | from float_less1_mantissa_bound `0 < x` `x < 1` Float | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1086 | have "m < 2^?e" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1087 | with bitlen_bounds[OF `0 < m`, THEN conjunct1] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1088 | have "(2::int)^nat (bitlen m - 1) < 2^?e" by (rule order_le_less_trans) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1089 | from power_less_imp_less_exp[OF _ this] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1090 | have "bitlen m \<le> - e" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1091 | hence "(2::real)^?b \<le> 2^?e" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1092 | hence "(2::real)^?b * inverse (2^?b) \<le> 2^?e * inverse (2^?b)" by (rule mult_right_mono, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1093 | hence "(1::real) \<le> 2^?e * inverse (2^?b)" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1094 | also | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1095 | let ?d = "real (2 ^ nat (int prec + bitlen m - 1) div m) * inverse (2 ^ nat (int prec + bitlen m - 1))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1096 |   { have "2^(prec - 1) * m \<le> 2^(prec - 1) * 2^?b" using `m < 2^?b`[THEN less_imp_le] by (rule mult_left_mono, auto)
 | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1097 | also have "\<dots> = 2 ^ nat (int prec + bitlen m - 1)" unfolding pow_split zpower_zadd_distrib by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1098 | finally have "2^(prec - 1) * m div m \<le> 2 ^ nat (int prec + bitlen m - 1) div m" using `0 < m` by (rule zdiv_mono1) | 
| 30181 | 1099 | hence "2^(prec - 1) \<le> 2 ^ nat (int prec + bitlen m - 1) div m" unfolding div_mult_self2_is_id[OF `m \<noteq> 0`] . | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1100 | hence "2^(prec - 1) * inverse (2 ^ nat (int prec + bitlen m - 1)) \<le> ?d" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1101 | unfolding real_of_int_le_iff[of "2^(prec - 1)", symmetric] by auto } | 
| 36778 
739a9379e29b
avoid using real-specific versions of generic lemmas
 huffman parents: 
35344diff
changeset | 1102 | from mult_left_mono[OF this[unfolded pow_split power_add inverse_mult_distrib mult_assoc[symmetric] right_inverse[OF pow_not0] mult_1_left], of "2^?e"] | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1103 | have "2^?e * inverse (2^?b) \<le> 2^?e * ?d" unfolding pow_split power_add by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1104 | finally have "1 \<le> 2^?e * ?d" . | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1105 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1106 | have e_nat: "0 - e = int (nat (-e))" using `e < 0` by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1107 | have "bitlen 1 = 1" using bitlen.simps by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1108 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1109 | show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1110 | unfolding one_float_def Float float_divl.simps Let_def lapprox_rat.simps(2)[OF zero_le_one `0 < m`] lapprox_posrat_def `bitlen 1 = 1` | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1111 | unfolding le_float_def real_of_float_mult normfloat real_of_float_simp pow2_minus pow2_int e_nat | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1112 | using `1 \<le> 2^?e * ?d` by (auto simp add: pow2_def) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1113 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1114 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1115 | fun float_divr :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1116 | where | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1117 | "float_divr prec (Float m1 s1) (Float m2 s2) = | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1118 | (let | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1119 | r = rapprox_rat prec m1 m2; | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1120 | f = Float 1 (s1 - s2) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1121 | in | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1122 | f * r)" | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1123 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1124 | lemma float_divr: "real x / real y \<le> real (float_divr prec x y)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1125 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1126 | from float_split[of x] obtain mx sx where x: "x = Float mx sx" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1127 | from float_split[of y] obtain my sy where y: "y = Float my sy" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1128 | have "real mx / real my \<ge> (real mx * pow2 sx / (real my * pow2 sy)) / (pow2 (sx - sy))" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1129 | apply (case_tac "my = 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1130 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1131 | apply (case_tac "my > 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1132 | apply auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1133 | apply (subst pos_divide_le_eq) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1134 | apply (rule mult_pos_pos)+ | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1135 | apply simp_all | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1136 | apply (subst pow2_add[symmetric]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1137 | apply simp | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1138 | apply (subgoal_tac "my < 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1139 | apply auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1140 | apply (simp add: field_simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1141 | apply (subst pow2_add[symmetric]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1142 | apply (simp add: field_simps) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1143 | done | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1144 | then have "real (rapprox_rat prec mx my) \<ge> (real mx * pow2 sx / (real my * pow2 sy)) / (pow2 (sx - sy))" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1145 | by (rule order_trans[OF _ rapprox_rat]) | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1146 | then have "real (rapprox_rat prec mx my) * pow2 (sx - sy) \<ge> real mx * pow2 sx / (real my * pow2 sy)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1147 | apply (subst pos_divide_le_eq[symmetric]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1148 | apply simp_all | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1149 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1150 | then show ?thesis | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1151 | by (simp add: x y Let_def algebra_simps real_of_float_simp) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1152 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1153 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1154 | lemma float_divr_pos_less1_lower_bound: assumes "0 < x" and "x < 1" shows "1 \<le> float_divr prec 1 x" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1155 | proof - | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1156 | have "1 \<le> 1 / real x" using `0 < x` and `x < 1` unfolding less_float_def by auto | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1157 | also have "\<dots> \<le> real (float_divr prec 1 x)" using float_divr[where x=1 and y=x] by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1158 | finally show ?thesis unfolding le_float_def by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1159 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1160 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1161 | lemma float_divr_nonpos_pos_upper_bound: assumes "x \<le> 0" and "0 < y" shows "float_divr prec x y \<le> 0" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1162 | proof (cases x, cases y) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1163 | fix xm xe ym ye :: int | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1164 | assume x_eq: "x = Float xm xe" and y_eq: "y = Float ym ye" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1165 | have "xm \<le> 0" using `x \<le> 0`[unfolded x_eq le_float_def real_of_float_simp real_of_float_0 mult_le_0_iff] by auto | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1166 | have "0 < ym" using `0 < y`[unfolded y_eq less_float_def real_of_float_simp real_of_float_0 zero_less_mult_iff] by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1167 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1168 | have "\<And>n. 0 \<le> real (Float 1 n)" unfolding real_of_float_simp using zero_le_pow2 by auto | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1169 | moreover have "real (rapprox_rat prec xm ym) \<le> 0" using rapprox_rat_nonpos_pos[OF `xm \<le> 0` `0 < ym`] . | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1170 | ultimately show "float_divr prec x y \<le> 0" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1171 | unfolding x_eq y_eq float_divr.simps Let_def le_float_def real_of_float_0 real_of_float_mult by (auto intro!: mult_nonneg_nonpos) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1172 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1173 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1174 | lemma float_divr_nonneg_neg_upper_bound: assumes "0 \<le> x" and "y < 0" shows "float_divr prec x y \<le> 0" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1175 | proof (cases x, cases y) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1176 | fix xm xe ym ye :: int | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1177 | assume x_eq: "x = Float xm xe" and y_eq: "y = Float ym ye" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1178 | have "0 \<le> xm" using `0 \<le> x`[unfolded x_eq le_float_def real_of_float_simp real_of_float_0 zero_le_mult_iff] by auto | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1179 | have "ym < 0" using `y < 0`[unfolded y_eq less_float_def real_of_float_simp real_of_float_0 mult_less_0_iff] by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1180 | hence "0 < - ym" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1181 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1182 | have "\<And>n. 0 \<le> real (Float 1 n)" unfolding real_of_float_simp using zero_le_pow2 by auto | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1183 | moreover have "real (rapprox_rat prec xm ym) \<le> 0" using rapprox_rat_nonneg_neg[OF `0 \<le> xm` `ym < 0`] . | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1184 | ultimately show "float_divr prec x y \<le> 0" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1185 | unfolding x_eq y_eq float_divr.simps Let_def le_float_def real_of_float_0 real_of_float_mult by (auto intro!: mult_nonneg_nonpos) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1186 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1187 | |
| 30960 | 1188 | primrec round_down :: "nat \<Rightarrow> float \<Rightarrow> float" where | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1189 | "round_down prec (Float m e) = (let d = bitlen m - int prec in | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1190 | if 0 < d then let P = 2^nat d ; n = m div P in Float n (e + d) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1191 | else Float m e)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1192 | |
| 30960 | 1193 | primrec round_up :: "nat \<Rightarrow> float \<Rightarrow> float" where | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1194 | "round_up prec (Float m e) = (let d = bitlen m - int prec in | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1195 | if 0 < d then let P = 2^nat d ; n = m div P ; r = m mod P in Float (n + (if r = 0 then 0 else 1)) (e + d) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1196 | else Float m e)" | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1197 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1198 | lemma round_up: "real x \<le> real (round_up prec x)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1199 | proof (cases x) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1200 | case (Float m e) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1201 | let ?d = "bitlen m - int prec" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1202 | let ?p = "(2::int)^nat ?d" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1203 | have "0 < ?p" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1204 | show "?thesis" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1205 | proof (cases "0 < ?d") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1206 | case True | 
| 35344 
e0b46cd72414
moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
 huffman parents: 
35032diff
changeset | 1207 | hence pow_d: "pow2 ?d = real ?p" using pow2_int[symmetric] by simp | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1208 | show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1209 | proof (cases "m mod ?p = 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1210 | case True | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1211 | have m: "m = m div ?p * ?p + 0" unfolding True[symmetric] using zdiv_zmod_equality2[where k=0, unfolded monoid_add_class.add_0_right, symmetric] . | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1212 | have "real (Float m e) = real (Float (m div ?p) (e + ?d))" unfolding real_of_float_simp arg_cong[OF m, of real] | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 1213 | by (auto simp add: pow2_add `0 < ?d` pow_d) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1214 | thus ?thesis | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 1215 | unfolding Float round_up.simps Let_def if_P[OF `m mod ?p = 0`] if_P[OF `0 < ?d`] | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 1216 | by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1217 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1218 | case False | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1219 | have "m = m div ?p * ?p + m mod ?p" unfolding zdiv_zmod_equality2[where k=0, unfolded monoid_add_class.add_0_right] .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1220 | also have "\<dots> \<le> (m div ?p + 1) * ?p" unfolding left_distrib zmult_1 by (rule add_left_mono, rule pos_mod_bound[OF `0 < ?p`, THEN less_imp_le]) | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1221 | finally have "real (Float m e) \<le> real (Float (m div ?p + 1) (e + ?d))" unfolding real_of_float_simp add_commute[of e] | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 1222 | unfolding pow2_add mult_assoc[symmetric] real_of_int_le_iff[of m, symmetric] | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 1223 | by (auto intro!: mult_mono simp add: pow2_add `0 < ?d` pow_d) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1224 | thus ?thesis | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 1225 | unfolding Float round_up.simps Let_def if_not_P[OF `\<not> m mod ?p = 0`] if_P[OF `0 < ?d`] . | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1226 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1227 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1228 | case False | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1229 | show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1230 | unfolding Float round_up.simps Let_def if_not_P[OF False] .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1231 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1232 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1233 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1234 | lemma round_down: "real (round_down prec x) \<le> real x" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1235 | proof (cases x) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1236 | case (Float m e) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1237 | let ?d = "bitlen m - int prec" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1238 | let ?p = "(2::int)^nat ?d" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1239 | have "0 < ?p" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1240 | show "?thesis" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1241 | proof (cases "0 < ?d") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1242 | case True | 
| 35344 
e0b46cd72414
moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
 huffman parents: 
35032diff
changeset | 1243 | hence pow_d: "pow2 ?d = real ?p" using pow2_int[symmetric] by simp | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1244 | have "m div ?p * ?p \<le> m div ?p * ?p + m mod ?p" by (auto simp add: pos_mod_bound[OF `0 < ?p`, THEN less_imp_le]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1245 | also have "\<dots> \<le> m" unfolding zdiv_zmod_equality2[where k=0, unfolded monoid_add_class.add_0_right] .. | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1246 | finally have "real (Float (m div ?p) (e + ?d)) \<le> real (Float m e)" unfolding real_of_float_simp add_commute[of e] | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1247 | unfolding pow2_add mult_assoc[symmetric] real_of_int_le_iff[of _ m, symmetric] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1248 | by (auto intro!: mult_mono simp add: pow2_add `0 < ?d` pow_d) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1249 | thus ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1250 | unfolding Float round_down.simps Let_def if_P[OF `0 < ?d`] . | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1251 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1252 | case False | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1253 | show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1254 | unfolding Float round_down.simps Let_def if_not_P[OF False] .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1255 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1256 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1257 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1258 | definition lb_mult :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1259 | "lb_mult prec x y = (case normfloat (x * y) of Float m e \<Rightarrow> let | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1260 | l = bitlen m - int prec | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1261 | in if l > 0 then Float (m div (2^nat l)) (e + l) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1262 | else Float m e)" | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1263 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1264 | definition ub_mult :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1265 | "ub_mult prec x y = (case normfloat (x * y) of Float m e \<Rightarrow> let | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1266 | l = bitlen m - int prec | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1267 | in if l > 0 then Float (m div (2^nat l) + 1) (e + l) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1268 | else Float m e)" | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1269 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1270 | lemma lb_mult: "real (lb_mult prec x y) \<le> real (x * y)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1271 | proof (cases "normfloat (x * y)") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1272 | case (Float m e) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1273 | hence "odd m \<or> (m = 0 \<and> e = 0)" by (rule normfloat_imp_odd_or_zero) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1274 | let ?l = "bitlen m - int prec" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1275 | have "real (lb_mult prec x y) \<le> real (normfloat (x * y))" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1276 | proof (cases "?l > 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1277 | case False thus ?thesis unfolding lb_mult_def Float Let_def float.cases by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1278 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1279 | case True | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1280 | have "real (m div 2^(nat ?l)) * pow2 ?l \<le> real m" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1281 | proof - | 
| 35344 
e0b46cd72414
moved some lemmas from RealPow to RealDef; changed orientation of real_of_int_power
 huffman parents: 
35032diff
changeset | 1282 | have "real (m div 2^(nat ?l)) * pow2 ?l = real (2^(nat ?l) * (m div 2^(nat ?l)))" unfolding real_of_int_mult real_of_int_power real_number_of unfolding pow2_int[symmetric] | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32069diff
changeset | 1283 | using `?l > 0` by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1284 | also have "\<dots> \<le> real (2^(nat ?l) * (m div 2^(nat ?l)) + m mod 2^(nat ?l))" unfolding real_of_int_add by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1285 | also have "\<dots> = real m" unfolding zmod_zdiv_equality[symmetric] .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1286 | finally show ?thesis by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1287 | qed | 
| 36778 
739a9379e29b
avoid using real-specific versions of generic lemmas
 huffman parents: 
35344diff
changeset | 1288 | thus ?thesis unfolding lb_mult_def Float Let_def float.cases if_P[OF True] real_of_float_simp pow2_add mult_commute mult_assoc by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1289 | qed | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1290 | also have "\<dots> = real (x * y)" unfolding normfloat .. | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1291 | finally show ?thesis . | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1292 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1293 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1294 | lemma ub_mult: "real (x * y) \<le> real (ub_mult prec x y)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1295 | proof (cases "normfloat (x * y)") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1296 | case (Float m e) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1297 | hence "odd m \<or> (m = 0 \<and> e = 0)" by (rule normfloat_imp_odd_or_zero) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1298 | let ?l = "bitlen m - int prec" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1299 | have "real (x * y) = real (normfloat (x * y))" unfolding normfloat .. | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1300 | also have "\<dots> \<le> real (ub_mult prec x y)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1301 | proof (cases "?l > 0") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1302 | case False thus ?thesis unfolding ub_mult_def Float Let_def float.cases by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1303 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1304 | case True | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1305 | have "real m \<le> real (m div 2^(nat ?l) + 1) * pow2 ?l" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1306 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1307 | have "m mod 2^(nat ?l) < 2^(nat ?l)" by (rule pos_mod_bound) auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1308 | hence mod_uneq: "real (m mod 2^(nat ?l)) \<le> 1 * 2^(nat ?l)" unfolding zmult_1 real_of_int_less_iff[symmetric] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1309 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1310 | have "real m = real (2^(nat ?l) * (m div 2^(nat ?l)) + m mod 2^(nat ?l))" unfolding zmod_zdiv_equality[symmetric] .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1311 | also have "\<dots> = real (m div 2^(nat ?l)) * 2^(nat ?l) + real (m mod 2^(nat ?l))" unfolding real_of_int_add by auto | 
| 36778 
739a9379e29b
avoid using real-specific versions of generic lemmas
 huffman parents: 
35344diff
changeset | 1312 | also have "\<dots> \<le> (real (m div 2^(nat ?l)) + 1) * 2^(nat ?l)" unfolding left_distrib using mod_uneq by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1313 | finally show ?thesis unfolding pow2_int[symmetric] using True by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1314 | qed | 
| 36778 
739a9379e29b
avoid using real-specific versions of generic lemmas
 huffman parents: 
35344diff
changeset | 1315 | thus ?thesis unfolding ub_mult_def Float Let_def float.cases if_P[OF True] real_of_float_simp pow2_add mult_commute mult_assoc by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1316 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1317 | finally show ?thesis . | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1318 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1319 | |
| 30960 | 1320 | primrec float_abs :: "float \<Rightarrow> float" where | 
| 1321 | "float_abs (Float m e) = Float \<bar>m\<bar> e" | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1322 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1323 | instantiation float :: abs begin | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1324 | definition abs_float_def: "\<bar>x\<bar> = float_abs x" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1325 | instance .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1326 | end | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1327 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1328 | lemma real_of_float_abs: "real \<bar>x :: float\<bar> = \<bar>real x\<bar>" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1329 | proof (cases x) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1330 | case (Float m e) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1331 | have "\<bar>real m\<bar> * pow2 e = \<bar>real m * pow2 e\<bar>" unfolding abs_mult by auto | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1332 | thus ?thesis unfolding Float abs_float_def float_abs.simps real_of_float_simp by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1333 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1334 | |
| 30960 | 1335 | primrec floor_fl :: "float \<Rightarrow> float" where | 
| 1336 | "floor_fl (Float m e) = (if 0 \<le> e then Float m e | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1337 | else Float (m div (2 ^ (nat (-e)))) 0)" | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1338 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1339 | lemma floor_fl: "real (floor_fl x) \<le> real x" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1340 | proof (cases x) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1341 | case (Float m e) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1342 | show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1343 | proof (cases "0 \<le> e") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1344 | case False | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1345 | hence me_eq: "pow2 (-e) = pow2 (int (nat (-e)))" by auto | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1346 | have "real (Float (m div (2 ^ (nat (-e)))) 0) = real (m div 2 ^ (nat (-e)))" unfolding real_of_float_simp by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1347 | also have "\<dots> \<le> real m / real ((2::int) ^ (nat (-e)))" using real_of_int_div4 . | 
| 36778 
739a9379e29b
avoid using real-specific versions of generic lemmas
 huffman parents: 
35344diff
changeset | 1348 | also have "\<dots> = real m * inverse (2 ^ (nat (-e)))" unfolding real_of_int_power real_number_of divide_inverse .. | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1349 | also have "\<dots> = real (Float m e)" unfolding real_of_float_simp me_eq pow2_int pow2_neg[of e] .. | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1350 | finally show ?thesis unfolding Float floor_fl.simps if_not_P[OF `\<not> 0 \<le> e`] . | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1351 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1352 | case True thus ?thesis unfolding Float by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1353 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1354 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1355 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1356 | lemma floor_pos_exp: assumes floor: "Float m e = floor_fl x" shows "0 \<le> e" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1357 | proof (cases x) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1358 | case (Float mx me) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1359 | from floor[unfolded Float floor_fl.simps] show ?thesis by (cases "0 \<le> me", auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1360 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1361 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1362 | declare floor_fl.simps[simp del] | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1363 | |
| 30960 | 1364 | primrec ceiling_fl :: "float \<Rightarrow> float" where | 
| 1365 | "ceiling_fl (Float m e) = (if 0 \<le> e then Float m e | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1366 | else Float (m div (2 ^ (nat (-e))) + 1) 0)" | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1367 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1368 | lemma ceiling_fl: "real x \<le> real (ceiling_fl x)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1369 | proof (cases x) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1370 | case (Float m e) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1371 | show ?thesis | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1372 | proof (cases "0 \<le> e") | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1373 | case False | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1374 | hence me_eq: "pow2 (-e) = pow2 (int (nat (-e)))" by auto | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1375 | have "real (Float m e) = real m * inverse (2 ^ (nat (-e)))" unfolding real_of_float_simp me_eq pow2_int pow2_neg[of e] .. | 
| 36778 
739a9379e29b
avoid using real-specific versions of generic lemmas
 huffman parents: 
35344diff
changeset | 1376 | also have "\<dots> = real m / real ((2::int) ^ (nat (-e)))" unfolding real_of_int_power real_number_of divide_inverse .. | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1377 | also have "\<dots> \<le> 1 + real (m div 2 ^ (nat (-e)))" using real_of_int_div3[unfolded diff_le_eq] . | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1378 | also have "\<dots> = real (Float (m div (2 ^ (nat (-e))) + 1) 0)" unfolding real_of_float_simp by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1379 | finally show ?thesis unfolding Float ceiling_fl.simps if_not_P[OF `\<not> 0 \<le> e`] . | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1380 | next | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1381 | case True thus ?thesis unfolding Float by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1382 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1383 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1384 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1385 | declare ceiling_fl.simps[simp del] | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1386 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1387 | definition lb_mod :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1388 | "lb_mod prec x ub lb = x - ceiling_fl (float_divr prec x lb) * ub" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1389 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1390 | definition ub_mod :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1391 | "ub_mod prec x ub lb = x - floor_fl (float_divl prec x ub) * lb" | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1392 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1393 | lemma lb_mod: fixes k :: int assumes "0 \<le> real x" and "real k * y \<le> real x" (is "?k * y \<le> ?x") | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1394 | assumes "0 < real lb" "real lb \<le> y" (is "?lb \<le> y") "y \<le> real ub" (is "y \<le> ?ub") | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1395 | shows "real (lb_mod prec x ub lb) \<le> ?x - ?k * y" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1396 | proof - | 
| 33555 
a0a8a40385a2
eliminated some old uses of cumulative prems (!) in proof methods;
 wenzelm parents: 
32960diff
changeset | 1397 | have "?lb \<le> ?ub" using assms by auto | 
| 
a0a8a40385a2
eliminated some old uses of cumulative prems (!) in proof methods;
 wenzelm parents: 
32960diff
changeset | 1398 | have "0 \<le> ?lb" and "?lb \<noteq> 0" using assms by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1399 | have "?k * y \<le> ?x" using assms by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1400 | also have "\<dots> \<le> ?x / ?lb * ?ub" by (metis mult_left_mono[OF `?lb \<le> ?ub` `0 \<le> ?x`] divide_right_mono[OF _ `0 \<le> ?lb` ] times_divide_eq_left nonzero_mult_divide_cancel_right[OF `?lb \<noteq> 0`]) | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1401 | also have "\<dots> \<le> real (ceiling_fl (float_divr prec x lb)) * ?ub" by (metis mult_right_mono order_trans `0 \<le> ?lb` `?lb \<le> ?ub` float_divr ceiling_fl) | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1402 | finally show ?thesis unfolding lb_mod_def real_of_float_sub real_of_float_mult by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1403 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1404 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1405 | lemma ub_mod: fixes k :: int and x :: float assumes "0 \<le> real x" and "real x \<le> real k * y" (is "?x \<le> ?k * y") | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1406 | assumes "0 < real lb" "real lb \<le> y" (is "?lb \<le> y") "y \<le> real ub" (is "y \<le> ?ub") | 
| 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1407 | shows "?x - ?k * y \<le> real (ub_mod prec x ub lb)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1408 | proof - | 
| 33555 
a0a8a40385a2
eliminated some old uses of cumulative prems (!) in proof methods;
 wenzelm parents: 
32960diff
changeset | 1409 | have "?lb \<le> ?ub" using assms by auto | 
| 
a0a8a40385a2
eliminated some old uses of cumulative prems (!) in proof methods;
 wenzelm parents: 
32960diff
changeset | 1410 | hence "0 \<le> ?lb" and "0 \<le> ?ub" and "?ub \<noteq> 0" using assms by auto | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1411 | have "real (floor_fl (float_divl prec x ub)) * ?lb \<le> ?x / ?ub * ?lb" by (metis mult_right_mono order_trans `0 \<le> ?lb` `?lb \<le> ?ub` float_divl floor_fl) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1412 | also have "\<dots> \<le> ?x" by (metis mult_left_mono[OF `?lb \<le> ?ub` `0 \<le> ?x`] divide_right_mono[OF _ `0 \<le> ?ub` ] times_divide_eq_left nonzero_mult_divide_cancel_right[OF `?ub \<noteq> 0`]) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1413 | also have "\<dots> \<le> ?k * y" using assms by auto | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1414 | finally show ?thesis unfolding ub_mod_def real_of_float_sub real_of_float_mult by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1415 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1416 | |
| 39161 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 1417 | lemma le_float_def'[code]: "f \<le> g = (case f - g of Float a b \<Rightarrow> a \<le> 0)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1418 | proof - | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1419 | have le_transfer: "(f \<le> g) = (real (f - g) \<le> 0)" by (auto simp add: le_float_def) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1420 | from float_split[of "f - g"] obtain a b where f_diff_g: "f - g = Float a b" by auto | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1421 | with le_transfer have le_transfer': "f \<le> g = (real (Float a b) \<le> 0)" by simp | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1422 | show ?thesis by (simp add: le_transfer' f_diff_g float_le_zero) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1423 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1424 | |
| 39161 
75849a560c09
When comparing Floats use integers instead of reals (represented as rationals), generates less code when Floats are used.
 hoelzl parents: 
36778diff
changeset | 1425 | lemma less_float_def'[code]: "f < g = (case f - g of Float a b \<Rightarrow> a < 0)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1426 | proof - | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1427 | have less_transfer: "(f < g) = (real (f - g) < 0)" by (auto simp add: less_float_def) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1428 | from float_split[of "f - g"] obtain a b where f_diff_g: "f - g = Float a b" by auto | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1429 | with less_transfer have less_transfer': "f < g = (real (Float a b) < 0)" by simp | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1430 | show ?thesis by (simp add: less_transfer' f_diff_g float_less_zero) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1431 | qed | 
| 20771 | 1432 | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1433 | end |