| author | wenzelm | 
| Tue, 29 Mar 2016 23:45:28 +0200 | |
| changeset 62755 | 7fde2461f9ef | 
| parent 62421 | 28d2c75dd180 | 
| child 63040 | eb4ddd18d635 | 
| permissions | -rw-r--r-- | 
| 47615 | 1 | (* Title: HOL/Library/Float.thy | 
| 2 | Author: Johannes Hölzl, Fabian Immler | |
| 3 | Copyright 2012 TU München | |
| 4 | *) | |
| 5 | ||
| 60500 | 6 | section \<open>Floating-Point Numbers\<close> | 
| 29988 | 7 | |
| 20485 | 8 | theory Float | 
| 51542 | 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 | |
| 49812 
e3945ddcb9aa
eliminated some remaining uses of typedef with implicit set definition;
 wenzelm parents: 
47937diff
changeset | 12 | definition "float = {m * 2 powr e | (m :: int) (e :: int). True}"
 | 
| 
e3945ddcb9aa
eliminated some remaining uses of typedef with implicit set definition;
 wenzelm parents: 
47937diff
changeset | 13 | |
| 49834 | 14 | typedef float = float | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 15 | morphisms real_of_float float_of | 
| 49812 
e3945ddcb9aa
eliminated some remaining uses of typedef with implicit set definition;
 wenzelm parents: 
47937diff
changeset | 16 | unfolding float_def by auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 17 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 18 | setup_lifting type_definition_float | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 19 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 20 | declare real_of_float [code_unfold] | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 21 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 22 | lemmas float_of_inject[simp] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 23 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 24 | declare [[coercion "real_of_float :: float \<Rightarrow> real"]] | 
| 47600 | 25 | |
| 26 | lemma real_of_float_eq: | |
| 60698 | 27 | fixes f1 f2 :: float | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 28 | shows "f1 = f2 \<longleftrightarrow> real_of_float f1 = real_of_float f2" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 29 | unfolding real_of_float_inject .. | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 30 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 31 | declare real_of_float_inverse[simp] float_of_inverse [simp] | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 32 | declare real_of_float [simp] | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 33 | |
| 60500 | 34 | subsection \<open>Real operations preserving the representation as floating point number\<close> | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 35 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 36 | lemma floatI: fixes m e :: int shows "m * 2 powr e = x \<Longrightarrow> x \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 37 | by (auto simp: float_def) | 
| 19765 | 38 | |
| 60698 | 39 | lemma zero_float[simp]: "0 \<in> float" | 
| 40 | by (auto simp: float_def) | |
| 41 | lemma one_float[simp]: "1 \<in> float" | |
| 42 | by (intro floatI[of 1 0]) simp | |
| 43 | lemma numeral_float[simp]: "numeral i \<in> float" | |
| 44 | by (intro floatI[of "numeral i" 0]) simp | |
| 45 | lemma neg_numeral_float[simp]: "- numeral i \<in> float" | |
| 46 | by (intro floatI[of "- numeral i" 0]) simp | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 47 | lemma real_of_int_float[simp]: "real_of_int (x :: int) \<in> float" | 
| 60698 | 48 | by (intro floatI[of x 0]) simp | 
| 49 | lemma real_of_nat_float[simp]: "real (x :: nat) \<in> float" | |
| 50 | by (intro floatI[of x 0]) simp | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 51 | lemma two_powr_int_float[simp]: "2 powr (real_of_int (i::int)) \<in> float" | 
| 60698 | 52 | by (intro floatI[of 1 i]) simp | 
| 53 | lemma two_powr_nat_float[simp]: "2 powr (real (i::nat)) \<in> float" | |
| 54 | by (intro floatI[of 1 i]) simp | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 55 | lemma two_powr_minus_int_float[simp]: "2 powr - (real_of_int (i::int)) \<in> float" | 
| 60698 | 56 | by (intro floatI[of 1 "-i"]) simp | 
| 57 | lemma two_powr_minus_nat_float[simp]: "2 powr - (real (i::nat)) \<in> float" | |
| 58 | by (intro floatI[of 1 "-i"]) simp | |
| 59 | lemma two_powr_numeral_float[simp]: "2 powr numeral i \<in> float" | |
| 60 | by (intro floatI[of 1 "numeral i"]) simp | |
| 61 | lemma two_powr_neg_numeral_float[simp]: "2 powr - numeral i \<in> float" | |
| 62 | by (intro floatI[of 1 "- numeral i"]) simp | |
| 63 | lemma two_pow_float[simp]: "2 ^ n \<in> float" | |
| 64 | by (intro floatI[of 1 "n"]) (simp add: powr_realpow) | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 65 | |
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 66 | |
| 45495 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 hoelzl parents: 
44766diff
changeset | 67 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 68 | lemma plus_float[simp]: "r \<in> float \<Longrightarrow> p \<in> float \<Longrightarrow> r + p \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 69 | unfolding float_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 70 | proof (safe, simp) | 
| 60698 | 71 | have *: "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e" | 
| 72 | if "e1 \<le> e2" for e1 m1 e2 m2 :: int | |
| 73 | proof - | |
| 74 | from that have "m1 * 2 powr e1 + m2 * 2 powr e2 = (m1 + m2 * 2 ^ nat (e2 - e1)) * 2 powr e1" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 75 | by (simp add: powr_realpow[symmetric] powr_divide2[symmetric] field_simps) | 
| 60698 | 76 | then show ?thesis | 
| 77 | by blast | |
| 78 | qed | |
| 79 | fix e1 m1 e2 m2 :: int | |
| 80 | consider "e2 \<le> e1" | "e1 \<le> e2" by (rule linorder_le_cases) | |
| 81 | then show "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e" | |
| 82 | proof cases | |
| 83 | case 1 | |
| 84 | from *[OF this, of m2 m1] show ?thesis | |
| 85 | by (simp add: ac_simps) | |
| 86 | next | |
| 87 | case 2 | |
| 88 | then show ?thesis by (rule *) | |
| 89 | qed | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 90 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 91 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 92 | lemma uminus_float[simp]: "x \<in> float \<Longrightarrow> -x \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 93 | apply (auto simp: float_def) | 
| 57492 
74bf65a1910a
Hypsubst preserves equality hypotheses
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
56777diff
changeset | 94 | apply hypsubst_thin | 
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 95 | apply (rename_tac m e) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 96 | apply (rule_tac x="-m" in exI) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 97 | apply (rule_tac x="e" in exI) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 98 | apply (simp add: field_simps) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 99 | done | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 100 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 101 | lemma times_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> x * y \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 102 | apply (auto simp: float_def) | 
| 57492 
74bf65a1910a
Hypsubst preserves equality hypotheses
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
56777diff
changeset | 103 | apply hypsubst_thin | 
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 104 | apply (rename_tac mx my ex ey) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 105 | apply (rule_tac x="mx * my" in exI) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 106 | apply (rule_tac x="ex + ey" in exI) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 107 | apply (simp add: powr_add) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 108 | done | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 109 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 110 | lemma minus_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> x - y \<in> float" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53381diff
changeset | 111 | using plus_float [of x "- y"] by simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 112 | |
| 61945 | 113 | lemma abs_float[simp]: "x \<in> float \<Longrightarrow> \<bar>x\<bar> \<in> float" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 114 | by (cases x rule: linorder_cases[of 0]) auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 115 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 116 | lemma sgn_of_float[simp]: "x \<in> float \<Longrightarrow> sgn x \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 117 | by (cases x rule: linorder_cases[of 0]) (auto intro!: uminus_float) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21256diff
changeset | 118 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 119 | lemma div_power_2_float[simp]: "x \<in> float \<Longrightarrow> x / 2^d \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 120 | apply (auto simp add: float_def) | 
| 57492 
74bf65a1910a
Hypsubst preserves equality hypotheses
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
56777diff
changeset | 121 | apply hypsubst_thin | 
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 122 | apply (rename_tac m e) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 123 | apply (rule_tac x="m" in exI) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 124 | apply (rule_tac x="e - d" in exI) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 125 | apply (simp add: powr_realpow[symmetric] field_simps powr_add[symmetric]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 126 | done | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 127 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 128 | lemma div_power_2_int_float[simp]: "x \<in> float \<Longrightarrow> x / (2::int)^d \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 129 | apply (auto simp add: float_def) | 
| 57492 
74bf65a1910a
Hypsubst preserves equality hypotheses
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
56777diff
changeset | 130 | apply hypsubst_thin | 
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 131 | apply (rename_tac m e) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 132 | apply (rule_tac x="m" in exI) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 133 | apply (rule_tac x="e - d" in exI) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 134 | apply (simp add: powr_realpow[symmetric] field_simps powr_add[symmetric]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 135 | done | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 136 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 137 | lemma div_numeral_Bit0_float[simp]: | 
| 60698 | 138 | assumes x: "x / numeral n \<in> float" | 
| 139 | shows "x / (numeral (Num.Bit0 n)) \<in> float" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 140 | proof - | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 141 | have "(x / numeral n) / 2^1 \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 142 | by (intro x div_power_2_float) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 143 | also have "(x / numeral n) / 2^1 = x / (numeral (Num.Bit0 n))" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 144 | by (induct n) auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 145 | finally show ?thesis . | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 146 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 147 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 148 | lemma div_neg_numeral_Bit0_float[simp]: | 
| 60698 | 149 | assumes x: "x / numeral n \<in> float" | 
| 150 | shows "x / (- numeral (Num.Bit0 n)) \<in> float" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 151 | proof - | 
| 60698 | 152 | have "- (x / numeral (Num.Bit0 n)) \<in> float" | 
| 153 | using x by simp | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 154 | also have "- (x / numeral (Num.Bit0 n)) = x / - numeral (Num.Bit0 n)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 155 | by simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 156 | finally show ?thesis . | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 157 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 158 | |
| 60698 | 159 | lemma power_float[simp]: | 
| 160 | assumes "a \<in> float" | |
| 161 | shows "a ^ b \<in> float" | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 162 | proof - | 
| 60698 | 163 | from assms obtain m e :: int where "a = m * 2 powr e" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 164 | by (auto simp: float_def) | 
| 60698 | 165 | then show ?thesis | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 166 | by (auto intro!: floatI[where m="m^b" and e = "e*b"] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 167 | simp: power_mult_distrib powr_realpow[symmetric] powr_powr) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 168 | qed | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 169 | |
| 60698 | 170 | lift_definition Float :: "int \<Rightarrow> int \<Rightarrow> float" is "\<lambda>(m::int) (e::int). m * 2 powr e" | 
| 171 | by simp | |
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 172 | declare Float.rep_eq[simp] | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 173 | |
| 62419 
c7d6f4dded19
compute_real_of_float has not been used as code equation
 immler parents: 
62390diff
changeset | 174 | code_datatype Float | 
| 
c7d6f4dded19
compute_real_of_float has not been used as code equation
 immler parents: 
62390diff
changeset | 175 | |
| 47780 | 176 | lemma compute_real_of_float[code]: | 
| 177 | "real_of_float (Float m e) = (if e \<ge> 0 then m * 2 ^ nat e else m / 2 ^ (nat (-e)))" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 178 | by (simp add: powr_int) | 
| 47780 | 179 | |
| 60698 | 180 | |
| 60500 | 181 | subsection \<open>Arithmetic operations on floating point numbers\<close> | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 182 | |
| 47600 | 183 | instantiation float :: "{ring_1, linorder, linordered_ring, linordered_idom, numeral, equal}"
 | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 184 | begin | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 185 | |
| 47600 | 186 | lift_definition zero_float :: float is 0 by simp | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 187 | declare zero_float.rep_eq[simp] | 
| 47600 | 188 | lift_definition one_float :: float is 1 by simp | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 189 | declare one_float.rep_eq[simp] | 
| 47600 | 190 | lift_definition plus_float :: "float \<Rightarrow> float \<Rightarrow> float" is "op +" by simp | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 191 | declare plus_float.rep_eq[simp] | 
| 47600 | 192 | lift_definition times_float :: "float \<Rightarrow> float \<Rightarrow> float" is "op *" by simp | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 193 | declare times_float.rep_eq[simp] | 
| 47600 | 194 | lift_definition minus_float :: "float \<Rightarrow> float \<Rightarrow> float" is "op -" by simp | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 195 | declare minus_float.rep_eq[simp] | 
| 47600 | 196 | lift_definition uminus_float :: "float \<Rightarrow> float" is "uminus" by simp | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 197 | declare uminus_float.rep_eq[simp] | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 198 | |
| 47600 | 199 | lift_definition abs_float :: "float \<Rightarrow> float" is abs by simp | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 200 | declare abs_float.rep_eq[simp] | 
| 47600 | 201 | lift_definition sgn_float :: "float \<Rightarrow> float" is sgn by simp | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 202 | declare sgn_float.rep_eq[simp] | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 203 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 204 | lift_definition equal_float :: "float \<Rightarrow> float \<Rightarrow> bool" is "op = :: real \<Rightarrow> real \<Rightarrow> bool" . | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 205 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 206 | lift_definition less_eq_float :: "float \<Rightarrow> float \<Rightarrow> bool" is "op \<le>" . | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 207 | declare less_eq_float.rep_eq[simp] | 
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 208 | lift_definition less_float :: "float \<Rightarrow> float \<Rightarrow> bool" is "op <" . | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 209 | declare less_float.rep_eq[simp] | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 210 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 211 | instance | 
| 60698 | 212 | by (standard; transfer; fastforce simp add: field_simps intro: mult_left_mono mult_right_mono)+ | 
| 213 | ||
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 214 | end | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 215 | |
| 61639 
6ef461bee3fa
new conversion theorems for int, nat to float
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 216 | lemma real_of_float [simp]: "real_of_float (of_nat n) = of_nat n" | 
| 
6ef461bee3fa
new conversion theorems for int, nat to float
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 217 | by (induct n) simp_all | 
| 
6ef461bee3fa
new conversion theorems for int, nat to float
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 218 | |
| 
6ef461bee3fa
new conversion theorems for int, nat to float
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 219 | lemma real_of_float_of_int_eq [simp]: "real_of_float (of_int z) = of_int z" | 
| 
6ef461bee3fa
new conversion theorems for int, nat to float
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 220 | by (cases z rule: int_diff_cases) (simp_all add: of_rat_diff) | 
| 
6ef461bee3fa
new conversion theorems for int, nat to float
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 221 | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 222 | lemma Float_0_eq_0[simp]: "Float 0 e = 0" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 223 | by transfer simp | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 224 | |
| 60698 | 225 | lemma real_of_float_power[simp]: | 
| 226 | fixes f :: float | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 227 | shows "real_of_float (f^n) = real_of_float f^n" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 228 | by (induct n) simp_all | 
| 45495 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 hoelzl parents: 
44766diff
changeset | 229 | |
| 60698 | 230 | lemma | 
| 231 | fixes x y :: float | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 232 | shows real_of_float_min: "real_of_float (min x y) = min (real_of_float x) (real_of_float y)" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 233 | and real_of_float_max: "real_of_float (max x y) = max (real_of_float x) (real_of_float y)" | 
| 47600 | 234 | by (simp_all add: min_def max_def) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 235 | |
| 53215 
5e47c31c6f7c
renamed typeclass dense_linorder to unbounded_dense_linorder
 hoelzl parents: 
51542diff
changeset | 236 | instance float :: unbounded_dense_linorder | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 237 | proof | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 238 | fix a b :: float | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 239 | show "\<exists>c. a < c" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 240 | apply (intro exI[of _ "a + 1"]) | 
| 47600 | 241 | apply transfer | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 242 | apply simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 243 | done | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 244 | show "\<exists>c. c < a" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 245 | apply (intro exI[of _ "a - 1"]) | 
| 47600 | 246 | apply transfer | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 247 | apply simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 248 | done | 
| 60698 | 249 | show "\<exists>c. a < c \<and> c < b" if "a < b" | 
| 250 | apply (rule exI[of _ "(a + b) * Float 1 (- 1)"]) | |
| 251 | using that | |
| 47600 | 252 | apply transfer | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 253 | apply (simp add: powr_minus) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 254 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 255 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 256 | |
| 47600 | 257 | instantiation float :: lattice_ab_group_add | 
| 46573 | 258 | begin | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 259 | |
| 60698 | 260 | definition inf_float :: "float \<Rightarrow> float \<Rightarrow> float" | 
| 261 | where "inf_float a b = min a b" | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 262 | |
| 60698 | 263 | definition sup_float :: "float \<Rightarrow> float \<Rightarrow> float" | 
| 264 | where "sup_float a b = max a b" | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 265 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 266 | instance | 
| 60679 | 267 | by (standard; transfer; simp add: inf_float_def sup_float_def real_of_float_min real_of_float_max) | 
| 268 | ||
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 269 | end | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 270 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 271 | lemma float_numeral[simp]: "real_of_float (numeral x :: float) = numeral x" | 
| 47600 | 272 | apply (induct x) | 
| 273 | apply simp | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 274 | apply (simp_all only: numeral_Bit0 numeral_Bit1 real_of_float_eq float_of_inverse | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 275 | plus_float.rep_eq one_float.rep_eq plus_float numeral_float one_float) | 
| 47600 | 276 | done | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 277 | |
| 53381 | 278 | lemma transfer_numeral [transfer_rule]: | 
| 55945 | 279 | "rel_fun (op =) pcr_float (numeral :: _ \<Rightarrow> real) (numeral :: _ \<Rightarrow> float)" | 
| 60698 | 280 | by (simp add: rel_fun_def float.pcr_cr_eq cr_float_def) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 281 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 282 | lemma float_neg_numeral[simp]: "real_of_float (- numeral x :: float) = - numeral x" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 283 | by simp | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46670diff
changeset | 284 | |
| 53381 | 285 | lemma transfer_neg_numeral [transfer_rule]: | 
| 55945 | 286 | "rel_fun (op =) pcr_float (- numeral :: _ \<Rightarrow> real) (- numeral :: _ \<Rightarrow> float)" | 
| 60698 | 287 | by (simp add: rel_fun_def float.pcr_cr_eq cr_float_def) | 
| 47600 | 288 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 289 | lemma | 
| 47600 | 290 | shows float_of_numeral[simp]: "numeral k = float_of (numeral k)" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 291 | and float_of_neg_numeral[simp]: "- numeral k = float_of (- numeral k)" | 
| 47600 | 292 | unfolding real_of_float_eq by simp_all | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46670diff
changeset | 293 | |
| 60698 | 294 | |
| 60500 | 295 | subsection \<open>Quickcheck\<close> | 
| 58987 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 296 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 297 | instantiation float :: exhaustive | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 298 | begin | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 299 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 300 | definition exhaustive_float where | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 301 | "exhaustive_float f d = | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 302 | Quickcheck_Exhaustive.exhaustive (%x. Quickcheck_Exhaustive.exhaustive (%y. f (Float x y)) d) d" | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 303 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 304 | instance .. | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 305 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 306 | end | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 307 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 308 | definition (in term_syntax) [code_unfold]: | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 309 |   "valtermify_float x y = Code_Evaluation.valtermify Float {\<cdot>} x {\<cdot>} y"
 | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 310 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 311 | instantiation float :: full_exhaustive | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 312 | begin | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 313 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 314 | definition full_exhaustive_float where | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 315 | "full_exhaustive_float f d = | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 316 | Quickcheck_Exhaustive.full_exhaustive | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 317 | (\<lambda>x. Quickcheck_Exhaustive.full_exhaustive (\<lambda>y. f (valtermify_float x y)) d) d" | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 318 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 319 | instance .. | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 320 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 321 | end | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 322 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 323 | instantiation float :: random | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 324 | begin | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 325 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 326 | definition "Quickcheck_Random.random i = | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 327 | scomp (Quickcheck_Random.random (2 ^ nat_of_natural i)) | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 328 | (\<lambda>man. scomp (Quickcheck_Random.random i) (\<lambda>exp. Pair (valtermify_float man exp)))" | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 329 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 330 | instance .. | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 331 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 332 | end | 
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 333 | |
| 
119680ebf37c
quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
 immler parents: 
58985diff
changeset | 334 | |
| 60500 | 335 | subsection \<open>Represent floats as unique mantissa and exponent\<close> | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46670diff
changeset | 336 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 337 | lemma int_induct_abs[case_names less]: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 338 | fixes j :: int | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 339 | assumes H: "\<And>n. (\<And>i. \<bar>i\<bar> < \<bar>n\<bar> \<Longrightarrow> P i) \<Longrightarrow> P n" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 340 | shows "P j" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 341 | proof (induct "nat \<bar>j\<bar>" arbitrary: j rule: less_induct) | 
| 60698 | 342 | case less | 
| 343 | show ?case by (rule H[OF less]) simp | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 344 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 345 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 346 | lemma int_cancel_factors: | 
| 60698 | 347 | fixes n :: int | 
| 348 | assumes "1 < r" | |
| 349 | shows "n = 0 \<or> (\<exists>k i. n = k * r ^ i \<and> \<not> r dvd k)" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 350 | proof (induct n rule: int_induct_abs) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 351 | case (less n) | 
| 60698 | 352 | have "\<exists>k i. n = k * r ^ Suc i \<and> \<not> r dvd k" if "n \<noteq> 0" "n = m * r" for m | 
| 353 | proof - | |
| 354 | from that have "\<bar>m \<bar> < \<bar>n\<bar>" | |
| 60500 | 355 | using \<open>1 < r\<close> by (simp add: abs_mult) | 
| 60698 | 356 | from less[OF this] that show ?thesis by auto | 
| 357 | qed | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 358 | then show ?case | 
| 59554 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
59487diff
changeset | 359 | by (metis dvd_def monoid_mult_class.mult.right_neutral mult.commute power_0) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 360 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 361 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 362 | lemma mult_powr_eq_mult_powr_iff_asym: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 363 | fixes m1 m2 e1 e2 :: int | 
| 60698 | 364 | assumes m1: "\<not> 2 dvd m1" | 
| 365 | and "e1 \<le> e2" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 366 | shows "m1 * 2 powr e1 = m2 * 2 powr e2 \<longleftrightarrow> m1 = m2 \<and> e1 = e2" | 
| 60698 | 367 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 368 | proof | 
| 60698 | 369 | show ?rhs if eq: ?lhs | 
| 370 | proof - | |
| 371 | have "m1 \<noteq> 0" | |
| 372 | using m1 unfolding dvd_def by auto | |
| 373 | from \<open>e1 \<le> e2\<close> eq have "m1 = m2 * 2 powr nat (e2 - e1)" | |
| 374 | by (simp add: powr_divide2[symmetric] field_simps) | |
| 375 | also have "\<dots> = m2 * 2^nat (e2 - e1)" | |
| 376 | by (simp add: powr_realpow) | |
| 377 | finally have m1_eq: "m1 = m2 * 2^nat (e2 - e1)" | |
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61639diff
changeset | 378 | by linarith | 
| 60698 | 379 | with m1 have "m1 = m2" | 
| 380 | by (cases "nat (e2 - e1)") (auto simp add: dvd_def) | |
| 381 | then show ?thesis | |
| 382 | using eq \<open>m1 \<noteq> 0\<close> by (simp add: powr_inj) | |
| 383 | qed | |
| 384 | show ?lhs if ?rhs | |
| 385 | using that by simp | |
| 386 | qed | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 387 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 388 | lemma mult_powr_eq_mult_powr_iff: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 389 | fixes m1 m2 e1 e2 :: int | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 390 | shows "\<not> 2 dvd m1 \<Longrightarrow> \<not> 2 dvd m2 \<Longrightarrow> m1 * 2 powr e1 = m2 * 2 powr e2 \<longleftrightarrow> m1 = m2 \<and> e1 = e2" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 391 | using mult_powr_eq_mult_powr_iff_asym[of m1 e1 e2 m2] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 392 | using mult_powr_eq_mult_powr_iff_asym[of m2 e2 e1 m1] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 393 | by (cases e1 e2 rule: linorder_le_cases) auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 394 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 395 | lemma floatE_normed: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 396 | assumes x: "x \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 397 | obtains (zero) "x = 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 398 | | (powr) m e :: int where "x = m * 2 powr e" "\<not> 2 dvd m" "x \<noteq> 0" | 
| 60698 | 399 | proof - | 
| 400 |   {
 | |
| 401 | assume "x \<noteq> 0" | |
| 402 | from x obtain m e :: int where x: "x = m * 2 powr e" | |
| 403 | by (auto simp: float_def) | |
| 60500 | 404 | with \<open>x \<noteq> 0\<close> int_cancel_factors[of 2 m] obtain k i where "m = k * 2 ^ i" "\<not> 2 dvd k" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 405 | by auto | 
| 60500 | 406 | with \<open>\<not> 2 dvd k\<close> x have "\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 407 | by (rule_tac exI[of _ "k"], rule_tac exI[of _ "e + int i"]) | 
| 60698 | 408 | (simp add: powr_add powr_realpow) | 
| 409 | } | |
| 410 | with that show thesis by blast | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 411 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 412 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 413 | lemma float_normed_cases: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 414 | fixes f :: float | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 415 | obtains (zero) "f = 0" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 416 | | (powr) m e :: int where "real_of_float f = m * 2 powr e" "\<not> 2 dvd m" "f \<noteq> 0" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 417 | proof (atomize_elim, induct f) | 
| 60698 | 418 | case (float_of y) | 
| 419 | then show ?case | |
| 47600 | 420 | by (cases rule: floatE_normed) (auto simp: zero_float_def) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 421 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 422 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 423 | definition mantissa :: "float \<Rightarrow> int" where | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 424 | "mantissa f = fst (SOME p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 425 | \<or> (f \<noteq> 0 \<and> real_of_float f = real_of_int (fst p) * 2 powr real_of_int (snd p) \<and> \<not> 2 dvd fst p))" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 426 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 427 | definition exponent :: "float \<Rightarrow> int" where | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 428 | "exponent f = snd (SOME p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 429 | \<or> (f \<noteq> 0 \<and> real_of_float f = real_of_int (fst p) * 2 powr real_of_int (snd p) \<and> \<not> 2 dvd fst p))" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 430 | |
| 53381 | 431 | lemma | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 432 | shows exponent_0[simp]: "exponent (float_of 0) = 0" (is ?E) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 433 | and mantissa_0[simp]: "mantissa (float_of 0) = 0" (is ?M) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 434 | proof - | 
| 60698 | 435 | have "\<And>p::int \<times> int. fst p = 0 \<and> snd p = 0 \<longleftrightarrow> p = (0, 0)" | 
| 436 | by auto | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 437 | then show ?E ?M | 
| 47600 | 438 | by (auto simp add: mantissa_def exponent_def zero_float_def) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 439 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 440 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 441 | lemma | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 442 | shows mantissa_exponent: "real_of_float f = mantissa f * 2 powr exponent f" (is ?E) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 443 | and mantissa_not_dvd: "f \<noteq> (float_of 0) \<Longrightarrow> \<not> 2 dvd mantissa f" (is "_ \<Longrightarrow> ?D") | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 444 | proof cases | 
| 60698 | 445 | assume [simp]: "f \<noteq> float_of 0" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 446 | have "f = mantissa f * 2 powr exponent f \<and> \<not> 2 dvd mantissa f" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 447 | proof (cases f rule: float_normed_cases) | 
| 60698 | 448 | case zero | 
| 449 | then show ?thesis by (simp add: zero_float_def) | |
| 450 | next | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 451 | case (powr m e) | 
| 60698 | 452 | then have "\<exists>p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0) \<or> | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 453 | (f \<noteq> 0 \<and> real_of_float f = real_of_int (fst p) * 2 powr real_of_int (snd p) \<and> \<not> 2 dvd fst p)" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 454 | by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 455 | then show ?thesis | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 456 | unfolding exponent_def mantissa_def | 
| 47600 | 457 | by (rule someI2_ex) (simp add: zero_float_def) | 
| 60698 | 458 | qed | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 459 | then show ?E ?D by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 460 | qed simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 461 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 462 | lemma mantissa_noteq_0: "f \<noteq> float_of 0 \<Longrightarrow> mantissa f \<noteq> 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 463 | using mantissa_not_dvd[of f] by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 464 | |
| 53381 | 465 | lemma | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 466 | fixes m e :: int | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 467 | defines "f \<equiv> float_of (m * 2 powr e)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 468 | assumes dvd: "\<not> 2 dvd m" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 469 | shows mantissa_float: "mantissa f = m" (is "?M") | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 470 | and exponent_float: "m \<noteq> 0 \<Longrightarrow> exponent f = e" (is "_ \<Longrightarrow> ?E") | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 471 | proof cases | 
| 60698 | 472 | assume "m = 0" | 
| 473 | with dvd show "mantissa f = m" by auto | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 474 | next | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 475 | assume "m \<noteq> 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 476 | then have f_not_0: "f \<noteq> float_of 0" by (simp add: f_def) | 
| 60698 | 477 | from mantissa_exponent[of f] have "m * 2 powr e = mantissa f * 2 powr exponent f" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 478 | by (auto simp add: f_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 479 | then show "?M" "?E" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 480 | using mantissa_not_dvd[OF f_not_0] dvd | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 481 | by (auto simp: mult_powr_eq_mult_powr_iff) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 482 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 483 | |
| 60698 | 484 | |
| 60500 | 485 | subsection \<open>Compute arithmetic operations\<close> | 
| 47600 | 486 | |
| 487 | lemma Float_mantissa_exponent: "Float (mantissa f) (exponent f) = f" | |
| 488 | unfolding real_of_float_eq mantissa_exponent[of f] by simp | |
| 489 | ||
| 60698 | 490 | lemma Float_cases [cases type: float]: | 
| 47600 | 491 | fixes f :: float | 
| 492 | obtains (Float) m e :: int where "f = Float m e" | |
| 493 | using Float_mantissa_exponent[symmetric] | |
| 494 | by (atomize_elim) auto | |
| 495 | ||
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 496 | lemma denormalize_shift: | 
| 60698 | 497 | assumes f_def: "f \<equiv> Float m e" | 
| 498 | and not_0: "f \<noteq> float_of 0" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 499 | obtains i where "m = mantissa f * 2 ^ i" "e = exponent f - i" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 500 | proof | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 501 | from mantissa_exponent[of f] f_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 502 | have "m * 2 powr e = mantissa f * 2 powr exponent f" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 503 | by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 504 | then have eq: "m = mantissa f * 2 powr (exponent f - e)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 505 | by (simp add: powr_divide2[symmetric] field_simps) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 506 | moreover | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 507 | have "e \<le> exponent f" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 508 | proof (rule ccontr) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 509 | assume "\<not> e \<le> exponent f" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 510 | then have pos: "exponent f < e" by simp | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 511 | then have "2 powr (exponent f - e) = 2 powr - real_of_int (e - exponent f)" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 512 | by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 513 | also have "\<dots> = 1 / 2^nat (e - exponent f)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 514 | using pos by (simp add: powr_realpow[symmetric] powr_divide2[symmetric]) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 515 | finally have "m * 2^nat (e - exponent f) = real_of_int (mantissa f)" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 516 | using eq by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 517 | then have "mantissa f = m * 2^nat (e - exponent f)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 518 | by linarith | 
| 60500 | 519 | with \<open>exponent f < e\<close> have "2 dvd mantissa f" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 520 | apply (intro dvdI[where k="m * 2^(nat (e-exponent f)) div 2"]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 521 | apply (cases "nat (e - exponent f)") | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 522 | apply auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 523 | done | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 524 | then show False using mantissa_not_dvd[OF not_0] by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 525 | qed | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 526 | ultimately have "real_of_int m = mantissa f * 2^nat (exponent f - e)" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 527 | by (simp add: powr_realpow[symmetric]) | 
| 60500 | 528 | with \<open>e \<le> exponent f\<close> | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61639diff
changeset | 529 | show "m = mantissa f * 2 ^ nat (exponent f - e)" | 
| 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61639diff
changeset | 530 | by linarith | 
| 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61639diff
changeset | 531 | show "e = exponent f - nat (exponent f - e)" | 
| 61799 | 532 | using \<open>e \<le> exponent f\<close> by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 533 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 534 | |
| 60698 | 535 | context | 
| 536 | begin | |
| 47600 | 537 | |
| 60698 | 538 | qualified lemma compute_float_zero[code_unfold, code]: "0 = Float 0 0" | 
| 47600 | 539 | by transfer simp | 
| 60698 | 540 | |
| 541 | qualified lemma compute_float_one[code_unfold, code]: "1 = Float 1 0" | |
| 542 | by transfer simp | |
| 47600 | 543 | |
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 544 | lift_definition normfloat :: "float \<Rightarrow> float" is "\<lambda>x. x" . | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 545 | lemma normloat_id[simp]: "normfloat x = x" by transfer rule | 
| 47600 | 546 | |
| 60698 | 547 | qualified lemma compute_normfloat[code]: "normfloat (Float m e) = | 
| 47600 | 548 | (if m mod 2 = 0 \<and> m \<noteq> 0 then normfloat (Float (m div 2) (e + 1)) | 
| 549 | else if m = 0 then 0 else Float m e)" | |
| 550 | by transfer (auto simp add: powr_add zmod_eq_0_iff) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 551 | |
| 60698 | 552 | qualified lemma compute_float_numeral[code_abbrev]: "Float (numeral k) 0 = numeral k" | 
| 47600 | 553 | by transfer simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 554 | |
| 60698 | 555 | qualified lemma compute_float_neg_numeral[code_abbrev]: "Float (- numeral k) 0 = - numeral k" | 
| 47600 | 556 | by transfer simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 557 | |
| 60698 | 558 | qualified lemma compute_float_uminus[code]: "- Float m1 e1 = Float (- m1) e1" | 
| 47600 | 559 | by transfer simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 560 | |
| 60698 | 561 | qualified lemma compute_float_times[code]: "Float m1 e1 * Float m2 e2 = Float (m1 * m2) (e1 + e2)" | 
| 47600 | 562 | by transfer (simp add: field_simps powr_add) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 563 | |
| 60698 | 564 | qualified lemma compute_float_plus[code]: "Float m1 e1 + Float m2 e2 = | 
| 54783 
25860d89a044
Float: prevent unnecessary large numbers when adding 0
 immler parents: 
54782diff
changeset | 565 | (if m1 = 0 then Float m2 e2 else if m2 = 0 then Float m1 e1 else | 
| 
25860d89a044
Float: prevent unnecessary large numbers when adding 0
 immler parents: 
54782diff
changeset | 566 | if e1 \<le> e2 then Float (m1 + m2 * 2^nat (e2 - e1)) e1 | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 567 | else Float (m2 + m1 * 2^nat (e1 - e2)) e2)" | 
| 47600 | 568 | by transfer (simp add: field_simps powr_realpow[symmetric] powr_divide2[symmetric]) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 569 | |
| 60698 | 570 | qualified lemma compute_float_minus[code]: fixes f g::float shows "f - g = f + (-g)" | 
| 47600 | 571 | by simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 572 | |
| 60698 | 573 | qualified lemma compute_float_sgn[code]: "sgn (Float m1 e1) = (if 0 < m1 then 1 else if m1 < 0 then -1 else 0)" | 
| 47600 | 574 | by transfer (simp add: sgn_times) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 575 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 576 | lift_definition is_float_pos :: "float \<Rightarrow> bool" is "op < 0 :: real \<Rightarrow> bool" . | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 577 | |
| 60698 | 578 | qualified lemma compute_is_float_pos[code]: "is_float_pos (Float m e) \<longleftrightarrow> 0 < m" | 
| 47600 | 579 | by transfer (auto simp add: zero_less_mult_iff not_le[symmetric, of _ 0]) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 580 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 581 | lift_definition is_float_nonneg :: "float \<Rightarrow> bool" is "op \<le> 0 :: real \<Rightarrow> bool" . | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 582 | |
| 60698 | 583 | qualified lemma compute_is_float_nonneg[code]: "is_float_nonneg (Float m e) \<longleftrightarrow> 0 \<le> m" | 
| 47600 | 584 | by transfer (auto simp add: zero_le_mult_iff not_less[symmetric, of _ 0]) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 585 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 586 | lift_definition is_float_zero :: "float \<Rightarrow> bool" is "op = 0 :: real \<Rightarrow> bool" . | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 587 | |
| 60698 | 588 | qualified lemma compute_is_float_zero[code]: "is_float_zero (Float m e) \<longleftrightarrow> 0 = m" | 
| 47600 | 589 | by transfer (auto simp add: is_float_zero_def) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 590 | |
| 61945 | 591 | qualified lemma compute_float_abs[code]: "\<bar>Float m e\<bar> = Float \<bar>m\<bar> e" | 
| 47600 | 592 | by transfer (simp add: abs_mult) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 593 | |
| 60698 | 594 | qualified lemma compute_float_eq[code]: "equal_class.equal f g = is_float_zero (f - g)" | 
| 47600 | 595 | by transfer simp | 
| 60698 | 596 | |
| 597 | end | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 598 | |
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 599 | |
| 60500 | 600 | subsection \<open>Lemmas for types @{typ real}, @{typ nat}, @{typ int}\<close>
 | 
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 601 | |
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 602 | lemmas real_of_ints = | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 603 | of_int_add | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 604 | of_int_minus | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 605 | of_int_diff | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 606 | of_int_mult | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 607 | of_int_power | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 608 | of_int_numeral of_int_neg_numeral | 
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 609 | |
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 610 | lemmas int_of_reals = real_of_ints[symmetric] | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 611 | |
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 612 | |
| 60500 | 613 | subsection \<open>Rounding Real Numbers\<close> | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 614 | |
| 60698 | 615 | definition round_down :: "int \<Rightarrow> real \<Rightarrow> real" | 
| 61942 | 616 | where "round_down prec x = \<lfloor>x * 2 powr prec\<rfloor> * 2 powr -prec" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 617 | |
| 60698 | 618 | definition round_up :: "int \<Rightarrow> real \<Rightarrow> real" | 
| 61942 | 619 | where "round_up prec x = \<lceil>x * 2 powr prec\<rceil> * 2 powr -prec" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 620 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 621 | lemma round_down_float[simp]: "round_down prec x \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 622 | unfolding round_down_def | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 623 | by (auto intro!: times_float simp: of_int_minus[symmetric] simp del: of_int_minus) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 624 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 625 | lemma round_up_float[simp]: "round_up prec x \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 626 | unfolding round_up_def | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 627 | by (auto intro!: times_float simp: of_int_minus[symmetric] simp del: of_int_minus) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 628 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 629 | lemma round_up: "x \<le> round_up prec x" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 630 | by (simp add: powr_minus_divide le_divide_eq round_up_def ceiling_correct) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 631 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 632 | lemma round_down: "round_down prec x \<le> x" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 633 | by (simp add: powr_minus_divide divide_le_eq round_down_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 634 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 635 | lemma round_up_0[simp]: "round_up p 0 = 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 636 | unfolding round_up_def by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 637 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 638 | lemma round_down_0[simp]: "round_down p 0 = 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 639 | unfolding round_down_def by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 640 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 641 | lemma round_up_diff_round_down: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 642 | "round_up prec x - round_down prec x \<le> 2 powr -prec" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 643 | proof - | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 644 | have "round_up prec x - round_down prec x = | 
| 61942 | 645 | (\<lceil>x * 2 powr prec\<rceil> - \<lfloor>x * 2 powr prec\<rfloor>) * 2 powr -prec" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 646 | by (simp add: round_up_def round_down_def field_simps) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 647 | also have "\<dots> \<le> 1 * 2 powr -prec" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 648 | by (rule mult_mono) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 649 | (auto simp del: of_int_diff | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 650 | simp: of_int_diff[symmetric] ceiling_diff_floor_le_1) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 651 | finally show ?thesis by simp | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 652 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 653 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 654 | lemma round_down_shift: "round_down p (x * 2 powr k) = 2 powr k * round_down (p + k) x" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 655 | unfolding round_down_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 656 | by (simp add: powr_add powr_mult field_simps powr_divide2[symmetric]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 657 | (simp add: powr_add[symmetric]) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 658 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 659 | lemma round_up_shift: "round_up p (x * 2 powr k) = 2 powr k * round_up (p + k) x" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 660 | unfolding round_up_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 661 | by (simp add: powr_add powr_mult field_simps powr_divide2[symmetric]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 662 | (simp add: powr_add[symmetric]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 663 | |
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 664 | lemma round_up_uminus_eq: "round_up p (-x) = - round_down p x" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 665 | and round_down_uminus_eq: "round_down p (-x) = - round_up p x" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 666 | by (auto simp: round_up_def round_down_def ceiling_def) | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 667 | |
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 668 | lemma round_up_mono: "x \<le> y \<Longrightarrow> round_up p x \<le> round_up p y" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 669 | by (auto intro!: ceiling_mono simp: round_up_def) | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 670 | |
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 671 | lemma round_up_le1: | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 672 | assumes "x \<le> 1" "prec \<ge> 0" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 673 | shows "round_up prec x \<le> 1" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 674 | proof - | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 675 | have "real_of_int \<lceil>x * 2 powr prec\<rceil> \<le> real_of_int \<lceil>2 powr real_of_int prec\<rceil>" | 
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 676 | using assms by (auto intro!: ceiling_mono) | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 677 | also have "\<dots> = 2 powr prec" using assms by (auto simp: powr_int intro!: exI[where x="2^nat prec"]) | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 678 | finally show ?thesis | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 679 | by (simp add: round_up_def) (simp add: powr_minus inverse_eq_divide) | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 680 | qed | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 681 | |
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 682 | lemma round_up_less1: | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 683 | assumes "x < 1 / 2" "p > 0" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 684 | shows "round_up p x < 1" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 685 | proof - | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 686 | have "x * 2 powr p < 1 / 2 * 2 powr p" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 687 | using assms by simp | 
| 60500 | 688 | also have "\<dots> \<le> 2 powr p - 1" using \<open>p > 0\<close> | 
| 58989 | 689 | by (auto simp: powr_divide2[symmetric] powr_int field_simps self_le_power) | 
| 60500 | 690 | finally show ?thesis using \<open>p > 0\<close> | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 691 | by (simp add: round_up_def field_simps powr_minus powr_int ceiling_less_iff) | 
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 692 | qed | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 693 | |
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 694 | lemma round_down_ge1: | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 695 | assumes x: "x \<ge> 1" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 696 | assumes prec: "p \<ge> - log 2 x" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 697 | shows "1 \<le> round_down p x" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 698 | proof cases | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 699 | assume nonneg: "0 \<le> p" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 700 | have "2 powr p = real_of_int \<lfloor>2 powr real_of_int p\<rfloor>" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 701 | using nonneg by (auto simp: powr_int) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 702 | also have "\<dots> \<le> real_of_int \<lfloor>x * 2 powr p\<rfloor>" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 703 | using assms by (auto intro!: floor_mono) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 704 | finally show ?thesis | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 705 | by (simp add: round_down_def) (simp add: powr_minus inverse_eq_divide) | 
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 706 | next | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 707 | assume neg: "\<not> 0 \<le> p" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 708 | have "x = 2 powr (log 2 x)" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 709 | using x by simp | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 710 | also have "2 powr (log 2 x) \<ge> 2 powr - p" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 711 | using prec by auto | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 712 | finally have x_le: "x \<ge> 2 powr -p" . | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 713 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 714 | from neg have "2 powr real_of_int p \<le> 2 powr 0" | 
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 715 | by (intro powr_mono) auto | 
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 716 | also have "\<dots> \<le> \<lfloor>2 powr 0::real\<rfloor>" by simp | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 717 | also have "\<dots> \<le> \<lfloor>x * 2 powr (real_of_int p)\<rfloor>" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 718 | unfolding of_int_le_iff | 
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 719 | using x x_le by (intro floor_mono) (simp add: powr_minus_divide field_simps) | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 720 | finally show ?thesis | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 721 | using prec x | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 722 | by (simp add: round_down_def powr_minus_divide pos_le_divide_eq) | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 723 | qed | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 724 | |
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 725 | lemma round_up_le0: "x \<le> 0 \<Longrightarrow> round_up p x \<le> 0" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 726 | unfolding round_up_def | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 727 | by (auto simp: field_simps mult_le_0_iff zero_le_mult_iff) | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 728 | |
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 729 | |
| 60500 | 730 | subsection \<open>Rounding Floats\<close> | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 731 | |
| 60698 | 732 | definition div_twopow :: "int \<Rightarrow> nat \<Rightarrow> int" | 
| 733 | where [simp]: "div_twopow x n = x div (2 ^ n)" | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 734 | |
| 60698 | 735 | definition mod_twopow :: "int \<Rightarrow> nat \<Rightarrow> int" | 
| 736 | where [simp]: "mod_twopow x n = x mod (2 ^ n)" | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 737 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 738 | lemma compute_div_twopow[code]: | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 739 | "div_twopow x n = (if x = 0 \<or> x = -1 \<or> n = 0 then x else div_twopow (x div 2) (n - 1))" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 740 | by (cases n) (auto simp: zdiv_zmult2_eq div_eq_minus1) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 741 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 742 | lemma compute_mod_twopow[code]: | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 743 | "mod_twopow x n = (if n = 0 then 0 else x mod 2 + 2 * mod_twopow (x div 2) (n - 1))" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 744 | by (cases n) (auto simp: zmod_zmult2_eq) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 745 | |
| 47600 | 746 | lift_definition float_up :: "int \<Rightarrow> float \<Rightarrow> float" is round_up by simp | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 747 | declare float_up.rep_eq[simp] | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 748 | |
| 60698 | 749 | lemma round_up_correct: "round_up e f - f \<in> {0..2 powr -e}"
 | 
| 750 | unfolding atLeastAtMost_iff | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 751 | proof | 
| 60698 | 752 | have "round_up e f - f \<le> round_up e f - round_down e f" | 
| 753 | using round_down by simp | |
| 754 | also have "\<dots> \<le> 2 powr -e" | |
| 755 | using round_up_diff_round_down by simp | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 756 | finally show "round_up e f - f \<le> 2 powr - (real_of_int e)" | 
| 47600 | 757 | by simp | 
| 758 | qed (simp add: algebra_simps round_up) | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 759 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 760 | lemma float_up_correct: "real_of_float (float_up e f) - real_of_float f \<in> {0..2 powr -e}"
 | 
| 54782 | 761 | by transfer (rule round_up_correct) | 
| 762 | ||
| 47600 | 763 | lift_definition float_down :: "int \<Rightarrow> float \<Rightarrow> float" is round_down by simp | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 764 | declare float_down.rep_eq[simp] | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 765 | |
| 60698 | 766 | lemma round_down_correct: "f - (round_down e f) \<in> {0..2 powr -e}"
 | 
| 767 | unfolding atLeastAtMost_iff | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 768 | proof | 
| 60698 | 769 | have "f - round_down e f \<le> round_up e f - round_down e f" | 
| 770 | using round_up by simp | |
| 771 | also have "\<dots> \<le> 2 powr -e" | |
| 772 | using round_up_diff_round_down by simp | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 773 | finally show "f - round_down e f \<le> 2 powr - (real_of_int e)" | 
| 47600 | 774 | by simp | 
| 775 | qed (simp add: algebra_simps round_down) | |
| 24301 | 776 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 777 | lemma float_down_correct: "real_of_float f - real_of_float (float_down e f) \<in> {0..2 powr -e}"
 | 
| 54782 | 778 | by transfer (rule round_down_correct) | 
| 779 | ||
| 60698 | 780 | context | 
| 781 | begin | |
| 782 | ||
| 783 | qualified lemma compute_float_down[code]: | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 784 | "float_down p (Float m e) = | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 785 | (if p + e < 0 then Float (div_twopow m (nat (-(p + e)))) (-p) else Float m e)" | 
| 60698 | 786 | proof (cases "p + e < 0") | 
| 787 | case True | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 788 | then have "real_of_int ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 789 | using powr_realpow[of 2 "nat (-(p + e))"] by simp | 
| 60698 | 790 | also have "\<dots> = 1 / 2 powr p / 2 powr e" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 791 | unfolding powr_minus_divide of_int_minus by (simp add: powr_add) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 792 | finally show ?thesis | 
| 60500 | 793 | using \<open>p + e < 0\<close> | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 794 | apply transfer | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 795 | apply (simp add: ac_simps round_down_def floor_divide_of_int_eq[symmetric]) | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 796 | proof - (*FIXME*) | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 797 | fix pa :: int and ea :: int and ma :: int | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 798 | assume a1: "2 ^ nat (- pa - ea) = 1 / (2 powr real_of_int pa * 2 powr real_of_int ea)" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 799 | assume "pa + ea < 0" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 800 | have "\<lfloor>real_of_int ma / real_of_int (int 2 ^ nat (- (pa + ea)))\<rfloor> = \<lfloor>real_of_float (Float ma (pa + ea))\<rfloor>" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 801 | using a1 by (simp add: powr_add) | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 802 | thus "\<lfloor>real_of_int ma * (2 powr real_of_int pa * 2 powr real_of_int ea)\<rfloor> = ma div 2 ^ nat (- pa - ea)" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 803 | by (metis Float.rep_eq add_uminus_conv_diff floor_divide_of_int_eq minus_add_distrib of_int_simps(3) of_nat_numeral powr_add) | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 804 | qed | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 805 | next | 
| 60698 | 806 | case False | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 807 | then have r: "real_of_int e + real_of_int p = real (nat (e + p))" by simp | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 808 | have r: "\<lfloor>(m * 2 powr e) * 2 powr real_of_int p\<rfloor> = (m * 2 powr e) * 2 powr real_of_int p" | 
| 47600 | 809 | by (auto intro: exI[where x="m*2^nat (e+p)"] | 
| 810 | simp add: ac_simps powr_add[symmetric] r powr_realpow) | |
| 60500 | 811 | with \<open>\<not> p + e < 0\<close> show ?thesis | 
| 57862 | 812 | by transfer (auto simp add: round_down_def field_simps powr_add powr_minus) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 813 | qed | 
| 24301 | 814 | |
| 54782 | 815 | lemma abs_round_down_le: "\<bar>f - (round_down e f)\<bar> \<le> 2 powr -e" | 
| 816 | using round_down_correct[of f e] by simp | |
| 817 | ||
| 818 | lemma abs_round_up_le: "\<bar>f - (round_up e f)\<bar> \<le> 2 powr -e" | |
| 819 | using round_up_correct[of e f] by simp | |
| 820 | ||
| 821 | lemma round_down_nonneg: "0 \<le> s \<Longrightarrow> 0 \<le> round_down p s" | |
| 56536 | 822 | by (auto simp: round_down_def) | 
| 54782 | 823 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 824 | lemma ceil_divide_floor_conv: | 
| 60698 | 825 | assumes "b \<noteq> 0" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 826 | shows "\<lceil>real_of_int a / real_of_int b\<rceil> = (if b dvd a then a div b else \<lfloor>real_of_int a / real_of_int b\<rfloor> + 1)" | 
| 60698 | 827 | proof (cases "b dvd a") | 
| 828 | case True | |
| 829 | then show ?thesis | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 830 | by (simp add: ceiling_def of_int_minus[symmetric] divide_minus_left[symmetric] | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 831 | floor_divide_of_int_eq dvd_neg_div del: divide_minus_left of_int_minus) | 
| 60698 | 832 | next | 
| 833 | case False | |
| 834 | then have "a mod b \<noteq> 0" | |
| 835 | by auto | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 836 | then have ne: "real_of_int (a mod b) / real_of_int b \<noteq> 0" | 
| 60698 | 837 | using \<open>b \<noteq> 0\<close> by auto | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 838 | have "\<lceil>real_of_int a / real_of_int b\<rceil> = \<lfloor>real_of_int a / real_of_int b\<rfloor> + 1" | 
| 60698 | 839 | apply (rule ceiling_eq) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 840 | apply (auto simp: floor_divide_of_int_eq[symmetric]) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 841 | proof - | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 842 | have "real_of_int \<lfloor>real_of_int a / real_of_int b\<rfloor> \<le> real_of_int a / real_of_int b" | 
| 60698 | 843 | by simp | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 844 | moreover have "real_of_int \<lfloor>real_of_int a / real_of_int b\<rfloor> \<noteq> real_of_int a / real_of_int b" | 
| 60698 | 845 | apply (subst (2) real_of_int_div_aux) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 846 | unfolding floor_divide_of_int_eq | 
| 60698 | 847 | using ne \<open>b \<noteq> 0\<close> apply auto | 
| 848 | done | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 849 | ultimately show "real_of_int \<lfloor>real_of_int a / real_of_int b\<rfloor> < real_of_int a / real_of_int b" by arith | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 850 | qed | 
| 60698 | 851 | then show ?thesis | 
| 852 | using \<open>\<not> b dvd a\<close> by simp | |
| 853 | qed | |
| 19765 | 854 | |
| 60698 | 855 | qualified lemma compute_float_up[code]: "float_up p x = - float_down p (-x)" | 
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 856 | by transfer (simp add: round_down_uminus_eq) | 
| 60698 | 857 | |
| 858 | end | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 859 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 860 | |
| 60500 | 861 | subsection \<open>Compute bitlen of integers\<close> | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 862 | |
| 60698 | 863 | definition bitlen :: "int \<Rightarrow> int" | 
| 864 | where "bitlen a = (if a > 0 then \<lfloor>log 2 a\<rfloor> + 1 else 0)" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 865 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 866 | lemma bitlen_nonneg: "0 \<le> bitlen x" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 867 | proof - | 
| 60698 | 868 | have "-1 < log 2 (-x)" if "0 > x" | 
| 869 | proof - | |
| 870 | have "-1 = log 2 (inverse 2)" | |
| 871 | by (subst log_inverse) simp_all | |
| 872 | also have "\<dots> < log 2 (-x)" | |
| 873 | using \<open>0 > x\<close> by auto | |
| 874 | finally show ?thesis . | |
| 875 | qed | |
| 876 | then show ?thesis | |
| 877 | unfolding bitlen_def by (auto intro!: add_nonneg_nonneg) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 878 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 879 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 880 | lemma bitlen_bounds: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 881 | assumes "x > 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 882 | shows "2 ^ nat (bitlen x - 1) \<le> x \<and> x < 2 ^ nat (bitlen x)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 883 | proof | 
| 60698 | 884 | show "2 ^ nat (bitlen x - 1) \<le> x" | 
| 885 | proof - | |
| 61942 | 886 | have "(2::real) ^ nat \<lfloor>log 2 (real_of_int x)\<rfloor> = 2 powr real_of_int \<lfloor>log 2 (real_of_int x)\<rfloor>" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 887 | using powr_realpow[symmetric, of 2 "nat \<lfloor>log 2 (real_of_int x)\<rfloor>"] \<open>x > 0\<close> | 
| 60698 | 888 | by simp | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 889 | also have "\<dots> \<le> 2 powr log 2 (real_of_int x)" | 
| 60698 | 890 | by simp | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 891 | also have "\<dots> = real_of_int x" | 
| 60698 | 892 | using \<open>0 < x\<close> by simp | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 893 | finally have "2 ^ nat \<lfloor>log 2 (real_of_int x)\<rfloor> \<le> real_of_int x" | 
| 60698 | 894 | by simp | 
| 895 | then show ?thesis | |
| 896 | using \<open>0 < x\<close> by (simp add: bitlen_def) | |
| 897 | qed | |
| 898 | show "x < 2 ^ nat (bitlen x)" | |
| 899 | proof - | |
| 900 | have "x \<le> 2 powr (log 2 x)" | |
| 901 | using \<open>x > 0\<close> by simp | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 902 | also have "\<dots> < 2 ^ nat (\<lfloor>log 2 (real_of_int x)\<rfloor> + 1)" | 
| 60698 | 903 | apply (simp add: powr_realpow[symmetric]) | 
| 904 | using \<open>x > 0\<close> apply simp | |
| 905 | done | |
| 906 | finally show ?thesis | |
| 907 | using \<open>x > 0\<close> by (simp add: bitlen_def ac_simps) | |
| 908 | qed | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 909 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 910 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 911 | lemma bitlen_pow2[simp]: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 912 | assumes "b > 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 913 | shows "bitlen (b * 2 ^ c) = bitlen b + c" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 914 | proof - | 
| 60698 | 915 | from assms have "b * 2 ^ c > 0" | 
| 916 | by auto | |
| 917 | then show ?thesis | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 918 | using floor_add[of "log 2 b" c] assms | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 919 | apply (auto simp add: log_mult log_nat_power bitlen_def) | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 920 | by (metis add.right_neutral frac_lt_1 frac_of_int of_int_of_nat_eq) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 921 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 922 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 923 | lemma bitlen_Float: | 
| 53381 | 924 | fixes m e | 
| 925 | defines "f \<equiv> Float m e" | |
| 926 | shows "bitlen (\<bar>mantissa f\<bar>) + exponent f = (if m = 0 then 0 else bitlen \<bar>m\<bar> + e)" | |
| 927 | proof (cases "m = 0") | |
| 928 | case True | |
| 929 | then show ?thesis by (simp add: f_def bitlen_def Float_def) | |
| 930 | next | |
| 931 | case False | |
| 60698 | 932 | then have "f \<noteq> float_of 0" | 
| 47600 | 933 | unfolding real_of_float_eq by (simp add: f_def) | 
| 60698 | 934 | then have "mantissa f \<noteq> 0" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 935 | by (simp add: mantissa_noteq_0) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 936 | moreover | 
| 53381 | 937 | obtain i where "m = mantissa f * 2 ^ i" "e = exponent f - int i" | 
| 60500 | 938 | by (rule f_def[THEN denormalize_shift, OF \<open>f \<noteq> float_of 0\<close>]) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 939 | ultimately show ?thesis by (simp add: abs_mult) | 
| 53381 | 940 | qed | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 941 | |
| 60698 | 942 | context | 
| 943 | begin | |
| 944 | ||
| 945 | qualified lemma compute_bitlen[code]: "bitlen x = (if x > 0 then bitlen (x div 2) + 1 else 0)" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 946 | proof - | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 947 |   { assume "2 \<le> x"
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 948 | then have "\<lfloor>log 2 (x div 2)\<rfloor> + 1 = \<lfloor>log 2 (x - x mod 2)\<rfloor>" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 949 | by (simp add: log_mult zmod_zdiv_equality') | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 950 | also have "\<dots> = \<lfloor>log 2 (real_of_int x)\<rfloor>" | 
| 60698 | 951 | proof (cases "x mod 2 = 0") | 
| 952 | case True | |
| 953 | then show ?thesis by simp | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 954 | next | 
| 60698 | 955 | case False | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 956 | def n \<equiv> "\<lfloor>log 2 (real_of_int x)\<rfloor>" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 957 | then have "0 \<le> n" | 
| 60500 | 958 | using \<open>2 \<le> x\<close> by simp | 
| 60698 | 959 | from \<open>2 \<le> x\<close> False have "x mod 2 = 1" "\<not> 2 dvd x" | 
| 960 | by (auto simp add: dvd_eq_mod_eq_0) | |
| 961 | with \<open>2 \<le> x\<close> have "x \<noteq> 2 ^ nat n" | |
| 962 | by (cases "nat n") auto | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 963 | moreover | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 964 |       { have "real_of_int (2^nat n :: int) = 2 powr (nat n)"
 | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 965 | by (simp add: powr_realpow) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 966 | also have "\<dots> \<le> 2 powr (log 2 x)" | 
| 60500 | 967 | using \<open>2 \<le> x\<close> by (simp add: n_def del: powr_log_cancel) | 
| 968 | finally have "2^nat n \<le> x" using \<open>2 \<le> x\<close> by simp } | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 969 | ultimately have "2^nat n \<le> x - 1" by simp | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 970 | then have "2^nat n \<le> real_of_int (x - 1)" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 971 | using numeral_power_le_real_of_int_cancel_iff by blast | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 972 |       { have "n = \<lfloor>log 2 (2^nat n)\<rfloor>"
 | 
| 60500 | 973 | using \<open>0 \<le> n\<close> by (simp add: log_nat_power) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 974 | also have "\<dots> \<le> \<lfloor>log 2 (x - 1)\<rfloor>" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 975 | using \<open>2^nat n \<le> real_of_int (x - 1)\<close> \<open>0 \<le> n\<close> \<open>2 \<le> x\<close> by (auto intro: floor_mono) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 976 | finally have "n \<le> \<lfloor>log 2 (x - 1)\<rfloor>" . } | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 977 | moreover have "\<lfloor>log 2 (x - 1)\<rfloor> \<le> n" | 
| 60500 | 978 | using \<open>2 \<le> x\<close> by (auto simp add: n_def intro!: floor_mono) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 979 | ultimately show "\<lfloor>log 2 (x - x mod 2)\<rfloor> = \<lfloor>log 2 x\<rfloor>" | 
| 60500 | 980 | unfolding n_def \<open>x mod 2 = 1\<close> by auto | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 981 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 982 | finally have "\<lfloor>log 2 (x div 2)\<rfloor> + 1 = \<lfloor>log 2 x\<rfloor>" . } | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 983 | moreover | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 984 |   { assume "x < 2" "0 < x"
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 985 | then have "x = 1" by simp | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 986 | then have "\<lfloor>log 2 (real_of_int x)\<rfloor> = 0" by simp } | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 987 | ultimately show ?thesis | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 988 | unfolding bitlen_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 989 | by (auto simp: pos_imp_zdiv_pos_iff not_le) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 990 | qed | 
| 60698 | 991 | |
| 992 | end | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 993 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 994 | lemma float_gt1_scale: assumes "1 \<le> Float m e" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 995 | shows "0 \<le> e + (bitlen m - 1)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 996 | proof - | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 997 | have "0 < Float m e" using assms by auto | 
| 60698 | 998 | then have "0 < m" using powr_gt_zero[of 2 e] | 
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 999 | apply (auto simp: zero_less_mult_iff) | 
| 60698 | 1000 | using not_le powr_ge_pzero apply blast | 
| 1001 | done | |
| 1002 | then have "m \<noteq> 0" by auto | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1003 | show ?thesis | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1004 | proof (cases "0 \<le> e") | 
| 60698 | 1005 | case True | 
| 1006 | then show ?thesis | |
| 1007 | using \<open>0 < m\<close> by (simp add: bitlen_def) | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1008 | next | 
| 60698 | 1009 | case False | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1010 | have "(1::int) < 2" by simp | 
| 60698 | 1011 | let ?S = "2^(nat (-e))" | 
| 1012 | have "inverse (2 ^ nat (- e)) = 2 powr e" | |
| 1013 | using assms False powr_realpow[of 2 "nat (-e)"] | |
| 57862 | 1014 | by (auto simp: powr_minus field_simps) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1015 | then have "1 \<le> real_of_int m * inverse ?S" | 
| 60698 | 1016 | using assms False powr_realpow[of 2 "nat (-e)"] | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1017 | by (auto simp: powr_minus) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1018 | then have "1 * ?S \<le> real_of_int m * inverse ?S * ?S" | 
| 60698 | 1019 | by (rule mult_right_mono) auto | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1020 | then have "?S \<le> real_of_int m" | 
| 60698 | 1021 | unfolding mult.assoc by auto | 
| 1022 | then have "?S \<le> m" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1023 | unfolding of_int_le_iff[symmetric] by auto | 
| 60500 | 1024 | from this bitlen_bounds[OF \<open>0 < m\<close>, THEN conjunct2] | 
| 60698 | 1025 | have "nat (-e) < (nat (bitlen m))" | 
| 1026 | unfolding power_strict_increasing_iff[OF \<open>1 < 2\<close>, symmetric] | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1027 | by (rule order_le_less_trans) | 
| 60698 | 1028 | then have "-e < bitlen m" | 
| 1029 | using False by auto | |
| 1030 | then show ?thesis | |
| 1031 | by auto | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1032 | qed | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1033 | qed | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1034 | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1035 | lemma bitlen_div: | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1036 | assumes "0 < m" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1037 | shows "1 \<le> real_of_int m / 2^nat (bitlen m - 1)" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1038 | and "real_of_int m / 2^nat (bitlen m - 1) < 2" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1039 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1040 | let ?B = "2^nat(bitlen m - 1)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1041 | |
| 60500 | 1042 | have "?B \<le> m" using bitlen_bounds[OF \<open>0 <m\<close>] .. | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1043 | then have "1 * ?B \<le> real_of_int m" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1044 | unfolding of_int_le_iff[symmetric] by auto | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1045 | then show "1 \<le> real_of_int m / ?B" | 
| 60698 | 1046 | by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1047 | |
| 60698 | 1048 | have "m \<noteq> 0" | 
| 1049 | using assms by auto | |
| 1050 | have "0 \<le> bitlen m - 1" | |
| 1051 | using \<open>0 < m\<close> by (auto simp: bitlen_def) | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1052 | |
| 60698 | 1053 | have "m < 2^nat(bitlen m)" | 
| 1054 | using bitlen_bounds[OF \<open>0 <m\<close>] .. | |
| 1055 | also have "\<dots> = 2^nat(bitlen m - 1 + 1)" | |
| 1056 | using \<open>0 < m\<close> by (auto simp: bitlen_def) | |
| 1057 | also have "\<dots> = ?B * 2" | |
| 1058 | unfolding nat_add_distrib[OF \<open>0 \<le> bitlen m - 1\<close> zero_le_one] by auto | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1059 | finally have "real_of_int m < 2 * ?B" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1060 | by (metis (full_types) mult.commute power.simps(2) real_of_int_less_numeral_power_cancel_iff) | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1061 | then have "real_of_int m / ?B < 2 * ?B / ?B" | 
| 60698 | 1062 | by (rule divide_strict_right_mono) auto | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1063 | then show "real_of_int m / ?B < 2" | 
| 60698 | 1064 | by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1065 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1066 | |
| 60698 | 1067 | |
| 60500 | 1068 | subsection \<open>Truncating Real Numbers\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1069 | |
| 60698 | 1070 | definition truncate_down::"nat \<Rightarrow> real \<Rightarrow> real" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1071 | where "truncate_down prec x = round_down (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>) x" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1072 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1073 | lemma truncate_down: "truncate_down prec x \<le> x" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1074 | using round_down by (simp add: truncate_down_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1075 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1076 | lemma truncate_down_le: "x \<le> y \<Longrightarrow> truncate_down prec x \<le> y" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1077 | by (rule order_trans[OF truncate_down]) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1078 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1079 | lemma truncate_down_zero[simp]: "truncate_down prec 0 = 0" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1080 | by (simp add: truncate_down_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1081 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1082 | lemma truncate_down_float[simp]: "truncate_down p x \<in> float" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1083 | by (auto simp: truncate_down_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1084 | |
| 60698 | 1085 | definition truncate_up::"nat \<Rightarrow> real \<Rightarrow> real" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1086 | where "truncate_up prec x = round_up (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>) x" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1087 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1088 | lemma truncate_up: "x \<le> truncate_up prec x" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1089 | using round_up by (simp add: truncate_up_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1090 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1091 | lemma truncate_up_le: "x \<le> y \<Longrightarrow> x \<le> truncate_up prec y" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1092 | by (rule order_trans[OF _ truncate_up]) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1093 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1094 | lemma truncate_up_zero[simp]: "truncate_up prec 0 = 0" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1095 | by (simp add: truncate_up_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1096 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1097 | lemma truncate_up_uminus_eq: "truncate_up prec (-x) = - truncate_down prec x" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1098 | and truncate_down_uminus_eq: "truncate_down prec (-x) = - truncate_up prec x" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1099 | by (auto simp: truncate_up_def round_up_def truncate_down_def round_down_def ceiling_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1100 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1101 | lemma truncate_up_float[simp]: "truncate_up p x \<in> float" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1102 | by (auto simp: truncate_up_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1103 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1104 | lemma mult_powr_eq: "0 < b \<Longrightarrow> b \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow> x * b powr y = b powr (y + log b x)" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1105 | by (simp_all add: powr_add) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1106 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1107 | lemma truncate_down_pos: | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1108 | assumes "x > 0" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1109 | shows "truncate_down p x > 0" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1110 | proof - | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1111 | have "0 \<le> log 2 x - real_of_int \<lfloor>log 2 x\<rfloor>" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1112 | by (simp add: algebra_simps) | 
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61649diff
changeset | 1113 | with assms | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1114 | show ?thesis | 
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61649diff
changeset | 1115 | apply (auto simp: truncate_down_def round_down_def mult_powr_eq | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1116 | intro!: ge_one_powr_ge_zero mult_pos_pos) | 
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61649diff
changeset | 1117 | by linarith | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1118 | qed | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1119 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1120 | lemma truncate_down_nonneg: "0 \<le> y \<Longrightarrow> 0 \<le> truncate_down prec y" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1121 | by (auto simp: truncate_down_def round_down_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1122 | |
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1123 | lemma truncate_down_ge1: "1 \<le> x \<Longrightarrow> 1 \<le> truncate_down p x" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1124 | apply (auto simp: truncate_down_def algebra_simps intro!: round_down_ge1) | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1125 | apply linarith | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1126 | done | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1127 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1128 | lemma truncate_up_nonpos: "x \<le> 0 \<Longrightarrow> truncate_up prec x \<le> 0" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1129 | by (auto simp: truncate_up_def round_up_def intro!: mult_nonpos_nonneg) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1130 | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1131 | lemma truncate_up_le1: | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1132 | assumes "x \<le> 1" | 
| 60698 | 1133 | shows "truncate_up p x \<le> 1" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1134 | proof - | 
| 60698 | 1135 | consider "x \<le> 0" | "x > 0" | 
| 1136 | by arith | |
| 1137 | then show ?thesis | |
| 1138 | proof cases | |
| 1139 | case 1 | |
| 1140 | with truncate_up_nonpos[OF this, of p] show ?thesis | |
| 1141 | by simp | |
| 1142 | next | |
| 1143 | case 2 | |
| 1144 | then have le: "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<le> 0" | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1145 | using assms by (auto simp: log_less_iff) | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1146 | from assms have "0 \<le> int p" by simp | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1147 | from add_mono[OF this le] | 
| 60698 | 1148 | show ?thesis | 
| 1149 | using assms by (simp add: truncate_up_def round_up_le1 add_mono) | |
| 1150 | qed | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1151 | qed | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1152 | |
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1153 | lemma truncate_down_shift_int: "truncate_down p (x * 2 powr real_of_int k) = truncate_down p x * 2 powr k" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1154 | by (cases "x = 0") | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1155 | (simp_all add: algebra_simps abs_mult log_mult truncate_down_def round_down_shift[of _ _ k, simplified]) | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1156 | |
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1157 | lemma truncate_down_shift_nat: "truncate_down p (x * 2 powr real k) = truncate_down p x * 2 powr k" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1158 | by (metis of_int_of_nat_eq truncate_down_shift_int) | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1159 | |
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1160 | lemma truncate_up_shift_int: "truncate_up p (x * 2 powr real_of_int k) = truncate_up p x * 2 powr k" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1161 | by (cases "x = 0") | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1162 | (simp_all add: algebra_simps abs_mult log_mult truncate_up_def round_up_shift[of _ _ k, simplified]) | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1163 | |
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1164 | lemma truncate_up_shift_nat: "truncate_up p (x * 2 powr real k) = truncate_up p x * 2 powr k" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1165 | by (metis of_int_of_nat_eq truncate_up_shift_int) | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1166 | |
| 60698 | 1167 | |
| 60500 | 1168 | subsection \<open>Truncating Floats\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1169 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1170 | lift_definition float_round_up :: "nat \<Rightarrow> float \<Rightarrow> float" is truncate_up | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1171 | by (simp add: truncate_up_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1172 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1173 | lemma float_round_up: "real_of_float x \<le> real_of_float (float_round_up prec x)" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1174 | using truncate_up by transfer simp | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1175 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1176 | lemma float_round_up_zero[simp]: "float_round_up prec 0 = 0" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1177 | by transfer simp | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1178 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1179 | lift_definition float_round_down :: "nat \<Rightarrow> float \<Rightarrow> float" is truncate_down | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1180 | by (simp add: truncate_down_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1181 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1182 | lemma float_round_down: "real_of_float (float_round_down prec x) \<le> real_of_float x" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1183 | using truncate_down by transfer simp | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1184 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1185 | lemma float_round_down_zero[simp]: "float_round_down prec 0 = 0" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1186 | by transfer simp | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1187 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1188 | lemmas float_round_up_le = order_trans[OF _ float_round_up] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1189 | and float_round_down_le = order_trans[OF float_round_down] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1190 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1191 | lemma minus_float_round_up_eq: "- float_round_up prec x = float_round_down prec (- x)" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1192 | and minus_float_round_down_eq: "- float_round_down prec x = float_round_up prec (- x)" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1193 | by (transfer, simp add: truncate_down_uminus_eq truncate_up_uminus_eq)+ | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1194 | |
| 60698 | 1195 | context | 
| 1196 | begin | |
| 1197 | ||
| 1198 | qualified lemma compute_float_round_down[code]: | |
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1199 | "float_round_down prec (Float m e) = (let d = bitlen \<bar>m\<bar> - int prec - 1 in | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1200 | if 0 < d then Float (div_twopow m (nat d)) (e + d) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1201 | else Float m e)" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1202 | using Float.compute_float_down[of "Suc prec - bitlen \<bar>m\<bar> - e" m e, symmetric] | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1203 | by transfer | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1204 | (simp add: field_simps abs_mult log_mult bitlen_def truncate_down_def | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1205 | cong del: if_weak_cong) | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1206 | |
| 60698 | 1207 | qualified lemma compute_float_round_up[code]: | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1208 | "float_round_up prec x = - float_round_down prec (-x)" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1209 | by transfer (simp add: truncate_down_uminus_eq) | 
| 60698 | 1210 | |
| 1211 | end | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1212 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1213 | |
| 60500 | 1214 | subsection \<open>Approximation of positive rationals\<close> | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1215 | |
| 60698 | 1216 | lemma div_mult_twopow_eq: | 
| 1217 | fixes a b :: nat | |
| 1218 | shows "a div ((2::nat) ^ n) div b = a div (b * 2 ^ n)" | |
| 1219 | by (cases "b = 0") (simp_all add: div_mult2_eq[symmetric] ac_simps) | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1220 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1221 | lemma real_div_nat_eq_floor_of_divide: | 
| 59984 
4f1eccec320c
conversion between division on nat/int and division in archmedean fields
 haftmann parents: 
59554diff
changeset | 1222 | fixes a b :: nat | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1223 | shows "a div b = real_of_int \<lfloor>a / b\<rfloor>" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1224 | by (simp add: floor_divide_of_nat_eq [of a b]) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1225 | |
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1226 | definition "rat_precision prec x y = | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1227 | (let d = bitlen x - bitlen y in int prec - d + | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1228 | (if Float (abs x) 0 < Float (abs y) d then 1 else 0))" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1229 | |
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1230 | lemma floor_log_divide_eq: | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1231 | assumes "i > 0" "j > 0" "p > 1" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1232 | shows "\<lfloor>log p (i / j)\<rfloor> = floor (log p i) - floor (log p j) - | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1233 | (if i \<ge> j * p powr (floor (log p i) - floor (log p j)) then 0 else 1)" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1234 | proof - | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1235 | let ?l = "log p" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1236 | let ?fl = "\<lambda>x. floor (?l x)" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1237 | have "\<lfloor>?l (i / j)\<rfloor> = \<lfloor>?l i - ?l j\<rfloor>" using assms | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1238 | by (auto simp: log_divide) | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1239 | also have "\<dots> = floor (real_of_int (?fl i - ?fl j) + (?l i - ?fl i - (?l j - ?fl j)))" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1240 | (is "_ = floor (_ + ?r)") | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1241 | by (simp add: algebra_simps) | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1242 | also note floor_add2 | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1243 | also note \<open>p > 1\<close> | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1244 | note powr = powr_le_cancel_iff[symmetric, OF \<open>1 < p\<close>, THEN iffD2] | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1245 | note powr_strict = powr_less_cancel_iff[symmetric, OF \<open>1 < p\<close>, THEN iffD2] | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1246 | have "floor ?r = (if i \<ge> j * p powr (?fl i - ?fl j) then 0 else -1)" (is "_ = ?if") | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1247 | using assms | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1248 | by (linarith | | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1249 | auto | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1250 | intro!: floor_eq2 | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1251 | intro: powr_strict powr | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1252 | simp: powr_divide2[symmetric] powr_add divide_simps algebra_simps bitlen_def)+ | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1253 | finally | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1254 | show ?thesis by simp | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1255 | qed | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1256 | |
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1257 | lemma truncate_down_rat_precision: | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1258 | "truncate_down prec (real x / real y) = round_down (rat_precision prec x y) (real x / real y)" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1259 | and truncate_up_rat_precision: | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1260 | "truncate_up prec (real x / real y) = round_up (rat_precision prec x y) (real x / real y)" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1261 | unfolding truncate_down_def truncate_up_def rat_precision_def | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1262 | by (cases x; cases y) (auto simp: floor_log_divide_eq algebra_simps bitlen_def) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1263 | |
| 47600 | 1264 | lift_definition lapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1265 | is "\<lambda>prec (x::nat) (y::nat). truncate_down prec (x / y)" | 
| 60698 | 1266 | by simp | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1267 | |
| 60698 | 1268 | context | 
| 1269 | begin | |
| 1270 | ||
| 1271 | qualified lemma compute_lapprox_posrat[code]: | |
| 53381 | 1272 | fixes prec x y | 
| 1273 | shows "lapprox_posrat prec x y = | |
| 1274 | (let | |
| 60698 | 1275 | l = rat_precision prec x y; | 
| 1276 | d = if 0 \<le> l then x * 2^nat l div y else x div 2^nat (- l) div y | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1277 | in normfloat (Float d (- l)))" | 
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 1278 | unfolding div_mult_twopow_eq | 
| 47600 | 1279 | by transfer | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1280 | (simp add: round_down_def powr_int real_div_nat_eq_floor_of_divide field_simps Let_def | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1281 | truncate_down_rat_precision del: two_powr_minus_int_float) | 
| 60698 | 1282 | |
| 1283 | end | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1284 | |
| 47600 | 1285 | lift_definition rapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1286 | is "\<lambda>prec (x::nat) (y::nat). truncate_up prec (x / y)" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1287 | by simp | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1288 | |
| 60376 | 1289 | context | 
| 1290 | begin | |
| 1291 | ||
| 1292 | qualified lemma compute_rapprox_posrat[code]: | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1293 | fixes prec x y | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1294 | defines "l \<equiv> rat_precision prec x y" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1295 | shows "rapprox_posrat prec x y = (let | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1296 | l = l ; | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60698diff
changeset | 1297 | (r, s) = if 0 \<le> l then (x * 2^nat l, y) else (x, y * 2^nat(-l)) ; | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60698diff
changeset | 1298 | d = r div s ; | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60698diff
changeset | 1299 | m = r mod s | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1300 | in normfloat (Float (d + (if m = 0 \<or> y = 0 then 0 else 1)) (- l)))" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1301 | proof (cases "y = 0") | 
| 60698 | 1302 | assume "y = 0" | 
| 1303 | then show ?thesis by transfer simp | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1304 | next | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1305 | assume "y \<noteq> 0" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1306 | show ?thesis | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1307 | proof (cases "0 \<le> l") | 
| 60698 | 1308 | case True | 
| 56777 | 1309 | def x' \<equiv> "x * 2 ^ nat l" | 
| 60698 | 1310 | have "int x * 2 ^ nat l = x'" | 
| 62348 | 1311 | by (simp add: x'_def of_nat_mult of_nat_power) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1312 | moreover have "real x * 2 powr l = real x'" | 
| 60500 | 1313 | by (simp add: powr_realpow[symmetric] \<open>0 \<le> l\<close> x'_def) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1314 | ultimately show ?thesis | 
| 60500 | 1315 | using ceil_divide_floor_conv[of y x'] powr_realpow[of 2 "nat l"] \<open>0 \<le> l\<close> \<open>y \<noteq> 0\<close> | 
| 47600 | 1316 | l_def[symmetric, THEN meta_eq_to_obj_eq] | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1317 | apply transfer | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1318 | apply (auto simp add: round_up_def truncate_up_rat_precision) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1319 | by (metis floor_divide_of_int_eq of_int_of_nat_eq) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1320 | next | 
| 60698 | 1321 | case False | 
| 56777 | 1322 | def y' \<equiv> "y * 2 ^ nat (- l)" | 
| 60500 | 1323 | from \<open>y \<noteq> 0\<close> have "y' \<noteq> 0" by (simp add: y'_def) | 
| 62348 | 1324 | have "int y * 2 ^ nat (- l) = y'" by (simp add: y'_def of_nat_mult of_nat_power) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1325 | moreover have "real x * real_of_int (2::int) powr real_of_int l / real y = x / real y'" | 
| 60500 | 1326 | using \<open>\<not> 0 \<le> l\<close> | 
| 57862 | 1327 | by (simp add: powr_realpow[symmetric] powr_minus y'_def field_simps) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1328 | ultimately show ?thesis | 
| 60500 | 1329 | using ceil_divide_floor_conv[of y' x] \<open>\<not> 0 \<le> l\<close> \<open>y' \<noteq> 0\<close> \<open>y \<noteq> 0\<close> | 
| 47600 | 1330 | l_def[symmetric, THEN meta_eq_to_obj_eq] | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1331 | apply transfer | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1332 | apply (auto simp add: round_up_def ceil_divide_floor_conv truncate_up_rat_precision) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1333 | by (metis floor_divide_of_int_eq of_int_of_nat_eq) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1334 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1335 | qed | 
| 60376 | 1336 | |
| 1337 | end | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1338 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1339 | lemma rat_precision_pos: | 
| 60698 | 1340 | assumes "0 \<le> x" | 
| 1341 | and "0 < y" | |
| 1342 | and "2 * x < y" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1343 | shows "rat_precision n (int x) (int y) > 0" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1344 | proof - | 
| 60698 | 1345 | have "0 < x \<Longrightarrow> log 2 x + 1 = log 2 (2 * x)" | 
| 1346 | by (simp add: log_mult) | |
| 1347 | then have "bitlen (int x) < bitlen (int y)" | |
| 1348 | using assms | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1349 | by (simp add: bitlen_def del: floor_add_one) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1350 | (auto intro!: floor_mono simp add: floor_add_one[symmetric] simp del: floor_add floor_add_one) | 
| 60698 | 1351 | then show ?thesis | 
| 1352 | using assms | |
| 1353 | by (auto intro!: pos_add_strict simp add: field_simps rat_precision_def) | |
| 29804 
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 | |
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1356 | lemma rapprox_posrat_less1: | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1357 | "0 \<le> x \<Longrightarrow> 0 < y \<Longrightarrow> 2 * x < y \<Longrightarrow> real_of_float (rapprox_posrat n x y) < 1" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1358 | by transfer (simp add: rat_precision_pos round_up_less1 truncate_up_rat_precision) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1359 | |
| 47600 | 1360 | lift_definition lapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1361 | "\<lambda>prec (x::int) (y::int). truncate_down prec (x / y)" | 
| 60698 | 1362 | by simp | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1363 | |
| 60698 | 1364 | context | 
| 1365 | begin | |
| 1366 | ||
| 1367 | qualified lemma compute_lapprox_rat[code]: | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1368 | "lapprox_rat prec x y = | 
| 60698 | 1369 | (if y = 0 then 0 | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1370 | else if 0 \<le> x then | 
| 60698 | 1371 | (if 0 < y then lapprox_posrat prec (nat x) (nat y) | 
| 53381 | 1372 | else - (rapprox_posrat prec (nat x) (nat (-y)))) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1373 | else (if 0 < y | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1374 | then - (rapprox_posrat prec (nat (-x)) (nat y)) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1375 | else lapprox_posrat prec (nat (-x)) (nat (-y))))" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1376 | by transfer (simp add: truncate_up_uminus_eq) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1377 | |
| 47600 | 1378 | lift_definition rapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1379 | "\<lambda>prec (x::int) (y::int). truncate_up prec (x / y)" | 
| 60698 | 1380 | by simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1381 | |
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 1382 | lemma "rapprox_rat = rapprox_posrat" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 1383 | by transfer auto | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 1384 | |
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 1385 | lemma "lapprox_rat = lapprox_posrat" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 1386 | by transfer auto | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 1387 | |
| 60698 | 1388 | qualified lemma compute_rapprox_rat[code]: | 
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 1389 | "rapprox_rat prec x y = - lapprox_rat prec (-x) y" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1390 | by transfer (simp add: truncate_down_uminus_eq) | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1391 | |
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1392 | qualified lemma compute_truncate_down[code]: "truncate_down p (Ratreal r) = (let (a, b) = quotient_of r in lapprox_rat p a b)" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1393 | by transfer (auto split: prod.split simp: of_rat_divide dest!: quotient_of_div) | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1394 | |
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1395 | qualified lemma compute_truncate_up[code]: "truncate_up p (Ratreal r) = (let (a, b) = quotient_of r in rapprox_rat p a b)" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1396 | by transfer (auto split: prod.split simp: of_rat_divide dest!: quotient_of_div) | 
| 60698 | 1397 | |
| 1398 | end | |
| 1399 | ||
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1400 | |
| 60500 | 1401 | subsection \<open>Division\<close> | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1402 | |
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1403 | definition "real_divl prec a b = truncate_down prec (a / b)" | 
| 54782 | 1404 | |
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1405 | definition "real_divr prec a b = truncate_up prec (a / b)" | 
| 54782 | 1406 | |
| 1407 | lift_definition float_divl :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is real_divl | |
| 1408 | by (simp add: real_divl_def) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1409 | |
| 60698 | 1410 | context | 
| 1411 | begin | |
| 1412 | ||
| 1413 | qualified lemma compute_float_divl[code]: | |
| 47600 | 1414 | "float_divl prec (Float m1 s1) (Float m2 s2) = lapprox_rat prec m1 m2 * Float 1 (s1 - s2)" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1415 | apply transfer | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1416 | unfolding real_divl_def of_int_1 mult_1 truncate_down_shift_int[symmetric] | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1417 | apply (simp add: powr_divide2[symmetric] powr_minus) | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1418 | done | 
| 47600 | 1419 | |
| 54782 | 1420 | lift_definition float_divr :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is real_divr | 
| 1421 | by (simp add: real_divr_def) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1422 | |
| 60698 | 1423 | qualified lemma compute_float_divr[code]: | 
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 1424 | "float_divr prec x y = - float_divl prec (-x) y" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1425 | by transfer (simp add: real_divr_def real_divl_def truncate_down_uminus_eq) | 
| 60698 | 1426 | |
| 1427 | end | |
| 47600 | 1428 | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1429 | |
| 60500 | 1430 | subsection \<open>Approximate Power\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1431 | |
| 60698 | 1432 | lemma div2_less_self[termination_simp]: | 
| 1433 | fixes n :: nat | |
| 1434 | shows "odd n \<Longrightarrow> n div 2 < n" | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1435 | by (simp add: odd_pos) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1436 | |
| 60698 | 1437 | fun power_down :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real" | 
| 1438 | where | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1439 | "power_down p x 0 = 1" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1440 | | "power_down p x (Suc n) = | 
| 60698 | 1441 | (if odd n then truncate_down (Suc p) ((power_down p x (Suc n div 2))\<^sup>2) | 
| 1442 | else truncate_down (Suc p) (x * power_down p x n))" | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1443 | |
| 60698 | 1444 | fun power_up :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real" | 
| 1445 | where | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1446 | "power_up p x 0 = 1" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1447 | | "power_up p x (Suc n) = | 
| 60698 | 1448 | (if odd n then truncate_up p ((power_up p x (Suc n div 2))\<^sup>2) | 
| 1449 | else truncate_up p (x * power_up p x n))" | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1450 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1451 | lift_definition power_up_fl :: "nat \<Rightarrow> float \<Rightarrow> nat \<Rightarrow> float" is power_up | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1452 | by (induct_tac rule: power_up.induct) simp_all | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1453 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1454 | lift_definition power_down_fl :: "nat \<Rightarrow> float \<Rightarrow> nat \<Rightarrow> float" is power_down | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1455 | by (induct_tac rule: power_down.induct) simp_all | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1456 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1457 | lemma power_float_transfer[transfer_rule]: | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1458 | "(rel_fun pcr_float (rel_fun op = pcr_float)) op ^ op ^" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1459 | unfolding power_def | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1460 | by transfer_prover | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1461 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1462 | lemma compute_power_up_fl[code]: | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1463 | "power_up_fl p x 0 = 1" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1464 | "power_up_fl p x (Suc n) = | 
| 60698 | 1465 | (if odd n then float_round_up p ((power_up_fl p x (Suc n div 2))\<^sup>2) | 
| 1466 | else float_round_up p (x * power_up_fl p x n))" | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1467 | and compute_power_down_fl[code]: | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1468 | "power_down_fl p x 0 = 1" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1469 | "power_down_fl p x (Suc n) = | 
| 60698 | 1470 | (if odd n then float_round_down (Suc p) ((power_down_fl p x (Suc n div 2))\<^sup>2) | 
| 1471 | else float_round_down (Suc p) (x * power_down_fl p x n))" | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1472 | unfolding atomize_conj | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1473 | by transfer simp | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1474 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1475 | lemma power_down_pos: "0 < x \<Longrightarrow> 0 < power_down p x n" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1476 | by (induct p x n rule: power_down.induct) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1477 | (auto simp del: odd_Suc_div_two intro!: truncate_down_pos) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1478 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1479 | lemma power_down_nonneg: "0 \<le> x \<Longrightarrow> 0 \<le> power_down p x n" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1480 | by (induct p x n rule: power_down.induct) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1481 | (auto simp del: odd_Suc_div_two intro!: truncate_down_nonneg mult_nonneg_nonneg) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1482 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1483 | lemma power_down: "0 \<le> x \<Longrightarrow> power_down p x n \<le> x ^ n" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1484 | proof (induct p x n rule: power_down.induct) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1485 | case (2 p x n) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1486 |   {
 | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1487 | assume "odd n" | 
| 60698 | 1488 | then have "(power_down p x (Suc n div 2)) ^ 2 \<le> (x ^ (Suc n div 2)) ^ 2" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1489 | using 2 | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1490 | by (auto intro: power_mono power_down_nonneg simp del: odd_Suc_div_two) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1491 | also have "\<dots> = x ^ (Suc n div 2 * 2)" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1492 | by (simp add: power_mult[symmetric]) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1493 | also have "Suc n div 2 * 2 = Suc n" | 
| 60500 | 1494 | using \<open>odd n\<close> by presburger | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1495 | finally have ?case | 
| 60500 | 1496 | using \<open>odd n\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1497 | by (auto intro!: truncate_down_le simp del: odd_Suc_div_two) | 
| 60698 | 1498 | } | 
| 1499 | then show ?case | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1500 | by (auto intro!: truncate_down_le mult_left_mono 2 mult_nonneg_nonneg power_down_nonneg) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1501 | qed simp | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1502 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1503 | lemma power_up: "0 \<le> x \<Longrightarrow> x ^ n \<le> power_up p x n" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1504 | proof (induct p x n rule: power_up.induct) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1505 | case (2 p x n) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1506 |   {
 | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1507 | assume "odd n" | 
| 60698 | 1508 | then have "Suc n = Suc n div 2 * 2" | 
| 60500 | 1509 | using \<open>odd n\<close> even_Suc by presburger | 
| 60698 | 1510 | then have "x ^ Suc n \<le> (x ^ (Suc n div 2))\<^sup>2" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1511 | by (simp add: power_mult[symmetric]) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1512 | also have "\<dots> \<le> (power_up p x (Suc n div 2))\<^sup>2" | 
| 60500 | 1513 | using 2 \<open>odd n\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1514 | by (auto intro: power_mono simp del: odd_Suc_div_two ) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1515 | finally have ?case | 
| 60500 | 1516 | using \<open>odd n\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1517 | by (auto intro!: truncate_up_le simp del: odd_Suc_div_two ) | 
| 60698 | 1518 | } | 
| 1519 | then show ?case | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1520 | by (auto intro!: truncate_up_le mult_left_mono 2) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1521 | qed simp | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1522 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1523 | lemmas power_up_le = order_trans[OF _ power_up] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1524 | and power_up_less = less_le_trans[OF _ power_up] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1525 | and power_down_le = order_trans[OF power_down] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1526 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1527 | lemma power_down_fl: "0 \<le> x \<Longrightarrow> power_down_fl p x n \<le> x ^ n" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1528 | by transfer (rule power_down) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1529 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1530 | lemma power_up_fl: "0 \<le> x \<Longrightarrow> x ^ n \<le> power_up_fl p x n" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1531 | by transfer (rule power_up) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1532 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1533 | lemma real_power_up_fl: "real_of_float (power_up_fl p x n) = power_up p x n" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1534 | by transfer simp | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1535 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1536 | lemma real_power_down_fl: "real_of_float (power_down_fl p x n) = power_down p x n" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1537 | by transfer simp | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1538 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1539 | |
| 60500 | 1540 | subsection \<open>Approximate Addition\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1541 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1542 | definition "plus_down prec x y = truncate_down prec (x + y)" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1543 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1544 | definition "plus_up prec x y = truncate_up prec (x + y)" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1545 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1546 | lemma float_plus_down_float[intro, simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> plus_down p x y \<in> float" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1547 | by (simp add: plus_down_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1548 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1549 | lemma float_plus_up_float[intro, simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> plus_up p x y \<in> float" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1550 | by (simp add: plus_up_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1551 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1552 | lift_definition float_plus_down::"nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is plus_down .. | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1553 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1554 | lift_definition float_plus_up::"nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is plus_up .. | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1555 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1556 | lemma plus_down: "plus_down prec x y \<le> x + y" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1557 | and plus_up: "x + y \<le> plus_up prec x y" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1558 | by (auto simp: plus_down_def truncate_down plus_up_def truncate_up) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1559 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1560 | lemma float_plus_down: "real_of_float (float_plus_down prec x y) \<le> x + y" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1561 | and float_plus_up: "x + y \<le> real_of_float (float_plus_up prec x y)" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1562 | by (transfer, rule plus_down plus_up)+ | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1563 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1564 | lemmas plus_down_le = order_trans[OF plus_down] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1565 | and plus_up_le = order_trans[OF _ plus_up] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1566 | and float_plus_down_le = order_trans[OF float_plus_down] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1567 | and float_plus_up_le = order_trans[OF _ float_plus_up] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1568 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1569 | lemma compute_plus_up[code]: "plus_up p x y = - plus_down p (-x) (-y)" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1570 | using truncate_down_uminus_eq[of p "x + y"] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1571 | by (auto simp: plus_down_def plus_up_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1572 | |
| 60698 | 1573 | lemma truncate_down_log2_eqI: | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1574 | assumes "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> = \<lfloor>log 2 \<bar>y\<bar>\<rfloor>" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1575 | assumes "\<lfloor>x * 2 powr (p - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>)\<rfloor> = \<lfloor>y * 2 powr (p - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>)\<rfloor>" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1576 | shows "truncate_down p x = truncate_down p y" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1577 | using assms by (auto simp: truncate_down_def round_down_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1578 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1579 | lemma bitlen_eq_zero_iff: "bitlen x = 0 \<longleftrightarrow> x \<le> 0" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1580 | by (clarsimp simp add: bitlen_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1581 | (metis Float.compute_bitlen add.commute bitlen_def bitlen_nonneg less_add_same_cancel2 not_less | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1582 | zero_less_one) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1583 | |
| 60698 | 1584 | lemma sum_neq_zeroI: | 
| 1585 | fixes a k :: real | |
| 61945 | 1586 | shows "\<bar>a\<bar> \<ge> k \<Longrightarrow> \<bar>b\<bar> < k \<Longrightarrow> a + b \<noteq> 0" | 
| 1587 | and "\<bar>a\<bar> > k \<Longrightarrow> \<bar>b\<bar> \<le> k \<Longrightarrow> a + b \<noteq> 0" | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1588 | by auto | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1589 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1590 | lemma abs_real_le_2_powr_bitlen[simp]: "\<bar>real_of_int m2\<bar> < 2 powr real_of_int (bitlen \<bar>m2\<bar>)" | 
| 60698 | 1591 | proof (cases "m2 = 0") | 
| 1592 | case True | |
| 1593 | then show ?thesis by simp | |
| 1594 | next | |
| 1595 | case False | |
| 1596 | then have "\<bar>m2\<bar> < 2 ^ nat (bitlen \<bar>m2\<bar>)" | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1597 | using bitlen_bounds[of "\<bar>m2\<bar>"] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1598 | by (auto simp: powr_add bitlen_nonneg) | 
| 60698 | 1599 | then show ?thesis | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61639diff
changeset | 1600 | by (metis bitlen_nonneg powr_int of_int_abs real_of_int_less_numeral_power_cancel_iff zero_less_numeral) | 
| 60698 | 1601 | qed | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1602 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1603 | lemma floor_sum_times_2_powr_sgn_eq: | 
| 60698 | 1604 | fixes ai p q :: int | 
| 1605 | and a b :: real | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1606 | assumes "a * 2 powr p = ai" | 
| 61945 | 1607 | and b_le_1: "\<bar>b * 2 powr (p + 1)\<bar> \<le> 1" | 
| 60698 | 1608 | and leqp: "q \<le> p" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1609 | shows "\<lfloor>(a + b) * 2 powr q\<rfloor> = \<lfloor>(2 * ai + sgn b) * 2 powr (q - p - 1)\<rfloor>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1610 | proof - | 
| 60698 | 1611 | consider "b = 0" | "b > 0" | "b < 0" by arith | 
| 1612 | then show ?thesis | |
| 1613 | proof cases | |
| 1614 | case 1 | |
| 1615 | then show ?thesis | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1616 | by (simp add: assms(1)[symmetric] powr_add[symmetric] algebra_simps powr_mult_base) | 
| 60698 | 1617 | next | 
| 1618 | case 2 | |
| 61945 | 1619 | then have "b * 2 powr p < \<bar>b * 2 powr (p + 1)\<bar>" | 
| 60698 | 1620 | by simp | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1621 | also note b_le_1 | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1622 | finally have b_less_1: "b * 2 powr real_of_int p < 1" . | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1623 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1624 | from b_less_1 \<open>b > 0\<close> have floor_eq: "\<lfloor>b * 2 powr real_of_int p\<rfloor> = 0" "\<lfloor>sgn b / 2\<rfloor> = 0" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1625 | by (simp_all add: floor_eq_iff) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1626 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1627 | have "\<lfloor>(a + b) * 2 powr q\<rfloor> = \<lfloor>(a + b) * 2 powr p * 2 powr (q - p)\<rfloor>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1628 | by (simp add: algebra_simps powr_realpow[symmetric] powr_add[symmetric]) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1629 | also have "\<dots> = \<lfloor>(ai + b * 2 powr p) * 2 powr (q - p)\<rfloor>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1630 | by (simp add: assms algebra_simps) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1631 | also have "\<dots> = \<lfloor>(ai + b * 2 powr p) / real_of_int ((2::int) ^ nat (p - q))\<rfloor>" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1632 | using assms | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1633 | by (simp add: algebra_simps powr_realpow[symmetric] divide_powr_uminus powr_add[symmetric]) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1634 | also have "\<dots> = \<lfloor>ai / real_of_int ((2::int) ^ nat (p - q))\<rfloor>" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1635 | by (simp del: of_int_power add: floor_divide_real_eq_div floor_eq) | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1636 | finally have "\<lfloor>(a + b) * 2 powr real_of_int q\<rfloor> = \<lfloor>real_of_int ai / real_of_int ((2::int) ^ nat (p - q))\<rfloor>" . | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1637 | moreover | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1638 |     {
 | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1639 | have "\<lfloor>(2 * ai + sgn b) * 2 powr (real_of_int (q - p) - 1)\<rfloor> = \<lfloor>(ai + sgn b / 2) * 2 powr (q - p)\<rfloor>" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1640 | by (subst powr_divide2[symmetric]) (simp add: field_simps) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1641 | also have "\<dots> = \<lfloor>(ai + sgn b / 2) / real_of_int ((2::int) ^ nat (p - q))\<rfloor>" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1642 | using leqp by (simp add: powr_realpow[symmetric] powr_divide2[symmetric]) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1643 | also have "\<dots> = \<lfloor>ai / real_of_int ((2::int) ^ nat (p - q))\<rfloor>" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1644 | by (simp del: of_int_power add: floor_divide_real_eq_div floor_eq) | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1645 | finally | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1646 | have "\<lfloor>(2 * ai + (sgn b)) * 2 powr (real_of_int (q - p) - 1)\<rfloor> = | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1647 | \<lfloor>real_of_int ai / real_of_int ((2::int) ^ nat (p - q))\<rfloor>" . | 
| 60698 | 1648 | } | 
| 1649 | ultimately show ?thesis by simp | |
| 1650 | next | |
| 1651 | case 3 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1652 | then have floor_eq: "\<lfloor>b * 2 powr (real_of_int p + 1)\<rfloor> = -1" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1653 | using b_le_1 | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1654 | by (auto simp: floor_eq_iff algebra_simps pos_divide_le_eq[symmetric] abs_if divide_powr_uminus | 
| 62390 | 1655 | intro!: mult_neg_pos split: if_split_asm) | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1656 | have "\<lfloor>(a + b) * 2 powr q\<rfloor> = \<lfloor>(2*a + 2*b) * 2 powr p * 2 powr (q - p - 1)\<rfloor>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1657 | by (simp add: algebra_simps powr_realpow[symmetric] powr_add[symmetric] powr_mult_base) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1658 | also have "\<dots> = \<lfloor>(2 * (a * 2 powr p) + 2 * b * 2 powr p) * 2 powr (q - p - 1)\<rfloor>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1659 | by (simp add: algebra_simps) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1660 | also have "\<dots> = \<lfloor>(2 * ai + b * 2 powr (p + 1)) / 2 powr (1 - q + p)\<rfloor>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1661 | using assms by (simp add: algebra_simps powr_mult_base divide_powr_uminus) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1662 | also have "\<dots> = \<lfloor>(2 * ai + b * 2 powr (p + 1)) / real_of_int ((2::int) ^ nat (p - q + 1))\<rfloor>" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1663 | using assms by (simp add: algebra_simps powr_realpow[symmetric]) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1664 | also have "\<dots> = \<lfloor>(2 * ai - 1) / real_of_int ((2::int) ^ nat (p - q + 1))\<rfloor>" | 
| 60500 | 1665 | using \<open>b < 0\<close> assms | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1666 | by (simp add: floor_divide_of_int_eq floor_eq floor_divide_real_eq_div | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1667 | del: of_int_mult of_int_power of_int_diff) | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1668 | also have "\<dots> = \<lfloor>(2 * ai - 1) * 2 powr (q - p - 1)\<rfloor>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1669 | using assms by (simp add: algebra_simps divide_powr_uminus powr_realpow[symmetric]) | 
| 60698 | 1670 | finally show ?thesis | 
| 1671 | using \<open>b < 0\<close> by simp | |
| 1672 | qed | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1673 | qed | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1674 | |
| 60698 | 1675 | lemma log2_abs_int_add_less_half_sgn_eq: | 
| 1676 | fixes ai :: int | |
| 1677 | and b :: real | |
| 61945 | 1678 | assumes "\<bar>b\<bar> \<le> 1/2" | 
| 60698 | 1679 | and "ai \<noteq> 0" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1680 | shows "\<lfloor>log 2 \<bar>real_of_int ai + b\<bar>\<rfloor> = \<lfloor>log 2 \<bar>ai + sgn b / 2\<bar>\<rfloor>" | 
| 60698 | 1681 | proof (cases "b = 0") | 
| 1682 | case True | |
| 1683 | then show ?thesis by simp | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1684 | next | 
| 60698 | 1685 | case False | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1686 | def k \<equiv> "\<lfloor>log 2 \<bar>ai\<bar>\<rfloor>" | 
| 60698 | 1687 | then have "\<lfloor>log 2 \<bar>ai\<bar>\<rfloor> = k" | 
| 1688 | by simp | |
| 1689 | then have k: "2 powr k \<le> \<bar>ai\<bar>" "\<bar>ai\<bar> < 2 powr (k + 1)" | |
| 60500 | 1690 | by (simp_all add: floor_log_eq_powr_iff \<open>ai \<noteq> 0\<close>) | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1691 | have "k \<ge> 0" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1692 | using assms by (auto simp: k_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1693 | def r \<equiv> "\<bar>ai\<bar> - 2 ^ nat k" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1694 | have r: "0 \<le> r" "r < 2 powr k" | 
| 60500 | 1695 | using \<open>k \<ge> 0\<close> k | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1696 | by (auto simp: r_def k_def algebra_simps powr_add abs_if powr_int) | 
| 60698 | 1697 | then have "r \<le> (2::int) ^ nat k - 1" | 
| 60500 | 1698 | using \<open>k \<ge> 0\<close> by (auto simp: powr_int) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1699 | from this[simplified of_int_le_iff[symmetric]] \<open>0 \<le> k\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1700 | have r_le: "r \<le> 2 powr k - 1" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1701 | apply (auto simp: algebra_simps powr_int) | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1702 | by (metis of_int_1 of_int_add real_of_int_le_numeral_power_cancel_iff) | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1703 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1704 | have "\<bar>ai\<bar> = 2 powr k + r" | 
| 60500 | 1705 | using \<open>k \<ge> 0\<close> by (auto simp: k_def r_def powr_realpow[symmetric]) | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1706 | |
| 61945 | 1707 | have pos: "\<bar>b\<bar> < 1 \<Longrightarrow> 0 < 2 powr k + (r + b)" for b :: real | 
| 60500 | 1708 | using \<open>0 \<le> k\<close> \<open>ai \<noteq> 0\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1709 | by (auto simp add: r_def powr_realpow[symmetric] abs_if sgn_if algebra_simps | 
| 62390 | 1710 | split: if_split_asm) | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1711 | have less: "\<bar>sgn ai * b\<bar> < 1" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1712 | and less': "\<bar>sgn (sgn ai * b) / 2\<bar> < 1" | 
| 62390 | 1713 | using \<open>\<bar>b\<bar> \<le> _\<close> by (auto simp: abs_if sgn_if split: if_split_asm) | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1714 | |
| 61945 | 1715 | have floor_eq: "\<And>b::real. \<bar>b\<bar> \<le> 1 / 2 \<Longrightarrow> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1716 | \<lfloor>log 2 (1 + (r + b) / 2 powr k)\<rfloor> = (if r = 0 \<and> b < 0 then -1 else 0)" | 
| 60500 | 1717 | using \<open>k \<ge> 0\<close> r r_le | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1718 | by (auto simp: floor_log_eq_powr_iff powr_minus_divide field_simps sgn_if) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1719 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1720 | from \<open>real_of_int \<bar>ai\<bar> = _\<close> have "\<bar>ai + b\<bar> = 2 powr k + (r + sgn ai * b)" | 
| 61945 | 1721 | using \<open>\<bar>b\<bar> <= _\<close> \<open>0 \<le> k\<close> r | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1722 | by (auto simp add: sgn_if abs_if) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1723 | also have "\<lfloor>log 2 \<dots>\<rfloor> = \<lfloor>log 2 (2 powr k + r + sgn (sgn ai * b) / 2)\<rfloor>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1724 | proof - | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1725 | have "2 powr k + (r + (sgn ai) * b) = 2 powr k * (1 + (r + sgn ai * b) / 2 powr k)" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1726 | by (simp add: field_simps) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1727 | also have "\<lfloor>log 2 \<dots>\<rfloor> = k + \<lfloor>log 2 (1 + (r + sgn ai * b) / 2 powr k)\<rfloor>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1728 | using pos[OF less] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1729 | by (subst log_mult) (simp_all add: log_mult powr_mult field_simps) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1730 | also | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1731 | let ?if = "if r = 0 \<and> sgn ai * b < 0 then -1 else 0" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1732 | have "\<lfloor>log 2 (1 + (r + sgn ai * b) / 2 powr k)\<rfloor> = ?if" | 
| 61945 | 1733 | using \<open>\<bar>b\<bar> <= _\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1734 | by (intro floor_eq) (auto simp: abs_mult sgn_if) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1735 | also | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1736 | have "\<dots> = \<lfloor>log 2 (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k)\<rfloor>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1737 | by (subst floor_eq) (auto simp: sgn_if) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1738 | also have "k + \<dots> = \<lfloor>log 2 (2 powr k * (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k))\<rfloor>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1739 | unfolding floor_add2[symmetric] | 
| 61945 | 1740 | using pos[OF less'] \<open>\<bar>b\<bar> \<le> _\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1741 | by (simp add: field_simps add_log_eq_powr) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1742 | also have "2 powr k * (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k) = | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1743 | 2 powr k + r + sgn (sgn ai * b) / 2" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1744 | by (simp add: sgn_if field_simps) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1745 | finally show ?thesis . | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1746 | qed | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1747 | also have "2 powr k + r + sgn (sgn ai * b) / 2 = \<bar>ai + sgn b / 2\<bar>" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1748 | unfolding \<open>real_of_int \<bar>ai\<bar> = _\<close>[symmetric] using \<open>ai \<noteq> 0\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1749 | by (auto simp: abs_if sgn_if algebra_simps) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1750 | finally show ?thesis . | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1751 | qed | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1752 | |
| 60698 | 1753 | context | 
| 1754 | begin | |
| 1755 | ||
| 1756 | qualified lemma compute_far_float_plus_down: | |
| 1757 | fixes m1 e1 m2 e2 :: int | |
| 1758 | and p :: nat | |
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1759 | defines "k1 \<equiv> Suc p - nat (bitlen \<bar>m1\<bar>)" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1760 | assumes H: "bitlen \<bar>m2\<bar> \<le> e1 - e2 - k1 - 2" "m1 \<noteq> 0" "m2 \<noteq> 0" "e1 \<ge> e2" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1761 | shows "float_plus_down p (Float m1 e1) (Float m2 e2) = | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1762 | float_round_down p (Float (m1 * 2 ^ (Suc (Suc k1)) + sgn m2) (e1 - int k1 - 2))" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1763 | proof - | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1764 | let ?a = "real_of_float (Float m1 e1)" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1765 | let ?b = "real_of_float (Float m2 e2)" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1766 | let ?sum = "?a + ?b" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1767 | let ?shift = "real_of_int e2 - real_of_int e1 + real k1 + 1" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1768 | let ?m1 = "m1 * 2 ^ Suc k1" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1769 | let ?m2 = "m2 * 2 powr ?shift" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1770 | let ?m2' = "sgn m2 / 2" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1771 | let ?e = "e1 - int k1 - 1" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1772 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1773 | have sum_eq: "?sum = (?m1 + ?m2) * 2 powr ?e" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1774 | by (auto simp: powr_add[symmetric] powr_mult[symmetric] algebra_simps | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1775 | powr_realpow[symmetric] powr_mult_base) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1776 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1777 | have "\<bar>?m2\<bar> * 2 < 2 powr (bitlen \<bar>m2\<bar> + ?shift + 1)" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1778 | by (auto simp: field_simps powr_add powr_mult_base powr_numeral powr_divide2[symmetric] abs_mult) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1779 | also have "\<dots> \<le> 2 powr 0" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1780 | using H by (intro powr_mono) auto | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1781 | finally have abs_m2_less_half: "\<bar>?m2\<bar> < 1 / 2" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1782 | by simp | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1783 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1784 | then have "\<bar>real_of_int m2\<bar> < 2 powr -(?shift + 1)" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1785 | unfolding powr_minus_divide by (auto simp: bitlen_def field_simps powr_mult_base abs_mult) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1786 | also have "\<dots> \<le> 2 powr real_of_int (e1 - e2 - 2)" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1787 | by simp | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1788 | finally have b_less_quarter: "\<bar>?b\<bar> < 1/4 * 2 powr real_of_int e1" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1789 | by (simp add: powr_add field_simps powr_divide2[symmetric] powr_numeral abs_mult) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1790 | also have "1/4 < \<bar>real_of_int m1\<bar> / 2" using \<open>m1 \<noteq> 0\<close> by simp | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1791 | finally have b_less_half_a: "\<bar>?b\<bar> < 1/2 * \<bar>?a\<bar>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1792 | by (simp add: algebra_simps powr_mult_base abs_mult) | 
| 60698 | 1793 | then have a_half_less_sum: "\<bar>?a\<bar> / 2 < \<bar>?sum\<bar>" | 
| 62390 | 1794 | by (auto simp: field_simps abs_if split: if_split_asm) | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1795 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1796 | from b_less_half_a have "\<bar>?b\<bar> < \<bar>?a\<bar>" "\<bar>?b\<bar> \<le> \<bar>?a\<bar>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1797 | by simp_all | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1798 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1799 | have "\<bar>real_of_float (Float m1 e1)\<bar> \<ge> 1/4 * 2 powr real_of_int e1" | 
| 60500 | 1800 | using \<open>m1 \<noteq> 0\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1801 | by (auto simp: powr_add powr_int bitlen_nonneg divide_right_mono abs_mult) | 
| 60698 | 1802 | then have "?sum \<noteq> 0" using b_less_quarter | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1803 | by (rule sum_neq_zeroI) | 
| 60698 | 1804 | then have "?m1 + ?m2 \<noteq> 0" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1805 | unfolding sum_eq by (simp add: abs_mult zero_less_mult_iff) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1806 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1807 | have "\<bar>real_of_int ?m1\<bar> \<ge> 2 ^ Suc k1" "\<bar>?m2'\<bar> < 2 ^ Suc k1" | 
| 60500 | 1808 | using \<open>m1 \<noteq> 0\<close> \<open>m2 \<noteq> 0\<close> by (auto simp: sgn_if less_1_mult abs_mult simp del: power.simps) | 
| 60698 | 1809 | then have sum'_nz: "?m1 + ?m2' \<noteq> 0" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1810 | by (intro sum_neq_zeroI) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1811 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1812 | have "\<lfloor>log 2 \<bar>real_of_float (Float m1 e1) + real_of_float (Float m2 e2)\<bar>\<rfloor> = \<lfloor>log 2 \<bar>?m1 + ?m2\<bar>\<rfloor> + ?e" | 
| 60500 | 1813 | using \<open>?m1 + ?m2 \<noteq> 0\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1814 | unfolding floor_add[symmetric] sum_eq | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1815 | by (simp add: abs_mult log_mult) linarith | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1816 | also have "\<lfloor>log 2 \<bar>?m1 + ?m2\<bar>\<rfloor> = \<lfloor>log 2 \<bar>?m1 + sgn (real_of_int m2 * 2 powr ?shift) / 2\<bar>\<rfloor>" | 
| 60500 | 1817 | using abs_m2_less_half \<open>m1 \<noteq> 0\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1818 | by (intro log2_abs_int_add_less_half_sgn_eq) (auto simp: abs_mult) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1819 | also have "sgn (real_of_int m2 * 2 powr ?shift) = sgn m2" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1820 | by (auto simp: sgn_if zero_less_mult_iff less_not_sym) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1821 | also | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1822 | have "\<bar>?m1 + ?m2'\<bar> * 2 powr ?e = \<bar>?m1 * 2 + sgn m2\<bar> * 2 powr (?e - 1)" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1823 | by (auto simp: field_simps powr_minus[symmetric] powr_divide2[symmetric] powr_mult_base) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1824 | then have "\<lfloor>log 2 \<bar>?m1 + ?m2'\<bar>\<rfloor> + ?e = \<lfloor>log 2 \<bar>real_of_float (Float (?m1 * 2 + sgn m2) (?e - 1))\<bar>\<rfloor>" | 
| 60500 | 1825 | using \<open>?m1 + ?m2' \<noteq> 0\<close> | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1826 | unfolding floor_add_of_int[symmetric] | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1827 | by (simp add: log_add_eq_powr abs_mult_pos) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1828 | finally | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1829 | have "\<lfloor>log 2 \<bar>?sum\<bar>\<rfloor> = \<lfloor>log 2 \<bar>real_of_float (Float (?m1*2 + sgn m2) (?e - 1))\<bar>\<rfloor>" . | 
| 60698 | 1830 | then have "plus_down p (Float m1 e1) (Float m2 e2) = | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1831 | truncate_down p (Float (?m1*2 + sgn m2) (?e - 1))" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1832 | unfolding plus_down_def | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1833 | proof (rule truncate_down_log2_eqI) | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1834 | let ?f = "(int p - \<lfloor>log 2 \<bar>real_of_float (Float m1 e1) + real_of_float (Float m2 e2)\<bar>\<rfloor>)" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1835 | let ?ai = "m1 * 2 ^ (Suc k1)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1836 | have "\<lfloor>(?a + ?b) * 2 powr real_of_int ?f\<rfloor> = \<lfloor>(real_of_int (2 * ?ai) + sgn ?b) * 2 powr real_of_int (?f - - ?e - 1)\<rfloor>" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1837 | proof (rule floor_sum_times_2_powr_sgn_eq) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1838 | show "?a * 2 powr real_of_int (-?e) = real_of_int ?ai" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1839 | by (simp add: powr_add powr_realpow[symmetric] powr_divide2[symmetric]) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1840 | show "\<bar>?b * 2 powr real_of_int (-?e + 1)\<bar> \<le> 1" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1841 | using abs_m2_less_half | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1842 | by (simp add: abs_mult powr_add[symmetric] algebra_simps powr_mult_base) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1843 | next | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1844 | have "e1 + \<lfloor>log 2 \<bar>real_of_int m1\<bar>\<rfloor> - 1 = \<lfloor>log 2 \<bar>?a\<bar>\<rfloor> - 1" | 
| 60500 | 1845 | using \<open>m1 \<noteq> 0\<close> | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1846 | by (simp add: floor_add2[symmetric] algebra_simps log_mult abs_mult del: floor_add2) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1847 | also have "\<dots> \<le> \<lfloor>log 2 \<bar>?a + ?b\<bar>\<rfloor>" | 
| 60500 | 1848 | using a_half_less_sum \<open>m1 \<noteq> 0\<close> \<open>?sum \<noteq> 0\<close> | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1849 | unfolding floor_diff_of_int[symmetric] | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1850 | by (auto simp add: log_minus_eq_powr powr_minus_divide intro!: floor_mono) | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1851 | finally | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1852 | have "int p - \<lfloor>log 2 \<bar>?a + ?b\<bar>\<rfloor> \<le> p - (bitlen \<bar>m1\<bar>) - e1 + 2" | 
| 60500 | 1853 | by (auto simp: algebra_simps bitlen_def \<open>m1 \<noteq> 0\<close>) | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1854 | also have "\<dots> \<le> - ?e" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1855 | using bitlen_nonneg[of "\<bar>m1\<bar>"] by (simp add: k1_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1856 | finally show "?f \<le> - ?e" by simp | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1857 | qed | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1858 | also have "sgn ?b = sgn m2" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1859 | using powr_gt_zero[of 2 e2] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1860 | by (auto simp add: sgn_if zero_less_mult_iff simp del: powr_gt_zero) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1861 | also have "\<lfloor>(real_of_int (2 * ?m1) + real_of_int (sgn m2)) * 2 powr real_of_int (?f - - ?e - 1)\<rfloor> = | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1862 | \<lfloor>Float (?m1 * 2 + sgn m2) (?e - 1) * 2 powr ?f\<rfloor>" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1863 | by (simp add: powr_add[symmetric] algebra_simps powr_realpow[symmetric]) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1864 | finally | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1865 | show "\<lfloor>(?a + ?b) * 2 powr ?f\<rfloor> = \<lfloor>real_of_float (Float (?m1 * 2 + sgn m2) (?e - 1)) * 2 powr ?f\<rfloor>" . | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1866 | qed | 
| 60698 | 1867 | then show ?thesis | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1868 | by transfer (simp add: plus_down_def ac_simps Let_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1869 | qed | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1870 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1871 | lemma compute_float_plus_down_naive[code]: "float_plus_down p x y = float_round_down p (x + y)" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1872 | by transfer (auto simp: plus_down_def) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1873 | |
| 60698 | 1874 | qualified lemma compute_float_plus_down[code]: | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1875 | fixes p::nat and m1 e1 m2 e2::int | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1876 | shows "float_plus_down p (Float m1 e1) (Float m2 e2) = | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1877 | (if m1 = 0 then float_round_down p (Float m2 e2) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1878 | else if m2 = 0 then float_round_down p (Float m1 e1) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1879 | else (if e1 \<ge> e2 then | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1880 | (let | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1881 | k1 = Suc p - nat (bitlen \<bar>m1\<bar>) | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1882 | in | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1883 | if bitlen \<bar>m2\<bar> > e1 - e2 - k1 - 2 then float_round_down p ((Float m1 e1) + (Float m2 e2)) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1884 | else float_round_down p (Float (m1 * 2 ^ (Suc (Suc k1)) + sgn m2) (e1 - int k1 - 2))) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1885 | else float_plus_down p (Float m2 e2) (Float m1 e1)))" | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1886 | proof - | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1887 |   {
 | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1888 | assume "bitlen \<bar>m2\<bar> \<le> e1 - e2 - (Suc p - nat (bitlen \<bar>m1\<bar>)) - 2" "m1 \<noteq> 0" "m2 \<noteq> 0" "e1 \<ge> e2" | 
| 60698 | 1889 | note compute_far_float_plus_down[OF this] | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1890 | } | 
| 60698 | 1891 | then show ?thesis | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1892 | by transfer (simp add: Let_def plus_down_def ac_simps) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1893 | qed | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1894 | |
| 60698 | 1895 | qualified lemma compute_float_plus_up[code]: "float_plus_up p x y = - float_plus_down p (-x) (-y)" | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1896 | using truncate_down_uminus_eq[of p "x + y"] | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1897 | by transfer (simp add: plus_down_def plus_up_def ac_simps) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1898 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1899 | lemma mantissa_zero[simp]: "mantissa 0 = 0" | 
| 60698 | 1900 | by (metis mantissa_0 zero_float.abs_eq) | 
| 1901 | ||
| 62421 
28d2c75dd180
finite precision computation to determine sign for comparison
 immler parents: 
62420diff
changeset | 1902 | qualified lemma compute_float_less[code]: "a < b \<longleftrightarrow> is_float_pos (float_plus_down 0 b (- a))" | 
| 
28d2c75dd180
finite precision computation to determine sign for comparison
 immler parents: 
62420diff
changeset | 1903 | using truncate_down[of 0 "b - a"] truncate_down_pos[of "b - a" 0] | 
| 
28d2c75dd180
finite precision computation to determine sign for comparison
 immler parents: 
62420diff
changeset | 1904 | by transfer (auto simp: plus_down_def) | 
| 
28d2c75dd180
finite precision computation to determine sign for comparison
 immler parents: 
62420diff
changeset | 1905 | |
| 
28d2c75dd180
finite precision computation to determine sign for comparison
 immler parents: 
62420diff
changeset | 1906 | qualified lemma compute_float_le[code]: "a \<le> b \<longleftrightarrow> is_float_nonneg (float_plus_down 0 b (- a))" | 
| 
28d2c75dd180
finite precision computation to determine sign for comparison
 immler parents: 
62420diff
changeset | 1907 | using truncate_down[of 0 "b - a"] truncate_down_nonneg[of "b - a" 0] | 
| 
28d2c75dd180
finite precision computation to determine sign for comparison
 immler parents: 
62420diff
changeset | 1908 | by transfer (auto simp: plus_down_def) | 
| 
28d2c75dd180
finite precision computation to determine sign for comparison
 immler parents: 
62420diff
changeset | 1909 | |
| 60698 | 1910 | end | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1911 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 1912 | |
| 60500 | 1913 | subsection \<open>Lemmas needed by Approximate\<close> | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1914 | |
| 60698 | 1915 | lemma Float_num[simp]: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1916 | "real_of_float (Float 1 0) = 1" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1917 | "real_of_float (Float 1 1) = 2" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1918 | "real_of_float (Float 1 2) = 4" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1919 | "real_of_float (Float 1 (- 1)) = 1/2" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1920 | "real_of_float (Float 1 (- 2)) = 1/4" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1921 | "real_of_float (Float 1 (- 3)) = 1/8" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1922 | "real_of_float (Float (- 1) 0) = -1" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1923 | "real_of_float (Float (numeral n) 0) = numeral n" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1924 | "real_of_float (Float (- numeral n) 0) = - numeral n" | 
| 60698 | 1925 | using two_powr_int_float[of 2] two_powr_int_float[of "-1"] two_powr_int_float[of "-2"] | 
| 1926 | two_powr_int_float[of "-3"] | |
| 1927 | using powr_realpow[of 2 2] powr_realpow[of 2 3] | |
| 1928 | using powr_minus[of 2 1] powr_minus[of 2 2] powr_minus[of 2 3] | |
| 1929 | by auto | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1930 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1931 | lemma real_of_Float_int[simp]: "real_of_float (Float n 0) = real n" | 
| 60698 | 1932 | by simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1933 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1934 | lemma float_zero[simp]: "real_of_float (Float 0 e) = 0" | 
| 60698 | 1935 | by simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1936 | |
| 61945 | 1937 | lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> \<bar>(a::int) div 2\<bar> < \<bar>a\<bar>" | 
| 60698 | 1938 | by arith | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1939 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1940 | lemma lapprox_rat: "real_of_float (lapprox_rat prec x y) \<le> real_of_int x / real_of_int y" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1941 | by (simp add: lapprox_rat.rep_eq truncate_down) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1942 | |
| 60698 | 1943 | lemma mult_div_le: | 
| 1944 | fixes a b :: int | |
| 1945 | assumes "b > 0" | |
| 1946 | shows "a \<ge> b * (a div b)" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1947 | proof - | 
| 60698 | 1948 | from zmod_zdiv_equality'[of a b] have "a = b * (a div b) + a mod b" | 
| 1949 | by simp | |
| 1950 | also have "\<dots> \<ge> b * (a div b) + 0" | |
| 1951 | apply (rule add_left_mono) | |
| 1952 | apply (rule pos_mod_sign) | |
| 1953 | using assms apply simp | |
| 1954 | done | |
| 1955 | finally show ?thesis | |
| 1956 | by simp | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1957 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1958 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1959 | lemma lapprox_rat_nonneg: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1960 | fixes n x y | 
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 1961 | assumes "0 \<le> x" and "0 \<le> y" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1962 | shows "0 \<le> real_of_float (lapprox_rat n x y)" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1963 | using assms | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1964 | by transfer (simp add: truncate_down_nonneg) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1965 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1966 | lemma rapprox_rat: "real_of_int x / real_of_int y \<le> real_of_float (rapprox_rat prec x y)" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1967 | by transfer (simp add: truncate_up) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1968 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1969 | lemma rapprox_rat_le1: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1970 | fixes n x y | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1971 | assumes xy: "0 \<le> x" "0 < y" "x \<le> y" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1972 | shows "real_of_float (rapprox_rat n x y) \<le> 1" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1973 | using assms | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1974 | by transfer (simp add: truncate_up_le1) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1975 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1976 | lemma rapprox_rat_nonneg_nonpos: "0 \<le> x \<Longrightarrow> y \<le> 0 \<Longrightarrow> real_of_float (rapprox_rat n x y) \<le> 0" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1977 | by transfer (simp add: truncate_up_nonpos divide_nonneg_nonpos) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1978 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1979 | lemma rapprox_rat_nonpos_nonneg: "x \<le> 0 \<Longrightarrow> 0 \<le> y \<Longrightarrow> real_of_float (rapprox_rat n x y) \<le> 0" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1980 | by transfer (simp add: truncate_up_nonpos divide_nonpos_nonneg) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1981 | |
| 54782 | 1982 | lemma real_divl: "real_divl prec x y \<le> x / y" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1983 | by (simp add: real_divl_def truncate_down) | 
| 54782 | 1984 | |
| 1985 | lemma real_divr: "x / y \<le> real_divr prec x y" | |
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1986 | by (simp add: real_divr_def truncate_up) | 
| 54782 | 1987 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1988 | lemma float_divl: "real_of_float (float_divl prec x y) \<le> x / y" | 
| 54782 | 1989 | by transfer (rule real_divl) | 
| 1990 | ||
| 1991 | lemma real_divl_lower_bound: | |
| 1992 | "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> real_divl prec x y" | |
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 1993 | by (simp add: real_divl_def truncate_down_nonneg) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1994 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1995 | lemma float_divl_lower_bound: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 1996 | "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> real_of_float (float_divl prec x y)" | 
| 54782 | 1997 | by transfer (rule real_divl_lower_bound) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1998 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1999 | lemma exponent_1: "exponent 1 = 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2000 | using exponent_float[of 1 0] by (simp add: one_float_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2001 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2002 | lemma mantissa_1: "mantissa 1 = 1" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2003 | using mantissa_float[of 1 0] by (simp add: one_float_def) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 2004 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2005 | lemma bitlen_1: "bitlen 1 = 1" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2006 | by (simp add: bitlen_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2007 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2008 | lemma mantissa_eq_zero_iff: "mantissa x = 0 \<longleftrightarrow> x = 0" | 
| 60698 | 2009 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2010 | proof | 
| 60698 | 2011 | show ?rhs if ?lhs | 
| 2012 | proof - | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2013 | from that have z: "0 = real_of_float x" | 
| 60698 | 2014 | using mantissa_exponent by simp | 
| 2015 | show ?thesis | |
| 2016 | by (simp add: zero_float_def z) | |
| 2017 | qed | |
| 2018 | show ?lhs if ?rhs | |
| 2019 | using that by (simp add: zero_float_def) | |
| 2020 | qed | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 2021 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2022 | lemma float_upper_bound: "x \<le> 2 powr (bitlen \<bar>mantissa x\<bar> + exponent x)" | 
| 60698 | 2023 | proof (cases "x = 0") | 
| 2024 | case True | |
| 2025 | then show ?thesis by simp | |
| 2026 | next | |
| 2027 | case False | |
| 2028 | then have "mantissa x \<noteq> 0" | |
| 2029 | using mantissa_eq_zero_iff by auto | |
| 2030 | have "x = mantissa x * 2 powr (exponent x)" | |
| 2031 | by (rule mantissa_exponent) | |
| 2032 | also have "mantissa x \<le> \<bar>mantissa x\<bar>" | |
| 2033 | by simp | |
| 2034 | also have "\<dots> \<le> 2 powr (bitlen \<bar>mantissa x\<bar>)" | |
| 60500 | 2035 | using bitlen_bounds[of "\<bar>mantissa x\<bar>"] bitlen_nonneg \<open>mantissa x \<noteq> 0\<close> | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61639diff
changeset | 2036 | by (auto simp del: of_int_abs simp add: powr_int) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2037 | finally show ?thesis by (simp add: powr_add) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 2038 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 2039 | |
| 54782 | 2040 | lemma real_divl_pos_less1_bound: | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2041 | assumes "0 < x" "x \<le> 1" | 
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 2042 | shows "1 \<le> real_divl prec 1 x" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2043 | using assms | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2044 | by (auto intro!: truncate_down_ge1 simp: real_divl_def) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 2045 | |
| 54782 | 2046 | lemma float_divl_pos_less1_bound: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2047 | "0 < real_of_float x \<Longrightarrow> real_of_float x \<le> 1 \<Longrightarrow> prec \<ge> 1 \<Longrightarrow> 1 \<le> real_of_float (float_divl prec 1 x)" | 
| 60698 | 2048 | by transfer (rule real_divl_pos_less1_bound) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 2049 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2050 | lemma float_divr: "real_of_float x / real_of_float y \<le> real_of_float (float_divr prec x y)" | 
| 54782 | 2051 | by transfer (rule real_divr) | 
| 2052 | ||
| 60698 | 2053 | lemma real_divr_pos_less1_lower_bound: | 
| 2054 | assumes "0 < x" | |
| 2055 | and "x \<le> 1" | |
| 2056 | shows "1 \<le> real_divr prec 1 x" | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 2057 | proof - | 
| 60698 | 2058 | have "1 \<le> 1 / x" | 
| 2059 | using \<open>0 < x\<close> and \<open>x <= 1\<close> by auto | |
| 2060 | also have "\<dots> \<le> real_divr prec 1 x" | |
| 2061 | using real_divr[where x=1 and y=x] by auto | |
| 47600 | 2062 | finally show ?thesis by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 2063 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 2064 | |
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 2065 | lemma float_divr_pos_less1_lower_bound: "0 < x \<Longrightarrow> x \<le> 1 \<Longrightarrow> 1 \<le> float_divr prec 1 x" | 
| 54782 | 2066 | by transfer (rule real_divr_pos_less1_lower_bound) | 
| 2067 | ||
| 2068 | lemma real_divr_nonpos_pos_upper_bound: | |
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 2069 | "x \<le> 0 \<Longrightarrow> 0 \<le> y \<Longrightarrow> real_divr prec x y \<le> 0" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2070 | by (simp add: real_divr_def truncate_up_nonpos divide_le_0_iff) | 
| 54782 | 2071 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2072 | lemma float_divr_nonpos_pos_upper_bound: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2073 | "real_of_float x \<le> 0 \<Longrightarrow> 0 \<le> real_of_float y \<Longrightarrow> real_of_float (float_divr prec x y) \<le> 0" | 
| 54782 | 2074 | by transfer (rule real_divr_nonpos_pos_upper_bound) | 
| 2075 | ||
| 2076 | lemma real_divr_nonneg_neg_upper_bound: | |
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 2077 | "0 \<le> x \<Longrightarrow> y \<le> 0 \<Longrightarrow> real_divr prec x y \<le> 0" | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2078 | by (simp add: real_divr_def truncate_up_nonpos divide_le_0_iff) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 2079 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2080 | lemma float_divr_nonneg_neg_upper_bound: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2081 | "0 \<le> real_of_float x \<Longrightarrow> real_of_float y \<le> 0 \<Longrightarrow> real_of_float (float_divr prec x y) \<le> 0" | 
| 54782 | 2082 | by transfer (rule real_divr_nonneg_neg_upper_bound) | 
| 2083 | ||
| 54784 | 2084 | lemma truncate_up_nonneg_mono: | 
| 2085 | assumes "0 \<le> x" "x \<le> y" | |
| 2086 | shows "truncate_up prec x \<le> truncate_up prec y" | |
| 2087 | proof - | |
| 60698 | 2088 | consider "\<lfloor>log 2 x\<rfloor> = \<lfloor>log 2 y\<rfloor>" | "\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>" "0 < x" | "x \<le> 0" | 
| 2089 | by arith | |
| 2090 | then show ?thesis | |
| 2091 | proof cases | |
| 2092 | case 1 | |
| 2093 | then show ?thesis | |
| 54784 | 2094 | using assms | 
| 2095 | by (auto simp: truncate_up_def round_up_def intro!: ceiling_mono) | |
| 60698 | 2096 | next | 
| 2097 | case 2 | |
| 2098 | from assms \<open>0 < x\<close> have "log 2 x \<le> log 2 y" | |
| 2099 | by auto | |
| 2100 | with \<open>\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>\<close> | |
| 2101 | have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>" | |
| 2102 | by (metis floor_less_cancel linorder_cases not_le)+ | |
| 54784 | 2103 | have "truncate_up prec x = | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2104 | real_of_int \<lceil>x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor> )\<rceil> * 2 powr - real_of_int (int prec - \<lfloor>log 2 x\<rfloor>)" | 
| 54784 | 2105 | using assms by (simp add: truncate_up_def round_up_def) | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2106 | also have "\<lceil>x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>)\<rceil> \<le> (2 ^ (Suc prec))" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2107 | proof (unfold ceiling_le_iff) | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2108 | have "x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>) \<le> x * (2 powr real (Suc prec) / (2 powr log 2 x))" | 
| 54784 | 2109 | using real_of_int_floor_add_one_ge[of "log 2 x"] assms | 
| 2110 | by (auto simp add: algebra_simps powr_divide2 intro!: mult_left_mono) | |
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2111 | then show "x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>) \<le> real_of_int ((2::int) ^ (Suc prec))" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2112 | using \<open>0 < x\<close> by (simp add: powr_realpow powr_add) | 
| 54784 | 2113 | qed | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2114 | then have "real_of_int \<lceil>x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>)\<rceil> \<le> 2 powr int (Suc prec)" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2115 | apply (auto simp: powr_realpow powr_add) | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2116 | by (metis power_Suc real_of_int_le_numeral_power_cancel_iff) | 
| 54784 | 2117 | also | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2118 | have "2 powr - real_of_int (int prec - \<lfloor>log 2 x\<rfloor>) \<le> 2 powr - real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1)" | 
| 54784 | 2119 | using logless flogless by (auto intro!: floor_mono) | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2120 | also have "2 powr real_of_int (int (Suc prec)) \<le> 2 powr (log 2 y + real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1))" | 
| 60500 | 2121 | using assms \<open>0 < x\<close> | 
| 54784 | 2122 | by (auto simp: algebra_simps) | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2123 | finally have "truncate_up prec x \<le> 2 powr (log 2 y + real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1)) * 2 powr - real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1)" | 
| 54784 | 2124 | by simp | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2125 | also have "\<dots> = 2 powr (log 2 y + real_of_int (int prec - \<lfloor>log 2 y\<rfloor>) - real_of_int (int prec - \<lfloor>log 2 y\<rfloor>))" | 
| 54784 | 2126 | by (subst powr_add[symmetric]) simp | 
| 2127 | also have "\<dots> = y" | |
| 60500 | 2128 | using \<open>0 < x\<close> assms | 
| 54784 | 2129 | by (simp add: powr_add) | 
| 2130 | also have "\<dots> \<le> truncate_up prec y" | |
| 2131 | by (rule truncate_up) | |
| 60698 | 2132 | finally show ?thesis . | 
| 2133 | next | |
| 2134 | case 3 | |
| 2135 | then show ?thesis | |
| 54784 | 2136 | using assms | 
| 2137 | by (auto intro!: truncate_up_le) | |
| 60698 | 2138 | qed | 
| 54784 | 2139 | qed | 
| 2140 | ||
| 2141 | lemma truncate_up_switch_sign_mono: | |
| 2142 | assumes "x \<le> 0" "0 \<le> y" | |
| 2143 | shows "truncate_up prec x \<le> truncate_up prec y" | |
| 2144 | proof - | |
| 60500 | 2145 | note truncate_up_nonpos[OF \<open>x \<le> 0\<close>] | 
| 2146 | also note truncate_up_le[OF \<open>0 \<le> y\<close>] | |
| 54784 | 2147 | finally show ?thesis . | 
| 2148 | qed | |
| 2149 | ||
| 2150 | lemma truncate_down_switch_sign_mono: | |
| 60698 | 2151 | assumes "x \<le> 0" | 
| 2152 | and "0 \<le> y" | |
| 2153 | and "x \<le> y" | |
| 54784 | 2154 | shows "truncate_down prec x \<le> truncate_down prec y" | 
| 2155 | proof - | |
| 60500 | 2156 | note truncate_down_le[OF \<open>x \<le> 0\<close>] | 
| 2157 | also note truncate_down_nonneg[OF \<open>0 \<le> y\<close>] | |
| 54784 | 2158 | finally show ?thesis . | 
| 2159 | qed | |
| 2160 | ||
| 2161 | lemma truncate_down_nonneg_mono: | |
| 2162 | assumes "0 \<le> x" "x \<le> y" | |
| 2163 | shows "truncate_down prec x \<le> truncate_down prec y" | |
| 2164 | proof - | |
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2165 | consider "x \<le> 0" | "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> = \<lfloor>log 2 \<bar>y\<bar>\<rfloor>" | | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2166 | "0 < x" "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>" | 
| 60698 | 2167 | by arith | 
| 2168 | then show ?thesis | |
| 2169 | proof cases | |
| 2170 | case 1 | |
| 54784 | 2171 | with assms have "x = 0" "0 \<le> y" by simp_all | 
| 60698 | 2172 | then show ?thesis | 
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 2173 | by (auto intro!: truncate_down_nonneg) | 
| 60698 | 2174 | next | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2175 | case 2 | 
| 60698 | 2176 | then show ?thesis | 
| 54784 | 2177 | using assms | 
| 2178 | by (auto simp: truncate_down_def round_down_def intro!: floor_mono) | |
| 60698 | 2179 | next | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2180 | case 3 | 
| 60698 | 2181 | from \<open>0 < x\<close> have "log 2 x \<le> log 2 y" "0 < y" "0 \<le> y" | 
| 2182 | using assms by auto | |
| 2183 | with \<open>\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>\<close> | |
| 2184 | have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>" | |
| 60500 | 2185 | unfolding atomize_conj abs_of_pos[OF \<open>0 < x\<close>] abs_of_pos[OF \<open>0 < y\<close>] | 
| 54784 | 2186 | by (metis floor_less_cancel linorder_cases not_le) | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2187 | have "2 powr prec \<le> y * 2 powr real prec / (2 powr log 2 y)" | 
| 60698 | 2188 | using \<open>0 < y\<close> by simp | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2189 | also have "\<dots> \<le> y * 2 powr real (Suc prec) / (2 powr (real_of_int \<lfloor>log 2 y\<rfloor> + 1))" | 
| 60500 | 2190 | using \<open>0 \<le> y\<close> \<open>0 \<le> x\<close> assms(2) | 
| 56544 | 2191 | by (auto intro!: powr_mono divide_left_mono | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2192 | simp: of_nat_diff powr_add | 
| 54784 | 2193 | powr_divide2[symmetric]) | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2194 | also have "\<dots> = y * 2 powr real (Suc prec) / (2 powr real_of_int \<lfloor>log 2 y\<rfloor> * 2)" | 
| 54784 | 2195 | by (auto simp: powr_add) | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2196 | finally have "(2 ^ prec) \<le> \<lfloor>y * 2 powr real_of_int (int (Suc prec) - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)\<rfloor>" | 
| 60500 | 2197 | using \<open>0 \<le> y\<close> | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2198 | by (auto simp: powr_divide2[symmetric] le_floor_iff powr_realpow powr_add) | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2199 | then have "(2 ^ (prec)) * 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor>) \<le> truncate_down prec y" | 
| 54784 | 2200 | by (auto simp: truncate_down_def round_down_def) | 
| 2201 | moreover | |
| 2202 |     {
 | |
| 60500 | 2203 | have "x = 2 powr (log 2 \<bar>x\<bar>)" using \<open>0 < x\<close> by simp | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2204 | also have "\<dots> \<le> (2 ^ (Suc prec )) * 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>)" | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2205 | using real_of_int_floor_add_one_ge[of "log 2 \<bar>x\<bar>"] \<open>0 < x\<close> | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2206 | by (auto simp: powr_realpow[symmetric] powr_add[symmetric] algebra_simps | 
| 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2207 | powr_mult_base le_powr_iff) | 
| 54784 | 2208 | also | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2209 | have "2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>) \<le> 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> + 1)" | 
| 60500 | 2210 | using logless flogless \<open>x > 0\<close> \<open>y > 0\<close> | 
| 54784 | 2211 | by (auto intro!: floor_mono) | 
| 62420 
c7666166c24e
positive precision for truncate; fixed precision for approximation of rationals; code for truncate
 immler parents: 
62419diff
changeset | 2212 | finally have "x \<le> (2 ^ prec) * 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor>)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2213 | by (auto simp: powr_realpow[symmetric] powr_divide2[symmetric] assms of_nat_diff) | 
| 60698 | 2214 | } | 
| 2215 | ultimately show ?thesis | |
| 54784 | 2216 | by (metis dual_order.trans truncate_down) | 
| 60698 | 2217 | qed | 
| 54784 | 2218 | qed | 
| 2219 | ||
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 2220 | lemma truncate_down_eq_truncate_up: "truncate_down p x = - truncate_up p (-x)" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 2221 | and truncate_up_eq_truncate_down: "truncate_up p x = - truncate_down p (-x)" | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 2222 | by (auto simp: truncate_up_uminus_eq truncate_down_uminus_eq) | 
| 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 2223 | |
| 54784 | 2224 | lemma truncate_down_mono: "x \<le> y \<Longrightarrow> truncate_down p x \<le> truncate_down p y" | 
| 2225 | apply (cases "0 \<le> x") | |
| 2226 | apply (rule truncate_down_nonneg_mono, assumption+) | |
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 2227 | apply (simp add: truncate_down_eq_truncate_up) | 
| 54784 | 2228 | apply (cases "0 \<le> y") | 
| 2229 | apply (auto intro: truncate_up_nonneg_mono truncate_up_switch_sign_mono) | |
| 2230 | done | |
| 2231 | ||
| 2232 | lemma truncate_up_mono: "x \<le> y \<Longrightarrow> truncate_up p x \<le> truncate_up p y" | |
| 58982 
27e7e3f9e665
simplified computations based on round_up by reducing to round_down;
 immler parents: 
58881diff
changeset | 2233 | by (simp add: truncate_up_eq_truncate_down truncate_down_mono) | 
| 54784 | 2234 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2235 | lemma Float_le_zero_iff: "Float a b \<le> 0 \<longleftrightarrow> a \<le> 0" | 
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59984diff
changeset | 2236 | by (auto simp: zero_float_def mult_le_0_iff) (simp add: not_less [symmetric]) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2237 | |
| 60698 | 2238 | lemma real_of_float_pprt[simp]: | 
| 2239 | fixes a :: float | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2240 | shows "real_of_float (pprt a) = pprt (real_of_float a)" | 
| 47600 | 2241 | unfolding pprt_def sup_float_def max_def sup_real_def by auto | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2242 | |
| 60698 | 2243 | lemma real_of_float_nprt[simp]: | 
| 2244 | fixes a :: float | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2245 | shows "real_of_float (nprt a) = nprt (real_of_float a)" | 
| 47600 | 2246 | unfolding nprt_def inf_float_def min_def inf_real_def by auto | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2247 | |
| 60698 | 2248 | context | 
| 2249 | begin | |
| 2250 | ||
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 2251 | lift_definition int_floor_fl :: "float \<Rightarrow> int" is floor . | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 2252 | |
| 60698 | 2253 | qualified lemma compute_int_floor_fl[code]: | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 2254 | "int_floor_fl (Float m e) = (if 0 \<le> e then m * 2 ^ nat e else m div (2 ^ (nat (-e))))" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2255 | apply transfer | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2256 | apply (simp add: powr_int floor_divide_of_int_eq) | 
| 61942 | 2257 | apply (metis (no_types, hide_lams)floor_divide_of_int_eq of_int_numeral of_int_power floor_of_int of_int_mult) | 
| 2258 | done | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2259 | |
| 61942 | 2260 | lift_definition floor_fl :: "float \<Rightarrow> float" is "\<lambda>x. real_of_int \<lfloor>x\<rfloor>" | 
| 2261 | by simp | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2262 | |
| 60698 | 2263 | qualified lemma compute_floor_fl[code]: | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 2264 | "floor_fl (Float m e) = (if 0 \<le> e then Float m e else Float (m div (2 ^ (nat (-e)))) 0)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2265 | apply transfer | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2266 | apply (simp add: powr_int floor_divide_of_int_eq) | 
| 61942 | 2267 | apply (metis (no_types, hide_lams)floor_divide_of_int_eq of_int_numeral of_int_power of_int_mult) | 
| 2268 | done | |
| 60698 | 2269 | |
| 2270 | end | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 2271 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2272 | lemma floor_fl: "real_of_float (floor_fl x) \<le> real_of_float x" | 
| 60698 | 2273 | by transfer simp | 
| 47600 | 2274 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2275 | lemma int_floor_fl: "real_of_int (int_floor_fl x) \<le> real_of_float x" | 
| 60698 | 2276 | by transfer simp | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 2277 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 2278 | lemma floor_pos_exp: "exponent (floor_fl x) \<ge> 0" | 
| 53381 | 2279 | proof (cases "floor_fl x = float_of 0") | 
| 2280 | case True | |
| 60698 | 2281 | then show ?thesis | 
| 2282 | by (simp add: floor_fl_def) | |
| 53381 | 2283 | next | 
| 2284 | case False | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2285 | have eq: "floor_fl x = Float \<lfloor>real_of_float x\<rfloor> 0" | 
| 60698 | 2286 | by transfer simp | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60868diff
changeset | 2287 | obtain i where "\<lfloor>real_of_float x\<rfloor> = mantissa (floor_fl x) * 2 ^ i" "0 = exponent (floor_fl x) - int i" | 
| 53381 | 2288 | by (rule denormalize_shift[OF eq[THEN eq_reflection] False]) | 
| 60698 | 2289 | then show ?thesis | 
| 2290 | by simp | |
| 53381 | 2291 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 2292 | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 2293 | lemma compute_mantissa[code]: | 
| 60698 | 2294 | "mantissa (Float m e) = | 
| 2295 | (if m = 0 then 0 else if 2 dvd m then mantissa (normfloat (Float m e)) else m)" | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 2296 | by (auto simp: mantissa_float Float.abs_eq) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 2297 | |
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 2298 | lemma compute_exponent[code]: | 
| 60698 | 2299 | "exponent (Float m e) = | 
| 2300 | (if m = 0 then 0 else if 2 dvd m then exponent (normfloat (Float m e)) else e)" | |
| 58985 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 2301 | by (auto simp: exponent_float Float.abs_eq) | 
| 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 immler parents: 
58982diff
changeset | 2302 | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 2303 | end |