| author | desharna | 
| Tue, 14 Oct 2014 16:17:34 +0200 | |
| changeset 58675 | 69571f0a93df | 
| parent 58410 | 6d46ad54a2ab | 
| child 58834 | 773b378d9313 | 
| 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 | ||
| 29988 | 6 | header {* Floating-Point Numbers *}
 | 
| 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 | |
| 58042 
ffa9e39763e3
introduce real_of typeclass for real :: 'a => real
 hoelzl parents: 
57862diff
changeset | 18 | instantiation float :: real_of | 
| 
ffa9e39763e3
introduce real_of typeclass for real :: 'a => real
 hoelzl parents: 
57862diff
changeset | 19 | begin | 
| 
ffa9e39763e3
introduce real_of typeclass for real :: 'a => real
 hoelzl parents: 
57862diff
changeset | 20 | |
| 
ffa9e39763e3
introduce real_of typeclass for real :: 'a => real
 hoelzl parents: 
57862diff
changeset | 21 | definition real_float :: "float \<Rightarrow> real" where | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 22 | real_of_float_def[code_unfold]: "real \<equiv> real_of_float" | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 23 | |
| 58042 
ffa9e39763e3
introduce real_of typeclass for real :: 'a => real
 hoelzl parents: 
57862diff
changeset | 24 | instance .. | 
| 
ffa9e39763e3
introduce real_of typeclass for real :: 'a => real
 hoelzl parents: 
57862diff
changeset | 25 | end | 
| 
ffa9e39763e3
introduce real_of typeclass for real :: 'a => real
 hoelzl parents: 
57862diff
changeset | 26 | |
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 27 | lemma type_definition_float': "type_definition real float_of float" | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 28 | using type_definition_float unfolding real_of_float_def . | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 29 | |
| 47937 
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
 kuncar parents: 
47780diff
changeset | 30 | setup_lifting (no_code) type_definition_float' | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 31 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 32 | lemmas float_of_inject[simp] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 33 | |
| 47600 | 34 | declare [[coercion "real :: float \<Rightarrow> real"]] | 
| 35 | ||
| 36 | lemma real_of_float_eq: | |
| 37 | fixes f1 f2 :: float shows "f1 = f2 \<longleftrightarrow> real f1 = real f2" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 38 | unfolding real_of_float_def real_of_float_inject .. | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 39 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 40 | lemma float_of_real[simp]: "float_of (real x) = x" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 41 | unfolding real_of_float_def by (rule real_of_float_inverse) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 42 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 43 | lemma real_float[simp]: "x \<in> float \<Longrightarrow> real (float_of x) = x" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 44 | unfolding real_of_float_def by (rule float_of_inverse) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 45 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 46 | subsection {* Real operations preserving the representation as floating point number *}
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 47 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 48 | 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 | 49 | by (auto simp: float_def) | 
| 19765 | 50 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 51 | lemma zero_float[simp]: "0 \<in> float" by (auto simp: float_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 52 | lemma one_float[simp]: "1 \<in> float" by (intro floatI[of 1 0]) simp | 
| 53381 | 53 | lemma numeral_float[simp]: "numeral i \<in> float" by (intro floatI[of "numeral i" 0]) simp | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 54 | lemma neg_numeral_float[simp]: "- numeral i \<in> float" by (intro floatI[of "- numeral i" 0]) simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 55 | lemma real_of_int_float[simp]: "real (x :: int) \<in> float" by (intro floatI[of x 0]) simp | 
| 47600 | 56 | lemma real_of_nat_float[simp]: "real (x :: nat) \<in> float" by (intro floatI[of x 0]) simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 57 | lemma two_powr_int_float[simp]: "2 powr (real (i::int)) \<in> float" by (intro floatI[of 1 i]) simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 58 | lemma two_powr_nat_float[simp]: "2 powr (real (i::nat)) \<in> float" by (intro floatI[of 1 i]) simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 59 | lemma two_powr_minus_int_float[simp]: "2 powr - (real (i::int)) \<in> float" by (intro floatI[of 1 "-i"]) simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 60 | lemma two_powr_minus_nat_float[simp]: "2 powr - (real (i::nat)) \<in> float" by (intro floatI[of 1 "-i"]) simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 61 | lemma two_powr_numeral_float[simp]: "2 powr numeral i \<in> float" by (intro floatI[of 1 "numeral i"]) simp | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 62 | lemma two_powr_neg_numeral_float[simp]: "2 powr - numeral i \<in> float" by (intro floatI[of 1 "- numeral i"]) simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 63 | lemma two_pow_float[simp]: "2 ^ n \<in> float" by (intro floatI[of 1 "n"]) (simp add: powr_realpow) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 64 | lemma real_of_float_float[simp]: "real (f::float) \<in> float" by (cases f) simp | 
| 45495 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 hoelzl parents: 
44766diff
changeset | 65 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 66 | 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 | 67 | unfolding float_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 68 | proof (safe, simp) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 69 | fix e1 m1 e2 m2 :: int | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 70 |   { fix e1 m1 e2 m2 :: int assume "e1 \<le> e2"
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 71 | then have "m1 * 2 powr e1 + m2 * 2 powr e2 = (m1 + m2 * 2 ^ nat (e2 - e1)) * 2 powr e1" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 72 | by (simp add: powr_realpow[symmetric] powr_divide2[symmetric] field_simps) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 73 | then have "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 74 | by blast } | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 75 | note * = this | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 76 | show "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 77 | proof (cases e1 e2 rule: linorder_le_cases) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 78 | assume "e2 \<le> e1" from *[OF this, of m2 m1] show ?thesis by (simp add: ac_simps) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 79 | qed (rule *) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 80 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 81 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 82 | 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 | 83 | apply (auto simp: float_def) | 
| 57492 
74bf65a1910a
Hypsubst preserves equality hypotheses
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
56777diff
changeset | 84 | apply hypsubst_thin | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 85 | apply (rule_tac x="-x" in exI) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 86 | apply (rule_tac x="xa" in exI) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 87 | apply (simp add: field_simps) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 88 | done | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 89 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 90 | 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 | 91 | apply (auto simp: float_def) | 
| 57492 
74bf65a1910a
Hypsubst preserves equality hypotheses
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
56777diff
changeset | 92 | apply hypsubst_thin | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 93 | apply (rule_tac x="x * xa" in exI) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 94 | apply (rule_tac x="xb + xc" in exI) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 95 | apply (simp add: powr_add) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 96 | done | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 97 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 98 | 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 | 99 | using plus_float [of x "- y"] by simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 100 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 101 | lemma abs_float[simp]: "x \<in> float \<Longrightarrow> abs x \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 102 | by (cases x rule: linorder_cases[of 0]) auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 103 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 104 | 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 | 105 | 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 | 106 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 107 | 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 | 108 | apply (auto simp add: float_def) | 
| 57492 
74bf65a1910a
Hypsubst preserves equality hypotheses
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
56777diff
changeset | 109 | apply hypsubst_thin | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 110 | apply (rule_tac x="x" in exI) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 111 | apply (rule_tac x="xa - d" in exI) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 112 | 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 | 113 | done | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 114 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 115 | 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 | 116 | apply (auto simp add: float_def) | 
| 57492 
74bf65a1910a
Hypsubst preserves equality hypotheses
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
56777diff
changeset | 117 | apply hypsubst_thin | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 118 | apply (rule_tac x="x" in exI) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 119 | apply (rule_tac x="xa - d" in exI) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 120 | 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 | 121 | done | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 122 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 123 | lemma div_numeral_Bit0_float[simp]: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 124 | assumes x: "x / numeral n \<in> float" shows "x / (numeral (Num.Bit0 n)) \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 125 | proof - | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 126 | have "(x / numeral n) / 2^1 \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 127 | by (intro x div_power_2_float) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 128 | 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 | 129 | by (induct n) auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 130 | finally show ?thesis . | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 131 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 132 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 133 | lemma div_neg_numeral_Bit0_float[simp]: | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 134 | assumes x: "x / numeral n \<in> float" shows "x / (- numeral (Num.Bit0 n)) \<in> float" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 135 | proof - | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 136 | have "- (x / numeral (Num.Bit0 n)) \<in> float" using x by simp | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 137 | also have "- (x / numeral (Num.Bit0 n)) = x / - numeral (Num.Bit0 n)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 138 | by simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 139 | finally show ?thesis . | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 140 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 141 | |
| 47600 | 142 | lift_definition Float :: "int \<Rightarrow> int \<Rightarrow> float" is "\<lambda>(m::int) (e::int). m * 2 powr e" by simp | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 143 | 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 | 144 | |
| 47780 | 145 | lemma compute_real_of_float[code]: | 
| 146 | "real_of_float (Float m e) = (if e \<ge> 0 then m * 2 ^ nat e else m / 2 ^ (nat (-e)))" | |
| 147 | by (simp add: real_of_float_def[symmetric] powr_int) | |
| 148 | ||
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 149 | code_datatype Float | 
| 47600 | 150 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 151 | subsection {* Arithmetic operations on floating point numbers *}
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 152 | |
| 47600 | 153 | 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 | 154 | begin | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 155 | |
| 47600 | 156 | 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 | 157 | declare zero_float.rep_eq[simp] | 
| 47600 | 158 | 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 | 159 | declare one_float.rep_eq[simp] | 
| 47600 | 160 | 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 | 161 | declare plus_float.rep_eq[simp] | 
| 47600 | 162 | 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 | 163 | declare times_float.rep_eq[simp] | 
| 47600 | 164 | 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 | 165 | declare minus_float.rep_eq[simp] | 
| 47600 | 166 | 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 | 167 | declare uminus_float.rep_eq[simp] | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 168 | |
| 47600 | 169 | 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 | 170 | declare abs_float.rep_eq[simp] | 
| 47600 | 171 | 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 | 172 | declare sgn_float.rep_eq[simp] | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 173 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 174 | 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 | 175 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 176 | 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 | 177 | declare less_eq_float.rep_eq[simp] | 
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 178 | 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 | 179 | declare less_float.rep_eq[simp] | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 180 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 181 | instance | 
| 47600 | 182 | proof qed (transfer, fastforce simp add: field_simps intro: mult_left_mono mult_right_mono)+ | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 183 | end | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 184 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 185 | lemma real_of_float_power[simp]: fixes f::float shows "real (f^n) = real f^n" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 186 | by (induct n) simp_all | 
| 45495 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 hoelzl parents: 
44766diff
changeset | 187 | |
| 53381 | 188 | lemma fixes x y::float | 
| 47600 | 189 | shows real_of_float_min: "real (min x y) = min (real x) (real y)" | 
| 190 | and real_of_float_max: "real (max x y) = max (real x) (real y)" | |
| 191 | by (simp_all add: min_def max_def) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 192 | |
| 53215 
5e47c31c6f7c
renamed typeclass dense_linorder to unbounded_dense_linorder
 hoelzl parents: 
51542diff
changeset | 193 | instance float :: unbounded_dense_linorder | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 194 | proof | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 195 | fix a b :: float | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 196 | show "\<exists>c. a < c" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 197 | apply (intro exI[of _ "a + 1"]) | 
| 47600 | 198 | apply transfer | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 199 | apply simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 200 | done | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 201 | show "\<exists>c. c < a" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 202 | apply (intro exI[of _ "a - 1"]) | 
| 47600 | 203 | apply transfer | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 204 | apply simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 205 | done | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 206 | assume "a < b" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 207 | then show "\<exists>c. a < c \<and> c < b" | 
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
58042diff
changeset | 208 | apply (intro exI[of _ "(a + b) * Float 1 (- 1)"]) | 
| 47600 | 209 | apply transfer | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 210 | apply (simp add: powr_minus) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 211 | done | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 212 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 213 | |
| 47600 | 214 | instantiation float :: lattice_ab_group_add | 
| 46573 | 215 | begin | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 216 | |
| 47600 | 217 | definition inf_float::"float\<Rightarrow>float\<Rightarrow>float" | 
| 218 | 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 | 219 | |
| 47600 | 220 | definition sup_float::"float\<Rightarrow>float\<Rightarrow>float" | 
| 221 | 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 | 222 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 223 | instance | 
| 47600 | 224 | by default | 
| 225 | (transfer, simp_all add: inf_float_def sup_float_def real_of_float_min real_of_float_max)+ | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 226 | end | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 227 | |
| 47600 | 228 | lemma float_numeral[simp]: "real (numeral x :: float) = numeral x" | 
| 229 | apply (induct x) | |
| 230 | apply simp | |
| 231 | apply (simp_all only: numeral_Bit0 numeral_Bit1 real_of_float_eq real_float | |
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 232 | plus_float.rep_eq one_float.rep_eq plus_float numeral_float one_float) | 
| 47600 | 233 | done | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 234 | |
| 53381 | 235 | lemma transfer_numeral [transfer_rule]: | 
| 55945 | 236 | "rel_fun (op =) pcr_float (numeral :: _ \<Rightarrow> real) (numeral :: _ \<Rightarrow> float)" | 
| 237 | unfolding rel_fun_def float.pcr_cr_eq cr_float_def by simp | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 238 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 239 | lemma float_neg_numeral[simp]: "real (- numeral x :: float) = - numeral x" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 240 | by simp | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46670diff
changeset | 241 | |
| 53381 | 242 | lemma transfer_neg_numeral [transfer_rule]: | 
| 55945 | 243 | "rel_fun (op =) pcr_float (- numeral :: _ \<Rightarrow> real) (- numeral :: _ \<Rightarrow> float)" | 
| 244 | unfolding rel_fun_def float.pcr_cr_eq cr_float_def by simp | |
| 47600 | 245 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 246 | lemma | 
| 47600 | 247 | shows float_of_numeral[simp]: "numeral k = float_of (numeral k)" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 248 | and float_of_neg_numeral[simp]: "- numeral k = float_of (- numeral k)" | 
| 47600 | 249 | unfolding real_of_float_eq by simp_all | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46670diff
changeset | 250 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 251 | subsection {* Represent floats as unique mantissa and exponent *}
 | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46670diff
changeset | 252 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 253 | lemma int_induct_abs[case_names less]: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 254 | fixes j :: int | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 255 | 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 | 256 | shows "P j" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 257 | proof (induct "nat \<bar>j\<bar>" arbitrary: j rule: less_induct) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 258 | case less show ?case by (rule H[OF less]) simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 259 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 260 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 261 | lemma int_cancel_factors: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 262 | fixes n :: int assumes "1 < r" shows "n = 0 \<or> (\<exists>k i. n = k * r ^ i \<and> \<not> r dvd k)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 263 | proof (induct n rule: int_induct_abs) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 264 | case (less n) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 265 |   { fix m assume n: "n \<noteq> 0" "n = m * r"
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 266 | then have "\<bar>m \<bar> < \<bar>n\<bar>" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 267 | by (metis abs_dvd_iff abs_ge_self assms comm_semiring_1_class.normalizing_semiring_rules(7) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 268 | dvd_imp_le_int dvd_refl dvd_triv_right linorder_neq_iff linorder_not_le | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 269 | mult_eq_0_iff zdvd_mult_cancel1) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 270 | from less[OF this] n have "\<exists>k i. n = k * r ^ Suc i \<and> \<not> r dvd k" by auto } | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 271 | then show ?case | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 272 | by (metis comm_semiring_1_class.normalizing_semiring_rules(12,7) dvdE power_0) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 273 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 274 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 275 | lemma mult_powr_eq_mult_powr_iff_asym: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 276 | fixes m1 m2 e1 e2 :: int | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 277 | assumes m1: "\<not> 2 dvd m1" and "e1 \<le> e2" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 278 | shows "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 | 279 | proof | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 280 | have "m1 \<noteq> 0" using m1 unfolding dvd_def by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 281 | assume eq: "m1 * 2 powr e1 = m2 * 2 powr e2" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 282 | with `e1 \<le> e2` have "m1 = m2 * 2 powr nat (e2 - e1)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 283 | by (simp add: powr_divide2[symmetric] field_simps) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 284 | also have "\<dots> = m2 * 2^nat (e2 - e1)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 285 | by (simp add: powr_realpow) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 286 | finally have m1_eq: "m1 = m2 * 2^nat (e2 - e1)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 287 | unfolding real_of_int_inject . | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 288 | with m1 have "m1 = m2" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 289 | by (cases "nat (e2 - e1)") (auto simp add: dvd_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 290 | then show "m1 = m2 \<and> e1 = e2" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 291 | using eq `m1 \<noteq> 0` by (simp add: powr_inj) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 292 | qed simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 293 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 294 | lemma mult_powr_eq_mult_powr_iff: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 295 | fixes m1 m2 e1 e2 :: int | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 296 | 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 | 297 | 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 | 298 | 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 | 299 | by (cases e1 e2 rule: linorder_le_cases) auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 300 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 301 | lemma floatE_normed: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 302 | assumes x: "x \<in> float" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 303 | obtains (zero) "x = 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 304 | | (powr) m e :: int where "x = m * 2 powr e" "\<not> 2 dvd m" "x \<noteq> 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 305 | proof atomize_elim | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 306 |   { assume "x \<noteq> 0"
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 307 | from x obtain m e :: int where x: "x = m * 2 powr e" by (auto simp: float_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 308 | with `x \<noteq> 0` int_cancel_factors[of 2 m] obtain k i where "m = k * 2 ^ i" "\<not> 2 dvd k" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 309 | by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 310 | with `\<not> 2 dvd k` x have "\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 311 | by (rule_tac exI[of _ "k"], rule_tac exI[of _ "e + int i"]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 312 | (simp add: powr_add powr_realpow) } | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 313 | then show "x = 0 \<or> (\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m \<and> x \<noteq> 0)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 314 | by blast | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 315 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 316 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 317 | lemma float_normed_cases: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 318 | fixes f :: float | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 319 | obtains (zero) "f = 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 320 | | (powr) m e :: int where "real f = m * 2 powr e" "\<not> 2 dvd m" "f \<noteq> 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 321 | proof (atomize_elim, induct f) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 322 | case (float_of y) then show ?case | 
| 47600 | 323 | 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 | 324 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 325 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 326 | definition mantissa :: "float \<Rightarrow> int" where | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 327 | "mantissa f = fst (SOME p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 328 | \<or> (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p))" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 329 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 330 | definition exponent :: "float \<Rightarrow> int" where | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 331 | "exponent f = snd (SOME p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 332 | \<or> (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p))" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 333 | |
| 53381 | 334 | lemma | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 335 | 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 | 336 | 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 | 337 | proof - | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 338 | have "\<And>p::int \<times> int. fst p = 0 \<and> snd p = 0 \<longleftrightarrow> p = (0, 0)" by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 339 | then show ?E ?M | 
| 47600 | 340 | 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 | 341 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 342 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 343 | lemma | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 344 | shows mantissa_exponent: "real f = mantissa f * 2 powr exponent f" (is ?E) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 345 | 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 | 346 | proof cases | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 347 | assume [simp]: "f \<noteq> (float_of 0)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 348 | 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 | 349 | proof (cases f rule: float_normed_cases) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 350 | case (powr m e) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 351 | then have "\<exists>p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 352 | \<or> (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 353 | by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 354 | then show ?thesis | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 355 | unfolding exponent_def mantissa_def | 
| 47600 | 356 | by (rule someI2_ex) (simp add: zero_float_def) | 
| 357 | qed (simp add: zero_float_def) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 358 | then show ?E ?D by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 359 | qed simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 360 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 361 | 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 | 362 | using mantissa_not_dvd[of f] by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 363 | |
| 53381 | 364 | lemma | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 365 | fixes m e :: int | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 366 | defines "f \<equiv> float_of (m * 2 powr e)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 367 | assumes dvd: "\<not> 2 dvd m" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 368 | shows mantissa_float: "mantissa f = m" (is "?M") | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 369 | 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 | 370 | proof cases | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 371 | assume "m = 0" with dvd show "mantissa f = m" by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 372 | next | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 373 | assume "m \<noteq> 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 374 | then have f_not_0: "f \<noteq> float_of 0" by (simp add: f_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 375 | from mantissa_exponent[of f] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 376 | 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 | 377 | by (auto simp add: f_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 378 | then show "?M" "?E" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 379 | using mantissa_not_dvd[OF f_not_0] dvd | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 380 | by (auto simp: mult_powr_eq_mult_powr_iff) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 381 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 382 | |
| 47600 | 383 | subsection {* Compute arithmetic operations *}
 | 
| 384 | ||
| 385 | lemma Float_mantissa_exponent: "Float (mantissa f) (exponent f) = f" | |
| 386 | unfolding real_of_float_eq mantissa_exponent[of f] by simp | |
| 387 | ||
| 388 | lemma Float_cases[case_names Float, cases type: float]: | |
| 389 | fixes f :: float | |
| 390 | obtains (Float) m e :: int where "f = Float m e" | |
| 391 | using Float_mantissa_exponent[symmetric] | |
| 392 | by (atomize_elim) auto | |
| 393 | ||
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 394 | lemma denormalize_shift: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 395 | assumes f_def: "f \<equiv> Float m e" and not_0: "f \<noteq> float_of 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 396 | 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 | 397 | proof | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 398 | from mantissa_exponent[of f] f_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 399 | 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 | 400 | by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 401 | 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 | 402 | by (simp add: powr_divide2[symmetric] field_simps) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 403 | moreover | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 404 | have "e \<le> exponent f" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 405 | proof (rule ccontr) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 406 | assume "\<not> e \<le> exponent f" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 407 | then have pos: "exponent f < e" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 408 | then have "2 powr (exponent f - e) = 2 powr - real (e - exponent f)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 409 | by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 410 | also have "\<dots> = 1 / 2^nat (e - exponent f)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 411 | using pos by (simp add: powr_realpow[symmetric] powr_divide2[symmetric]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 412 | finally have "m * 2^nat (e - exponent f) = real (mantissa f)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 413 | using eq by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 414 | then have "mantissa f = m * 2^nat (e - exponent f)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 415 | unfolding real_of_int_inject by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 416 | with `exponent f < e` have "2 dvd mantissa f" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 417 | 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 | 418 | apply (cases "nat (e - exponent f)") | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 419 | apply auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 420 | done | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 421 | 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 | 422 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 423 | ultimately have "real m = mantissa f * 2^nat (exponent f - e)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 424 | by (simp add: powr_realpow[symmetric]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 425 | with `e \<le> exponent f` | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 426 | show "m = mantissa f * 2 ^ nat (exponent f - e)" "e = exponent f - nat (exponent f - e)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 427 | unfolding real_of_int_inject by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 428 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 429 | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 430 | lemma compute_float_zero[code_unfold, code]: "0 = Float 0 0" | 
| 47600 | 431 | by transfer simp | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 432 | hide_fact (open) compute_float_zero | 
| 47600 | 433 | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 434 | lemma compute_float_one[code_unfold, code]: "1 = Float 1 0" | 
| 47600 | 435 | by transfer simp | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 436 | hide_fact (open) compute_float_one | 
| 47600 | 437 | |
| 438 | definition normfloat :: "float \<Rightarrow> float" where | |
| 439 | [simp]: "normfloat x = x" | |
| 440 | ||
| 441 | lemma compute_normfloat[code]: "normfloat (Float m e) = | |
| 442 | (if m mod 2 = 0 \<and> m \<noteq> 0 then normfloat (Float (m div 2) (e + 1)) | |
| 443 | else if m = 0 then 0 else Float m e)" | |
| 444 | unfolding normfloat_def | |
| 445 | by transfer (auto simp add: powr_add zmod_eq_0_iff) | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 446 | hide_fact (open) compute_normfloat | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 447 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 448 | lemma compute_float_numeral[code_abbrev]: "Float (numeral k) 0 = numeral k" | 
| 47600 | 449 | by transfer simp | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 450 | hide_fact (open) compute_float_numeral | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 451 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 452 | lemma compute_float_neg_numeral[code_abbrev]: "Float (- numeral k) 0 = - numeral k" | 
| 47600 | 453 | by transfer simp | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 454 | hide_fact (open) compute_float_neg_numeral | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 455 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 456 | lemma compute_float_uminus[code]: "- Float m1 e1 = Float (- m1) e1" | 
| 47600 | 457 | by transfer simp | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 458 | hide_fact (open) compute_float_uminus | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 459 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 460 | lemma compute_float_times[code]: "Float m1 e1 * Float m2 e2 = Float (m1 * m2) (e1 + e2)" | 
| 47600 | 461 | by transfer (simp add: field_simps powr_add) | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 462 | hide_fact (open) compute_float_times | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 463 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 464 | 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 | 465 | (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 | 466 | 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 | 467 | else Float (m2 + m1 * 2^nat (e1 - e2)) e2)" | 
| 47600 | 468 | by transfer (simp add: field_simps powr_realpow[symmetric] powr_divide2[symmetric]) | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 469 | hide_fact (open) compute_float_plus | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 470 | |
| 47600 | 471 | lemma compute_float_minus[code]: fixes f g::float shows "f - g = f + (-g)" | 
| 472 | by simp | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 473 | hide_fact (open) compute_float_minus | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 474 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 475 | lemma compute_float_sgn[code]: "sgn (Float m1 e1) = (if 0 < m1 then 1 else if m1 < 0 then -1 else 0)" | 
| 47600 | 476 | by transfer (simp add: sgn_times) | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 477 | hide_fact (open) compute_float_sgn | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 478 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 479 | 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 | 480 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 481 | lemma compute_is_float_pos[code]: "is_float_pos (Float m e) \<longleftrightarrow> 0 < m" | 
| 47600 | 482 | by transfer (auto simp add: zero_less_mult_iff not_le[symmetric, of _ 0]) | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 483 | hide_fact (open) compute_is_float_pos | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 484 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 485 | lemma compute_float_less[code]: "a < b \<longleftrightarrow> is_float_pos (b - a)" | 
| 47600 | 486 | by transfer (simp add: field_simps) | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 487 | hide_fact (open) compute_float_less | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 488 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 489 | 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 | 490 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 491 | lemma compute_is_float_nonneg[code]: "is_float_nonneg (Float m e) \<longleftrightarrow> 0 \<le> m" | 
| 47600 | 492 | by transfer (auto simp add: zero_le_mult_iff not_less[symmetric, of _ 0]) | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 493 | hide_fact (open) compute_is_float_nonneg | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 494 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 495 | lemma compute_float_le[code]: "a \<le> b \<longleftrightarrow> is_float_nonneg (b - a)" | 
| 47600 | 496 | by transfer (simp add: field_simps) | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 497 | hide_fact (open) compute_float_le | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 498 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 499 | 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 | 500 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 501 | lemma compute_is_float_zero[code]: "is_float_zero (Float m e) \<longleftrightarrow> 0 = m" | 
| 47600 | 502 | by transfer (auto simp add: is_float_zero_def) | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 503 | hide_fact (open) compute_is_float_zero | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 504 | |
| 47600 | 505 | lemma compute_float_abs[code]: "abs (Float m e) = Float (abs m) e" | 
| 506 | by transfer (simp add: abs_mult) | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 507 | hide_fact (open) compute_float_abs | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 508 | |
| 47600 | 509 | lemma compute_float_eq[code]: "equal_class.equal f g = is_float_zero (f - g)" | 
| 510 | by transfer simp | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 511 | hide_fact (open) compute_float_eq | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 512 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 513 | subsection {* Rounding Real numbers *}
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 514 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 515 | definition round_down :: "int \<Rightarrow> real \<Rightarrow> real" where | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 516 | "round_down prec x = floor (x * 2 powr prec) * 2 powr -prec" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 517 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 518 | definition round_up :: "int \<Rightarrow> real \<Rightarrow> real" where | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 519 | "round_up prec x = ceiling (x * 2 powr prec) * 2 powr -prec" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 520 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 521 | 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 | 522 | unfolding round_down_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 523 | by (auto intro!: times_float simp: real_of_int_minus[symmetric] simp del: real_of_int_minus) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 524 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 525 | 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 | 526 | unfolding round_up_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 527 | by (auto intro!: times_float simp: real_of_int_minus[symmetric] simp del: real_of_int_minus) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 528 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 529 | lemma round_up: "x \<le> round_up prec x" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 530 | by (simp add: powr_minus_divide le_divide_eq round_up_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 531 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 532 | lemma round_down: "round_down prec x \<le> x" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 533 | 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 | 534 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 535 | 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 | 536 | unfolding round_up_def by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 537 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 538 | 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 | 539 | unfolding round_down_def by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 540 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 541 | lemma round_up_diff_round_down: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 542 | "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 | 543 | proof - | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 544 | have "round_up prec x - round_down prec x = | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 545 | (ceiling (x * 2 powr prec) - floor (x * 2 powr prec)) * 2 powr -prec" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 546 | 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 | 547 | also have "\<dots> \<le> 1 * 2 powr -prec" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 548 | by (rule mult_mono) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 549 | (auto simp del: real_of_int_diff | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 550 | simp: real_of_int_diff[symmetric] real_of_int_le_one_cancel_iff ceiling_diff_floor_le_1) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 551 | finally show ?thesis by simp | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 552 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 553 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 554 | 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 | 555 | unfolding round_down_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 556 | 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 | 557 | (simp add: powr_add[symmetric]) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 558 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 559 | 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 | 560 | unfolding round_up_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 561 | 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 | 562 | (simp add: powr_add[symmetric]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 563 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 564 | subsection {* Rounding Floats *}
 | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 565 | |
| 47600 | 566 | 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 | 567 | declare float_up.rep_eq[simp] | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 568 | |
| 54782 | 569 | lemma round_up_correct: | 
| 570 |   shows "round_up e f - f \<in> {0..2 powr -e}"
 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 571 | unfolding atLeastAtMost_iff | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 572 | proof | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 573 | have "round_up e f - f \<le> round_up e f - round_down e f" using round_down by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 574 | also have "\<dots> \<le> 2 powr -e" using round_up_diff_round_down by simp | 
| 54782 | 575 | finally show "round_up e f - f \<le> 2 powr real (- e)" | 
| 47600 | 576 | by simp | 
| 577 | 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 | 578 | |
| 54782 | 579 | lemma float_up_correct: | 
| 580 |   shows "real (float_up e f) - real f \<in> {0..2 powr -e}"
 | |
| 581 | by transfer (rule round_up_correct) | |
| 582 | ||
| 47600 | 583 | 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 | 584 | declare float_down.rep_eq[simp] | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 585 | |
| 54782 | 586 | lemma round_down_correct: | 
| 587 |   shows "f - (round_down e f) \<in> {0..2 powr -e}"
 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 588 | unfolding atLeastAtMost_iff | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 589 | proof | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 590 | have "f - round_down e f \<le> round_up e f - round_down e f" using round_up by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 591 | also have "\<dots> \<le> 2 powr -e" using round_up_diff_round_down by simp | 
| 54782 | 592 | finally show "f - round_down e f \<le> 2 powr real (- e)" | 
| 47600 | 593 | by simp | 
| 594 | qed (simp add: algebra_simps round_down) | |
| 24301 | 595 | |
| 54782 | 596 | lemma float_down_correct: | 
| 597 |   shows "real f - real (float_down e f) \<in> {0..2 powr -e}"
 | |
| 598 | by transfer (rule round_down_correct) | |
| 599 | ||
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 600 | lemma compute_float_down[code]: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 601 | "float_down p (Float m e) = | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 602 | (if p + e < 0 then Float (m div 2^nat (-(p + e))) (-p) else Float m e)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 603 | proof cases | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 604 | assume "p + e < 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 605 | hence "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 606 | using powr_realpow[of 2 "nat (-(p + e))"] by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 607 | also have "... = 1 / 2 powr p / 2 powr e" | 
| 47600 | 608 | unfolding powr_minus_divide real_of_int_minus by (simp add: powr_add) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 609 | finally show ?thesis | 
| 47600 | 610 | using `p + e < 0` | 
| 611 | by transfer (simp add: ac_simps round_down_def floor_divide_eq_div[symmetric]) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 612 | next | 
| 47600 | 613 | assume "\<not> p + e < 0" | 
| 614 | then have r: "real e + real p = real (nat (e + p))" by simp | |
| 615 | have r: "\<lfloor>(m * 2 powr e) * 2 powr real p\<rfloor> = (m * 2 powr e) * 2 powr real p" | |
| 616 | by (auto intro: exI[where x="m*2^nat (e+p)"] | |
| 617 | simp add: ac_simps powr_add[symmetric] r powr_realpow) | |
| 618 | with `\<not> p + e < 0` show ?thesis | |
| 57862 | 619 | 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 | 620 | qed | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 621 | hide_fact (open) compute_float_down | 
| 24301 | 622 | |
| 54782 | 623 | lemma abs_round_down_le: "\<bar>f - (round_down e f)\<bar> \<le> 2 powr -e" | 
| 624 | using round_down_correct[of f e] by simp | |
| 625 | ||
| 626 | lemma abs_round_up_le: "\<bar>f - (round_up e f)\<bar> \<le> 2 powr -e" | |
| 627 | using round_up_correct[of e f] by simp | |
| 628 | ||
| 629 | lemma round_down_nonneg: "0 \<le> s \<Longrightarrow> 0 \<le> round_down p s" | |
| 56536 | 630 | by (auto simp: round_down_def) | 
| 54782 | 631 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 632 | lemma ceil_divide_floor_conv: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 633 | assumes "b \<noteq> 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 634 | shows "\<lceil>real a / real b\<rceil> = (if b dvd a then a div b else \<lfloor>real a / real b\<rfloor> + 1)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 635 | proof cases | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 636 | assume "\<not> b dvd a" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 637 | hence "a mod b \<noteq> 0" by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 638 | hence ne: "real (a mod b) / real b \<noteq> 0" using `b \<noteq> 0` by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 639 | have "\<lceil>real a / real b\<rceil> = \<lfloor>real a / real b\<rfloor> + 1" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 640 | apply (rule ceiling_eq) apply (auto simp: floor_divide_eq_div[symmetric]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 641 | proof - | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 642 | have "real \<lfloor>real a / real b\<rfloor> \<le> real a / real b" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 643 | moreover have "real \<lfloor>real a / real b\<rfloor> \<noteq> real a / real b" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 644 | apply (subst (2) real_of_int_div_aux) unfolding floor_divide_eq_div using ne `b \<noteq> 0` by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 645 | ultimately show "real \<lfloor>real a / real b\<rfloor> < real a / real b" by arith | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 646 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 647 | thus ?thesis using `\<not> b dvd a` by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 648 | qed (simp add: ceiling_def real_of_int_minus[symmetric] divide_minus_left[symmetric] | 
| 56479 
91958d4b30f7
revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
 hoelzl parents: 
56410diff
changeset | 649 | floor_divide_eq_div dvd_neg_div del: divide_minus_left real_of_int_minus) | 
| 19765 | 650 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 651 | lemma compute_float_up[code]: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 652 | "float_up p (Float m e) = | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 653 | (let P = 2^nat (-(p + e)); r = m mod P in | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 654 | if p + e < 0 then Float (m div P + (if r = 0 then 0 else 1)) (-p) else Float m e)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 655 | proof cases | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 656 | assume "p + e < 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 657 | hence "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 658 | using powr_realpow[of 2 "nat (-(p + e))"] by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 659 | also have "... = 1 / 2 powr p / 2 powr e" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 660 | unfolding powr_minus_divide real_of_int_minus by (simp add: powr_add) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 661 | finally have twopow_rewrite: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 662 | "real ((2::int) ^ nat (- (p + e))) = 1 / 2 powr real p / 2 powr real e" . | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 663 | with `p + e < 0` have powr_rewrite: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 664 | "2 powr real e * 2 powr real p = 1 / real ((2::int) ^ nat (- (p + e)))" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 665 | unfolding powr_divide2 by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 666 | show ?thesis | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 667 | proof cases | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 668 | assume "2^nat (-(p + e)) dvd m" | 
| 47615 | 669 | with `p + e < 0` twopow_rewrite show ?thesis | 
| 47600 | 670 | by transfer (auto simp: ac_simps round_up_def floor_divide_eq_div dvd_eq_mod_eq_0) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 671 | next | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 672 | assume ndvd: "\<not> 2 ^ nat (- (p + e)) dvd m" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 673 | have one_div: "real m * (1 / real ((2::int) ^ nat (- (p + e)))) = | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 674 | real m / real ((2::int) ^ nat (- (p + e)))" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 675 | by (simp add: field_simps) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 676 | have "real \<lceil>real m * (2 powr real e * 2 powr real p)\<rceil> = | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 677 | real \<lfloor>real m * (2 powr real e * 2 powr real p)\<rfloor> + 1" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 678 | using ndvd unfolding powr_rewrite one_div | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 679 | by (subst ceil_divide_floor_conv) (auto simp: field_simps) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 680 | thus ?thesis using `p + e < 0` twopow_rewrite | 
| 47600 | 681 | by transfer (auto simp: ac_simps round_up_def floor_divide_eq_div[symmetric]) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 682 | qed | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 683 | next | 
| 47600 | 684 | assume "\<not> p + e < 0" | 
| 685 | then have r1: "real e + real p = real (nat (e + p))" by simp | |
| 686 | have r: "\<lceil>(m * 2 powr e) * 2 powr real p\<rceil> = (m * 2 powr e) * 2 powr real p" | |
| 687 | by (auto simp add: ac_simps powr_add[symmetric] r1 powr_realpow | |
| 688 | intro: exI[where x="m*2^nat (e+p)"]) | |
| 689 | then show ?thesis using `\<not> p + e < 0` | |
| 57862 | 690 | by transfer (simp add: round_up_def floor_divide_eq_div field_simps powr_add powr_minus) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 691 | qed | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 692 | hide_fact (open) compute_float_up | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 693 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 694 | lemmas real_of_ints = | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 695 | real_of_int_zero | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 696 | real_of_one | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 697 | real_of_int_add | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 698 | real_of_int_minus | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 699 | real_of_int_diff | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 700 | real_of_int_mult | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 701 | real_of_int_power | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 702 | real_numeral | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 703 | lemmas real_of_nats = | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 704 | real_of_nat_zero | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 705 | real_of_nat_one | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 706 | real_of_nat_1 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 707 | real_of_nat_add | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 708 | real_of_nat_mult | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 709 | real_of_nat_power | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 710 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 711 | lemmas int_of_reals = real_of_ints[symmetric] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 712 | lemmas nat_of_reals = real_of_nats[symmetric] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 713 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 714 | lemma two_real_int: "(2::real) = real (2::int)" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 715 | lemma two_real_nat: "(2::real) = real (2::nat)" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 716 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 717 | lemma mult_cong: "a = c ==> b = d ==> a*b = c*d" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 718 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 719 | subsection {* Compute bitlen of integers *}
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 720 | |
| 47600 | 721 | definition bitlen :: "int \<Rightarrow> int" where | 
| 722 | "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 | 723 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 724 | 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 | 725 | proof - | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 726 |   {
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 727 | assume "0 > x" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 728 | have "-1 = log 2 (inverse 2)" by (subst log_inverse) simp_all | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 729 | also have "... < log 2 (-x)" using `0 > x` by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 730 | finally have "-1 < log 2 (-x)" . | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 731 | } thus "0 \<le> bitlen x" unfolding bitlen_def by (auto intro!: add_nonneg_nonneg) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 732 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 733 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 734 | lemma bitlen_bounds: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 735 | assumes "x > 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 736 | 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 | 737 | proof | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 738 | have "(2::real) ^ nat \<lfloor>log 2 (real x)\<rfloor> = 2 powr real (floor (log 2 (real x)))" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 739 | using powr_realpow[symmetric, of 2 "nat \<lfloor>log 2 (real x)\<rfloor>"] `x > 0` | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 740 | using real_nat_eq_real[of "floor (log 2 (real x))"] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 741 | by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 742 | also have "... \<le> 2 powr log 2 (real x)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 743 | by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 744 | also have "... = real x" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 745 | using `0 < x` by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 746 | finally have "2 ^ nat \<lfloor>log 2 (real x)\<rfloor> \<le> real x" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 747 | thus "2 ^ nat (bitlen x - 1) \<le> x" using `x > 0` | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 748 | by (simp add: bitlen_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 749 | next | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 750 | have "x \<le> 2 powr (log 2 x)" using `x > 0` by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 751 | also have "... < 2 ^ nat (\<lfloor>log 2 (real x)\<rfloor> + 1)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 752 | apply (simp add: powr_realpow[symmetric]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 753 | using `x > 0` by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 754 | finally show "x < 2 ^ nat (bitlen x)" using `x > 0` | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 755 | by (simp add: bitlen_def ac_simps int_of_reals del: real_of_ints) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 756 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 757 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 758 | lemma bitlen_pow2[simp]: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 759 | assumes "b > 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 760 | shows "bitlen (b * 2 ^ c) = bitlen b + c" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 761 | proof - | 
| 56544 | 762 | from assms have "b * 2 ^ c > 0" by auto | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 763 | thus ?thesis | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 764 | using floor_add[of "log 2 b" c] assms | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 765 | by (auto simp add: log_mult log_nat_power bitlen_def) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 766 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 767 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 768 | lemma bitlen_Float: | 
| 53381 | 769 | fixes m e | 
| 770 | defines "f \<equiv> Float m e" | |
| 771 | shows "bitlen (\<bar>mantissa f\<bar>) + exponent f = (if m = 0 then 0 else bitlen \<bar>m\<bar> + e)" | |
| 772 | proof (cases "m = 0") | |
| 773 | case True | |
| 774 | then show ?thesis by (simp add: f_def bitlen_def Float_def) | |
| 775 | next | |
| 776 | case False | |
| 47600 | 777 | hence "f \<noteq> float_of 0" | 
| 778 | unfolding real_of_float_eq by (simp add: f_def) | |
| 779 | hence "mantissa f \<noteq> 0" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 780 | by (simp add: mantissa_noteq_0) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 781 | moreover | 
| 53381 | 782 | obtain i where "m = mantissa f * 2 ^ i" "e = exponent f - int i" | 
| 783 | by (rule f_def[THEN denormalize_shift, OF `f \<noteq> float_of 0`]) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 784 | ultimately show ?thesis by (simp add: abs_mult) | 
| 53381 | 785 | qed | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 786 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 787 | lemma compute_bitlen[code]: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 788 | shows "bitlen x = (if x > 0 then bitlen (x div 2) + 1 else 0)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 789 | proof - | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 790 |   { assume "2 \<le> x"
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 791 | 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 | 792 | by (simp add: log_mult zmod_zdiv_equality') | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 793 | also have "\<dots> = \<lfloor>log 2 (real x)\<rfloor>" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 794 | proof cases | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 795 | assume "x mod 2 = 0" then show ?thesis by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 796 | next | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 797 | def n \<equiv> "\<lfloor>log 2 (real x)\<rfloor>" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 798 | then have "0 \<le> n" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 799 | using `2 \<le> x` by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 800 | assume "x mod 2 \<noteq> 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 801 | with `2 \<le> x` have "x mod 2 = 1" "\<not> 2 dvd x" by (auto simp add: dvd_eq_mod_eq_0) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 802 | with `2 \<le> x` have "x \<noteq> 2^nat n" by (cases "nat n") auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 803 | moreover | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 804 |       { have "real (2^nat n :: int) = 2 powr (nat n)"
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 805 | by (simp add: powr_realpow) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 806 | also have "\<dots> \<le> 2 powr (log 2 x)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 807 | using `2 \<le> x` by (simp add: n_def del: powr_log_cancel) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 808 | finally have "2^nat n \<le> x" using `2 \<le> x` by simp } | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 809 | ultimately have "2^nat n \<le> x - 1" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 810 | then have "2^nat n \<le> real (x - 1)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 811 | unfolding real_of_int_le_iff[symmetric] by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 812 |       { have "n = \<lfloor>log 2 (2^nat n)\<rfloor>"
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 813 | using `0 \<le> n` by (simp add: log_nat_power) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 814 | also have "\<dots> \<le> \<lfloor>log 2 (x - 1)\<rfloor>" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 815 | using `2^nat n \<le> real (x - 1)` `0 \<le> n` `2 \<le> x` by (auto intro: floor_mono) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 816 | 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 | 817 | moreover have "\<lfloor>log 2 (x - 1)\<rfloor> \<le> n" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 818 | using `2 \<le> x` by (auto simp add: n_def intro!: floor_mono) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 819 | ultimately show "\<lfloor>log 2 (x - x mod 2)\<rfloor> = \<lfloor>log 2 x\<rfloor>" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 820 | unfolding n_def `x mod 2 = 1` by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 821 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 822 | 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 | 823 | moreover | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 824 |   { assume "x < 2" "0 < x"
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 825 | then have "x = 1" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 826 | then have "\<lfloor>log 2 (real x)\<rfloor> = 0" by simp } | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 827 | ultimately show ?thesis | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 828 | unfolding bitlen_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 829 | 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 | 830 | qed | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 831 | hide_fact (open) compute_bitlen | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 832 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 833 | 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 | 834 | shows "0 \<le> e + (bitlen m - 1)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 835 | proof - | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 836 | have "0 < Float m e" using assms by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 837 | hence "0 < m" using powr_gt_zero[of 2 e] | 
| 47600 | 838 | by (auto simp: zero_less_mult_iff) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 839 | hence "m \<noteq> 0" by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 840 | show ?thesis | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 841 | proof (cases "0 \<le> e") | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 842 | case True thus ?thesis using `0 < m` by (simp add: bitlen_def) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 843 | next | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 844 | have "(1::int) < 2" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 845 | case False let ?S = "2^(nat (-e))" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 846 | have "inverse (2 ^ nat (- e)) = 2 powr e" using assms False powr_realpow[of 2 "nat (-e)"] | 
| 57862 | 847 | by (auto simp: powr_minus field_simps) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 848 | hence "1 \<le> real m * inverse ?S" using assms False powr_realpow[of 2 "nat (-e)"] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 849 | by (auto simp: powr_minus) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 850 | hence "1 * ?S \<le> real m * inverse ?S * ?S" by (rule mult_right_mono, auto) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 851 | hence "?S \<le> real m" unfolding mult.assoc by auto | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 852 | hence "?S \<le> m" unfolding real_of_int_le_iff[symmetric] by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 853 | from this bitlen_bounds[OF `0 < m`, THEN conjunct2] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 854 | have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF `1 < 2`, symmetric] by (rule order_le_less_trans) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 855 | hence "-e < bitlen m" using False by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 856 | thus ?thesis by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 857 | qed | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 858 | qed | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 859 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 860 | lemma bitlen_div: assumes "0 < m" shows "1 \<le> real m / 2^nat (bitlen m - 1)" and "real m / 2^nat (bitlen m - 1) < 2" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 861 | proof - | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 862 | let ?B = "2^nat(bitlen m - 1)" | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 863 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 864 | have "?B \<le> m" using bitlen_bounds[OF `0 <m`] .. | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 865 | hence "1 * ?B \<le> real m" unfolding real_of_int_le_iff[symmetric] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 866 | thus "1 \<le> real m / ?B" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 867 | |
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 868 | have "m \<noteq> 0" using assms by auto | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 869 | have "0 \<le> bitlen m - 1" using `0 < m` by (auto simp: bitlen_def) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 870 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 871 | have "m < 2^nat(bitlen m)" using bitlen_bounds[OF `0 <m`] .. | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 872 | also have "\<dots> = 2^nat(bitlen m - 1 + 1)" using `0 < m` by (auto simp: bitlen_def) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 873 | also have "\<dots> = ?B * 2" unfolding nat_add_distrib[OF `0 \<le> bitlen m - 1` zero_le_one] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 874 | finally have "real m < 2 * ?B" unfolding real_of_int_less_iff[symmetric] by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 875 | hence "real m / ?B < 2 * ?B / ?B" by (rule divide_strict_right_mono, auto) | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 876 | thus "real m / ?B < 2" by auto | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 877 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 878 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 879 | subsection {* Approximation of positive rationals *}
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 880 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 881 | lemma zdiv_zmult_twopow_eq: fixes a b::int shows "a div b div (2 ^ n) = a div (b * 2 ^ n)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 882 | by (simp add: zdiv_zmult2_eq) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 883 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 884 | lemma div_mult_twopow_eq: fixes a b::nat shows "a div ((2::nat) ^ n) div b = a div (b * 2 ^ n)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 885 | 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 | 886 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 887 | lemma real_div_nat_eq_floor_of_divide: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 888 | fixes a b::nat | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 889 | shows "a div b = real (floor (a/b))" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 890 | by (metis floor_divide_eq_div real_of_int_of_nat_eq zdiv_int) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 891 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 892 | definition "rat_precision prec x y = int prec - (bitlen x - bitlen y)" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 893 | |
| 47600 | 894 | lift_definition lapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float" | 
| 895 | is "\<lambda>prec (x::nat) (y::nat). round_down (rat_precision prec x y) (x / y)" by simp | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 896 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 897 | lemma compute_lapprox_posrat[code]: | 
| 53381 | 898 | fixes prec x y | 
| 899 | shows "lapprox_posrat prec x y = | |
| 900 | (let | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 901 | l = rat_precision prec x y; | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 902 | d = if 0 \<le> l then x * 2^nat l div y else x div 2^nat (- l) div y | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 903 | in normfloat (Float d (- l)))" | 
| 47615 | 904 | unfolding div_mult_twopow_eq normfloat_def | 
| 47600 | 905 | by transfer | 
| 47615 | 906 | (simp add: round_down_def powr_int real_div_nat_eq_floor_of_divide field_simps Let_def | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 907 | del: two_powr_minus_int_float) | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 908 | hide_fact (open) compute_lapprox_posrat | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 909 | |
| 47600 | 910 | lift_definition rapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float" | 
| 911 | is "\<lambda>prec (x::nat) (y::nat). round_up (rat_precision prec x y) (x / y)" by simp | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 912 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 913 | (* TODO: optimize using zmod_zmult2_eq, pdivmod ? *) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 914 | lemma compute_rapprox_posrat[code]: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 915 | fixes prec x y | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 916 | defines "l \<equiv> rat_precision prec x y" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 917 | shows "rapprox_posrat prec x y = (let | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 918 | l = l ; | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 919 | X = if 0 \<le> l then (x * 2^nat l, y) else (x, y * 2^nat(-l)) ; | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 920 | d = fst X div snd X ; | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 921 | m = fst X mod snd X | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 922 | 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 | 923 | proof (cases "y = 0") | 
| 47615 | 924 | assume "y = 0" thus ?thesis unfolding normfloat_def by transfer simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 925 | next | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 926 | assume "y \<noteq> 0" | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 927 | show ?thesis | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 928 | proof (cases "0 \<le> l") | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 929 | assume "0 \<le> l" | 
| 56777 | 930 | def x' \<equiv> "x * 2 ^ nat l" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 931 | have "int x * 2 ^ nat l = x'" by (simp add: x'_def int_mult int_power) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 932 | moreover have "real x * 2 powr real l = real x'" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 933 | by (simp add: powr_realpow[symmetric] `0 \<le> l` x'_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 934 | ultimately show ?thesis | 
| 47615 | 935 | unfolding normfloat_def | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 936 | using ceil_divide_floor_conv[of y x'] powr_realpow[of 2 "nat l"] `0 \<le> l` `y \<noteq> 0` | 
| 47600 | 937 | l_def[symmetric, THEN meta_eq_to_obj_eq] | 
| 938 | by transfer | |
| 939 | (simp add: floor_divide_eq_div[symmetric] dvd_eq_mod_eq_0 round_up_def) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 940 | next | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 941 | assume "\<not> 0 \<le> l" | 
| 56777 | 942 | def y' \<equiv> "y * 2 ^ nat (- l)" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 943 | from `y \<noteq> 0` have "y' \<noteq> 0" by (simp add: y'_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 944 | have "int y * 2 ^ nat (- l) = y'" by (simp add: y'_def int_mult int_power) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 945 | moreover have "real x * real (2::int) powr real l / real y = x / real y'" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 946 | using `\<not> 0 \<le> l` | 
| 57862 | 947 | 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 | 948 | ultimately show ?thesis | 
| 47615 | 949 | unfolding normfloat_def | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 950 | using ceil_divide_floor_conv[of y' x] `\<not> 0 \<le> l` `y' \<noteq> 0` `y \<noteq> 0` | 
| 47600 | 951 | l_def[symmetric, THEN meta_eq_to_obj_eq] | 
| 952 | by transfer | |
| 953 | (simp add: round_up_def ceil_divide_floor_conv floor_divide_eq_div[symmetric] dvd_eq_mod_eq_0) | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 954 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 955 | qed | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 956 | hide_fact (open) compute_rapprox_posrat | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 957 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 958 | lemma rat_precision_pos: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 959 | assumes "0 \<le> x" and "0 < y" and "2 * x < y" and "0 < n" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 960 | 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 | 961 | proof - | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 962 |   { assume "0 < x" hence "log 2 x + 1 = log 2 (2 * x)" by (simp add: log_mult) }
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 963 | hence "bitlen (int x) < bitlen (int y)" using assms | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 964 | by (simp add: bitlen_def del: floor_add_one) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 965 | (auto intro!: floor_mono simp add: floor_add_one[symmetric] simp del: floor_add floor_add_one) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 966 | thus ?thesis | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 967 | using assms 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 | 968 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 969 | |
| 56777 | 970 | lemma power_aux: | 
| 971 | assumes "x > 0" | |
| 972 | shows "(2::int) ^ nat (x - 1) \<le> 2 ^ nat x - 1" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 973 | proof - | 
| 56777 | 974 | def y \<equiv> "nat (x - 1)" | 
| 975 | moreover | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 976 | have "(2::int) ^ y \<le> (2 ^ (y + 1)) - 1" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 977 | ultimately show ?thesis using assms by simp | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 978 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 979 | |
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 980 | lemma rapprox_posrat_less1: | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 981 | assumes "0 \<le> x" and "0 < y" and "2 * x < y" and "0 < n" | 
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 982 | shows "real (rapprox_posrat n x y) < 1" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 983 | proof - | 
| 53381 | 984 | have powr1: "2 powr real (rat_precision n (int x) (int y)) = | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 985 | 2 ^ nat (rat_precision n (int x) (int y))" using rat_precision_pos[of x y n] assms | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 986 | by (simp add: powr_realpow[symmetric]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 987 | have "x * 2 powr real (rat_precision n (int x) (int y)) / y = (x / y) * | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 988 | 2 powr real (rat_precision n (int x) (int y))" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 989 | also have "... < (1 / 2) * 2 powr real (rat_precision n (int x) (int y))" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 990 | apply (rule mult_strict_right_mono) by (insert assms) auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 991 | also have "\<dots> = 2 powr real (rat_precision n (int x) (int y) - 1)" | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 992 | using powr_add [of 2 _ "- 1", simplified add_uminus_conv_diff] by (simp add: powr_minus) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 993 | also have "\<dots> = 2 ^ nat (rat_precision n (int x) (int y) - 1)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 994 | using rat_precision_pos[of x y n] assms by (simp add: powr_realpow[symmetric]) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 995 | also have "\<dots> \<le> 2 ^ nat (rat_precision n (int x) (int y)) - 1" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 996 | unfolding int_of_reals real_of_int_le_iff | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 997 | using rat_precision_pos[OF assms] by (rule power_aux) | 
| 47600 | 998 | finally show ?thesis | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 999 | apply (transfer fixing: n x y) | 
| 57862 | 1000 | apply (simp add: round_up_def field_simps powr_minus powr1) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1001 | unfolding int_of_reals real_of_int_less_iff | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1002 | apply (simp add: ceiling_less_eq) | 
| 47600 | 1003 | done | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1004 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1005 | |
| 47600 | 1006 | lift_definition lapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is | 
| 1007 | "\<lambda>prec (x::int) (y::int). round_down (rat_precision prec \<bar>x\<bar> \<bar>y\<bar>) (x / y)" by simp | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1008 | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1009 | lemma compute_lapprox_rat[code]: | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1010 | "lapprox_rat prec x y = | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1011 | (if y = 0 then 0 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1012 | else if 0 \<le> x then | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1013 | (if 0 < y then lapprox_posrat prec (nat x) (nat y) | 
| 53381 | 1014 | else - (rapprox_posrat prec (nat x) (nat (-y)))) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1015 | else (if 0 < y | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1016 | then - (rapprox_posrat prec (nat (-x)) (nat y)) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1017 | else lapprox_posrat prec (nat (-x)) (nat (-y))))" | 
| 56479 
91958d4b30f7
revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
 hoelzl parents: 
56410diff
changeset | 1018 | by transfer (auto simp: round_up_def round_down_def ceiling_def ac_simps) | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 1019 | hide_fact (open) compute_lapprox_rat | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1020 | |
| 47600 | 1021 | lift_definition rapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is | 
| 1022 | "\<lambda>prec (x::int) (y::int). round_up (rat_precision prec \<bar>x\<bar> \<bar>y\<bar>) (x / y)" by simp | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1023 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1024 | lemma compute_rapprox_rat[code]: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1025 | "rapprox_rat prec x y = | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1026 | (if y = 0 then 0 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1027 | else if 0 \<le> x then | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1028 | (if 0 < y then rapprox_posrat prec (nat x) (nat y) | 
| 53381 | 1029 | else - (lapprox_posrat prec (nat x) (nat (-y)))) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1030 | else (if 0 < y | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1031 | then - (lapprox_posrat prec (nat (-x)) (nat y)) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1032 | else rapprox_posrat prec (nat (-x)) (nat (-y))))" | 
| 56479 
91958d4b30f7
revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
 hoelzl parents: 
56410diff
changeset | 1033 | by transfer (auto simp: round_up_def round_down_def ceiling_def ac_simps) | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 1034 | hide_fact (open) compute_rapprox_rat | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1035 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1036 | subsection {* Division *}
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1037 | |
| 54782 | 1038 | definition "real_divl prec a b = round_down (int prec + \<lfloor> log 2 \<bar>b\<bar> \<rfloor> - \<lfloor> log 2 \<bar>a\<bar> \<rfloor>) (a / b)" | 
| 1039 | ||
| 1040 | definition "real_divr prec a b = round_up (int prec + \<lfloor> log 2 \<bar>b\<bar> \<rfloor> - \<lfloor> log 2 \<bar>a\<bar> \<rfloor>) (a / b)" | |
| 1041 | ||
| 1042 | lift_definition float_divl :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is real_divl | |
| 1043 | by (simp add: real_divl_def) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1044 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1045 | lemma compute_float_divl[code]: | 
| 47600 | 1046 | "float_divl prec (Float m1 s1) (Float m2 s2) = lapprox_rat prec m1 m2 * Float 1 (s1 - s2)" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1047 | proof cases | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1048 | let ?f1 = "real m1 * 2 powr real s1" and ?f2 = "real m2 * 2 powr real s2" | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1049 | let ?m = "real m1 / real m2" and ?s = "2 powr real (s1 - s2)" | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1050 | assume not_0: "m1 \<noteq> 0 \<and> m2 \<noteq> 0" | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1051 | then have eq2: "(int prec + \<lfloor>log 2 \<bar>?f2\<bar>\<rfloor> - \<lfloor>log 2 \<bar>?f1\<bar>\<rfloor>) = rat_precision prec \<bar>m1\<bar> \<bar>m2\<bar> + (s2 - s1)" | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1052 | by (simp add: abs_mult log_mult rat_precision_def bitlen_def) | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1053 | have eq1: "real m1 * 2 powr real s1 / (real m2 * 2 powr real s2) = ?m * ?s" | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1054 | by (simp add: field_simps powr_divide2[symmetric]) | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1055 | |
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1056 | show ?thesis | 
| 53381 | 1057 | using not_0 | 
| 54782 | 1058 | by (transfer fixing: m1 s1 m2 s2 prec) (unfold eq1 eq2 round_down_shift real_divl_def, | 
| 1059 | simp add: field_simps) | |
| 1060 | qed (transfer, auto simp: real_divl_def) | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 1061 | hide_fact (open) compute_float_divl | 
| 47600 | 1062 | |
| 54782 | 1063 | lift_definition float_divr :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is real_divr | 
| 1064 | by (simp add: real_divr_def) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1065 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1066 | lemma compute_float_divr[code]: | 
| 47600 | 1067 | "float_divr prec (Float m1 s1) (Float m2 s2) = rapprox_rat prec m1 m2 * Float 1 (s1 - s2)" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1068 | proof cases | 
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1069 | let ?f1 = "real m1 * 2 powr real s1" and ?f2 = "real m2 * 2 powr real s2" | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1070 | let ?m = "real m1 / real m2" and ?s = "2 powr real (s1 - s2)" | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1071 | assume not_0: "m1 \<noteq> 0 \<and> m2 \<noteq> 0" | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1072 | then have eq2: "(int prec + \<lfloor>log 2 \<bar>?f2\<bar>\<rfloor> - \<lfloor>log 2 \<bar>?f1\<bar>\<rfloor>) = rat_precision prec \<bar>m1\<bar> \<bar>m2\<bar> + (s2 - s1)" | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1073 | by (simp add: abs_mult log_mult rat_precision_def bitlen_def) | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1074 | have eq1: "real m1 * 2 powr real s1 / (real m2 * 2 powr real s2) = ?m * ?s" | 
| 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1075 | by (simp add: field_simps powr_divide2[symmetric]) | 
| 47600 | 1076 | |
| 47601 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 hoelzl parents: 
47600diff
changeset | 1077 | show ?thesis | 
| 53381 | 1078 | using not_0 | 
| 54782 | 1079 | by (transfer fixing: m1 s1 m2 s2 prec) (unfold eq1 eq2 round_up_shift real_divr_def, | 
| 1080 | simp add: field_simps) | |
| 1081 | qed (transfer, auto simp: real_divr_def) | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 1082 | hide_fact (open) compute_float_divr | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1083 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1084 | subsection {* Lemmas needed by Approximate *}
 | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1085 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1086 | lemma Float_num[simp]: shows | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1087 | "real (Float 1 0) = 1" and "real (Float 1 1) = 2" and "real (Float 1 2) = 4" and | 
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
58042diff
changeset | 1088 | "real (Float 1 (- 1)) = 1/2" and "real (Float 1 (- 2)) = 1/4" and "real (Float 1 (- 3)) = 1/8" and | 
| 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
58042diff
changeset | 1089 | "real (Float (- 1) 0) = -1" and "real (Float (number_of n) 0) = number_of n" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1090 | using two_powr_int_float[of 2] two_powr_int_float[of "-1"] two_powr_int_float[of "-2"] two_powr_int_float[of "-3"] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1091 | using powr_realpow[of 2 2] powr_realpow[of 2 3] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1092 | using powr_minus[of 2 1] powr_minus[of 2 2] powr_minus[of 2 3] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1093 | by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1094 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1095 | lemma real_of_Float_int[simp]: "real (Float n 0) = real n" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1096 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1097 | lemma float_zero[simp]: "real (Float 0 e) = 0" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1098 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1099 | lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> abs((a::int) div 2) < abs a" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1100 | by arith | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1101 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1102 | lemma lapprox_rat: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1103 | shows "real (lapprox_rat prec x y) \<le> real x / real y" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1104 | using round_down by (simp add: lapprox_rat_def) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1105 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1106 | lemma mult_div_le: fixes a b:: int assumes "b > 0" shows "a \<ge> b * (a div b)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1107 | proof - | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1108 | from zmod_zdiv_equality'[of a b] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1109 | have "a = b * (a div b) + a mod b" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1110 | also have "... \<ge> b * (a div b) + 0" apply (rule add_left_mono) apply (rule pos_mod_sign) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1111 | using assms by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1112 | finally show ?thesis by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1113 | qed | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1114 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1115 | lemma lapprox_rat_nonneg: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1116 | fixes n x y | 
| 56777 | 1117 | defines "p \<equiv> int n - ((bitlen \<bar>x\<bar>) - (bitlen \<bar>y\<bar>))" | 
| 1118 | assumes "0 \<le> x" and "0 < y" | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1119 | shows "0 \<le> real (lapprox_rat n x y)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1120 | using assms unfolding lapprox_rat_def p_def[symmetric] round_down_def real_of_int_minus[symmetric] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1121 | powr_int[of 2, simplified] | 
| 56571 
f4635657d66f
added divide_nonneg_nonneg and co; made it a simp rule
 hoelzl parents: 
56544diff
changeset | 1122 | by auto | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1123 | |
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1124 | lemma rapprox_rat: "real x / real y \<le> real (rapprox_rat prec x y)" | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1125 | using round_up by (simp add: rapprox_rat_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1126 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1127 | lemma rapprox_rat_le1: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1128 | fixes n x y | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1129 | assumes xy: "0 \<le> x" "0 < y" "x \<le> y" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1130 | shows "real (rapprox_rat n x y) \<le> 1" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1131 | proof - | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1132 | have "bitlen \<bar>x\<bar> \<le> bitlen \<bar>y\<bar>" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1133 | using xy unfolding bitlen_def by (auto intro!: floor_mono) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1134 | then have "0 \<le> rat_precision n \<bar>x\<bar> \<bar>y\<bar>" by (simp add: rat_precision_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1135 | have "real \<lceil>real x / real y * 2 powr real (rat_precision n \<bar>x\<bar> \<bar>y\<bar>)\<rceil> | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1136 | \<le> real \<lceil>2 powr real (rat_precision n \<bar>x\<bar> \<bar>y\<bar>)\<rceil>" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1137 | using xy by (auto intro!: ceiling_mono simp: field_simps) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1138 | also have "\<dots> = 2 powr real (rat_precision n \<bar>x\<bar> \<bar>y\<bar>)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1139 | using `0 \<le> rat_precision n \<bar>x\<bar> \<bar>y\<bar>` | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1140 | by (auto intro!: exI[of _ "2^nat (rat_precision n \<bar>x\<bar> \<bar>y\<bar>)"] simp: powr_int) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1141 | finally show ?thesis | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1142 | by (simp add: rapprox_rat_def round_up_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1143 | (simp add: powr_minus inverse_eq_divide) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1144 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1145 | |
| 53381 | 1146 | lemma rapprox_rat_nonneg_neg: | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1147 | "0 \<le> x \<Longrightarrow> y < 0 \<Longrightarrow> real (rapprox_rat n x y) \<le> 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1148 | unfolding rapprox_rat_def round_up_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1149 | by (auto simp: field_simps mult_le_0_iff zero_le_mult_iff) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1150 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1151 | lemma rapprox_rat_neg: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1152 | "x < 0 \<Longrightarrow> 0 < y \<Longrightarrow> real (rapprox_rat n x y) \<le> 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1153 | unfolding rapprox_rat_def round_up_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1154 | by (auto simp: field_simps mult_le_0_iff) | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1155 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1156 | lemma rapprox_rat_nonpos_pos: | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1157 | "x \<le> 0 \<Longrightarrow> 0 < y \<Longrightarrow> real (rapprox_rat n x y) \<le> 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1158 | unfolding rapprox_rat_def round_up_def | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1159 | by (auto simp: field_simps mult_le_0_iff) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1160 | |
| 54782 | 1161 | lemma real_divl: "real_divl prec x y \<le> x / y" | 
| 1162 | by (simp add: real_divl_def round_down) | |
| 1163 | ||
| 1164 | lemma real_divr: "x / y \<le> real_divr prec x y" | |
| 1165 | using round_up by (simp add: real_divr_def) | |
| 1166 | ||
| 31098 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 hoelzl parents: 
31021diff
changeset | 1167 | lemma float_divl: "real (float_divl prec x y) \<le> real x / real y" | 
| 54782 | 1168 | by transfer (rule real_divl) | 
| 1169 | ||
| 1170 | lemma real_divl_lower_bound: | |
| 1171 | "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> real_divl prec x y" | |
| 1172 | by (simp add: real_divl_def round_down_def zero_le_mult_iff zero_le_divide_iff) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1173 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1174 | lemma float_divl_lower_bound: | 
| 54782 | 1175 | "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> real (float_divl prec x y)" | 
| 1176 | by transfer (rule real_divl_lower_bound) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1177 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1178 | lemma exponent_1: "exponent 1 = 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1179 | 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 | 1180 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1181 | lemma mantissa_1: "mantissa 1 = 1" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1182 | 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 | 1183 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1184 | lemma bitlen_1: "bitlen 1 = 1" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1185 | by (simp add: bitlen_def) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1186 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1187 | lemma mantissa_eq_zero_iff: "mantissa x = 0 \<longleftrightarrow> x = 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1188 | proof | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1189 | assume "mantissa x = 0" hence z: "0 = real x" using mantissa_exponent by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1190 | show "x = 0" by (simp add: zero_float_def z) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1191 | qed (simp add: zero_float_def) | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1192 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1193 | lemma float_upper_bound: "x \<le> 2 powr (bitlen \<bar>mantissa x\<bar> + exponent x)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1194 | proof (cases "x = 0", simp) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1195 | assume "x \<noteq> 0" hence "mantissa x \<noteq> 0" using mantissa_eq_zero_iff by auto | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1196 | have "x = mantissa x * 2 powr (exponent x)" by (rule mantissa_exponent) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1197 | also have "mantissa x \<le> \<bar>mantissa x\<bar>" by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1198 | also have "... \<le> 2 powr (bitlen \<bar>mantissa x\<bar>)" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1199 | using bitlen_bounds[of "\<bar>mantissa x\<bar>"] bitlen_nonneg `mantissa x \<noteq> 0` | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1200 | by (simp add: powr_int) (simp only: two_real_int int_of_reals real_of_int_abs[symmetric] | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1201 | real_of_int_le_iff less_imp_le) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1202 | 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 | 1203 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1204 | |
| 54782 | 1205 | lemma real_divl_pos_less1_bound: | 
| 1206 | "0 < x \<Longrightarrow> x < 1 \<Longrightarrow> prec \<ge> 1 \<Longrightarrow> 1 \<le> real_divl prec 1 x" | |
| 1207 | proof (unfold real_divl_def) | |
| 1208 | fix prec :: nat and x :: real assume x: "0 < x" "x < 1" and prec: "1 \<le> prec" | |
| 53381 | 1209 | def p \<equiv> "int prec + \<lfloor>log 2 \<bar>x\<bar>\<rfloor>" | 
| 47600 | 1210 | show "1 \<le> round_down (int prec + \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - \<lfloor>log 2 \<bar>1\<bar>\<rfloor>) (1 / x) " | 
| 1211 | proof cases | |
| 1212 | assume nonneg: "0 \<le> p" | |
| 1213 | hence "2 powr real (p) = floor (real ((2::int) ^ nat p)) * floor (1::real)" | |
| 1214 | by (simp add: powr_int del: real_of_int_power) simp | |
| 1215 | also have "floor (1::real) \<le> floor (1 / x)" using x prec by simp | |
| 1216 | also have "floor (real ((2::int) ^ nat p)) * floor (1 / x) \<le> | |
| 1217 | floor (real ((2::int) ^ nat p) * (1 / x))" | |
| 1218 | by (rule le_mult_floor) (auto simp: x prec less_imp_le) | |
| 1219 | finally have "2 powr real p \<le> floor (2 powr nat p / x)" by (simp add: powr_realpow) | |
| 1220 | thus ?thesis unfolding p_def[symmetric] | |
| 1221 | using x prec nonneg by (simp add: powr_minus inverse_eq_divide round_down_def) | |
| 1222 | next | |
| 1223 | assume neg: "\<not> 0 \<le> p" | |
| 1224 | ||
| 1225 | have "x = 2 powr (log 2 x)" | |
| 1226 | using x by simp | |
| 1227 | also have "2 powr (log 2 x) \<le> 2 powr p" | |
| 1228 | proof (rule powr_mono) | |
| 1229 | have "log 2 x \<le> \<lceil>log 2 x\<rceil>" | |
| 1230 | by simp | |
| 1231 | also have "\<dots> \<le> \<lfloor>log 2 x\<rfloor> + 1" | |
| 1232 | using ceiling_diff_floor_le_1[of "log 2 x"] by simp | |
| 1233 | also have "\<dots> \<le> \<lfloor>log 2 x\<rfloor> + prec" | |
| 1234 | using prec by simp | |
| 1235 | finally show "log 2 x \<le> real p" | |
| 1236 | using x by (simp add: p_def) | |
| 1237 | qed simp | |
| 1238 | finally have x_le: "x \<le> 2 powr p" . | |
| 1239 | ||
| 1240 | from neg have "2 powr real p \<le> 2 powr 0" | |
| 1241 | by (intro powr_mono) auto | |
| 1242 | also have "\<dots> \<le> \<lfloor>2 powr 0\<rfloor>" by simp | |
| 1243 | also have "\<dots> \<le> \<lfloor>2 powr real p / x\<rfloor>" unfolding real_of_int_le_iff | |
| 56544 | 1244 | using x x_le by (intro floor_mono) (simp add: pos_le_divide_eq) | 
| 47600 | 1245 | finally show ?thesis | 
| 1246 | using prec x unfolding p_def[symmetric] | |
| 56544 | 1247 | by (simp add: round_down_def powr_minus_divide pos_le_divide_eq) | 
| 47600 | 1248 | qed | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1249 | qed | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1250 | |
| 54782 | 1251 | lemma float_divl_pos_less1_bound: | 
| 1252 | "0 < real x \<Longrightarrow> real x < 1 \<Longrightarrow> prec \<ge> 1 \<Longrightarrow> 1 \<le> real (float_divl prec 1 x)" | |
| 1253 | 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 | 1254 | |
| 54782 | 1255 | lemma float_divr: "real x / real y \<le> real (float_divr prec x y)" | 
| 1256 | by transfer (rule real_divr) | |
| 1257 | ||
| 1258 | lemma real_divr_pos_less1_lower_bound: assumes "0 < x" and "x < 1" 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 | 1259 | proof - | 
| 54782 | 1260 | have "1 \<le> 1 / x" using `0 < x` and `x < 1` by auto | 
| 1261 | also have "\<dots> \<le> real_divr prec 1 x" using real_divr[where x=1 and y=x] by auto | |
| 47600 | 1262 | finally show ?thesis by auto | 
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1263 | qed | 
| 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1264 | |
| 54782 | 1265 | lemma float_divr_pos_less1_lower_bound: "0 < x \<Longrightarrow> x < 1 \<Longrightarrow> 1 \<le> float_divr prec 1 x" | 
| 1266 | by transfer (rule real_divr_pos_less1_lower_bound) | |
| 1267 | ||
| 1268 | lemma real_divr_nonpos_pos_upper_bound: | |
| 1269 | "x \<le> 0 \<Longrightarrow> 0 < y \<Longrightarrow> real_divr prec x y \<le> 0" | |
| 1270 | by (auto simp: field_simps mult_le_0_iff divide_le_0_iff round_up_def real_divr_def) | |
| 1271 | ||
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1272 | lemma float_divr_nonpos_pos_upper_bound: | 
| 47600 | 1273 | "real x \<le> 0 \<Longrightarrow> 0 < real y \<Longrightarrow> real (float_divr prec x y) \<le> 0" | 
| 54782 | 1274 | by transfer (rule real_divr_nonpos_pos_upper_bound) | 
| 1275 | ||
| 1276 | lemma real_divr_nonneg_neg_upper_bound: | |
| 1277 | "0 \<le> x \<Longrightarrow> y < 0 \<Longrightarrow> real_divr prec x y \<le> 0" | |
| 1278 | by (auto simp: field_simps mult_le_0_iff zero_le_mult_iff divide_le_0_iff round_up_def real_divr_def) | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1279 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1280 | lemma float_divr_nonneg_neg_upper_bound: | 
| 47600 | 1281 | "0 \<le> real x \<Longrightarrow> real y < 0 \<Longrightarrow> real (float_divr prec x y) \<le> 0" | 
| 54782 | 1282 | by transfer (rule real_divr_nonneg_neg_upper_bound) | 
| 1283 | ||
| 1284 | definition truncate_down::"nat \<Rightarrow> real \<Rightarrow> real" where | |
| 1285 | "truncate_down prec x = round_down (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) x" | |
| 1286 | ||
| 1287 | lemma truncate_down: "truncate_down prec x \<le> x" | |
| 1288 | using round_down by (simp add: truncate_down_def) | |
| 1289 | ||
| 1290 | lemma truncate_down_le: "x \<le> y \<Longrightarrow> truncate_down prec x \<le> y" | |
| 1291 | by (rule order_trans[OF truncate_down]) | |
| 47600 | 1292 | |
| 54782 | 1293 | definition truncate_up::"nat \<Rightarrow> real \<Rightarrow> real" where | 
| 1294 | "truncate_up prec x = round_up (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) x" | |
| 1295 | ||
| 1296 | lemma truncate_up: "x \<le> truncate_up prec x" | |
| 1297 | using round_up by (simp add: truncate_up_def) | |
| 1298 | ||
| 1299 | lemma truncate_up_le: "x \<le> y \<Longrightarrow> x \<le> truncate_up prec y" | |
| 1300 | by (rule order_trans[OF _ truncate_up]) | |
| 1301 | ||
| 1302 | lemma truncate_up_zero[simp]: "truncate_up prec 0 = 0" | |
| 1303 | by (simp add: truncate_up_def) | |
| 1304 | ||
| 1305 | lift_definition float_round_up :: "nat \<Rightarrow> float \<Rightarrow> float" is truncate_up | |
| 1306 | by (simp add: truncate_up_def) | |
| 47600 | 1307 | |
| 1308 | lemma float_round_up: "real x \<le> real (float_round_up prec x)" | |
| 54782 | 1309 | using truncate_up by transfer simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1310 | |
| 54782 | 1311 | lift_definition float_round_down :: "nat \<Rightarrow> float \<Rightarrow> float" is truncate_down | 
| 1312 | by (simp add: truncate_down_def) | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1313 | |
| 47600 | 1314 | lemma float_round_down: "real (float_round_down prec x) \<le> real x" | 
| 54782 | 1315 | using truncate_down by transfer simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1316 | |
| 47600 | 1317 | lemma floor_add2[simp]: "\<lfloor> real i + x \<rfloor> = i + \<lfloor> x \<rfloor>" | 
| 1318 | using floor_add[of x i] by (simp del: floor_add add: ac_simps) | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1319 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1320 | lemma compute_float_round_down[code]: | 
| 47600 | 1321 | "float_round_down prec (Float m e) = (let d = bitlen (abs m) - int prec in | 
| 1322 | if 0 < d then let P = 2^nat d ; n = m div P in Float n (e + d) | |
| 1323 | else Float m e)" | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 1324 | using Float.compute_float_down[of "prec - bitlen \<bar>m\<bar> - e" m e, symmetric] | 
| 54782 | 1325 | by transfer (simp add: field_simps abs_mult log_mult bitlen_def truncate_down_def | 
| 1326 | cong del: if_weak_cong) | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 1327 | hide_fact (open) compute_float_round_down | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1328 | |
| 47600 | 1329 | lemma compute_float_round_up[code]: | 
| 1330 | "float_round_up prec (Float m e) = (let d = (bitlen (abs m) - int prec) in | |
| 1331 | if 0 < d then let P = 2^nat d ; n = m div P ; r = m mod P | |
| 1332 | in Float (n + (if r = 0 then 0 else 1)) (e + d) | |
| 1333 | else Float m e)" | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 1334 | using Float.compute_float_up[of "prec - bitlen \<bar>m\<bar> - e" m e, symmetric] | 
| 47600 | 1335 | unfolding Let_def | 
| 54782 | 1336 | by transfer (simp add: field_simps abs_mult log_mult bitlen_def truncate_up_def | 
| 1337 | cong del: if_weak_cong) | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 1338 | hide_fact (open) compute_float_round_up | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1339 | |
| 54784 | 1340 | lemma round_up_mono: "x \<le> y \<Longrightarrow> round_up p x \<le> round_up p y" | 
| 1341 | by (auto intro!: ceiling_mono simp: round_up_def) | |
| 1342 | ||
| 1343 | lemma truncate_up_nonneg_mono: | |
| 1344 | assumes "0 \<le> x" "x \<le> y" | |
| 1345 | shows "truncate_up prec x \<le> truncate_up prec y" | |
| 1346 | proof - | |
| 1347 |   {
 | |
| 1348 | assume "\<lfloor>log 2 x\<rfloor> = \<lfloor>log 2 y\<rfloor>" | |
| 1349 | hence ?thesis | |
| 1350 | using assms | |
| 1351 | by (auto simp: truncate_up_def round_up_def intro!: ceiling_mono) | |
| 1352 |   } moreover {
 | |
| 1353 | assume "0 < x" | |
| 1354 | hence "log 2 x \<le> log 2 y" using assms by auto | |
| 1355 | moreover | |
| 1356 | assume "\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>" | |
| 1357 | ultimately have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>" | |
| 1358 | unfolding atomize_conj | |
| 1359 | by (metis floor_less_cancel linorder_cases not_le) | |
| 1360 | have "truncate_up prec x = | |
| 1361 | real \<lceil>x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1)\<rceil> * 2 powr - real (int prec - \<lfloor>log 2 x\<rfloor> - 1)" | |
| 1362 | using assms by (simp add: truncate_up_def round_up_def) | |
| 1363 | also have "\<lceil>x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1)\<rceil> \<le> (2 ^ prec)" | |
| 1364 | proof (unfold ceiling_le_eq) | |
| 1365 | have "x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> x * (2 powr real prec / (2 powr log 2 x))" | |
| 1366 | using real_of_int_floor_add_one_ge[of "log 2 x"] assms | |
| 1367 | by (auto simp add: algebra_simps powr_divide2 intro!: mult_left_mono) | |
| 1368 | thus "x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> real ((2::int) ^ prec)" | |
| 1369 | using `0 < x` by (simp add: powr_realpow) | |
| 1370 | qed | |
| 1371 | hence "real \<lceil>x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1)\<rceil> \<le> 2 powr int prec" | |
| 1372 | by (auto simp: powr_realpow) | |
| 1373 | also | |
| 1374 | have "2 powr - real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> 2 powr - real (int prec - \<lfloor>log 2 y\<rfloor>)" | |
| 1375 | using logless flogless by (auto intro!: floor_mono) | |
| 1376 | also have "2 powr real (int prec) \<le> 2 powr (log 2 y + real (int prec - \<lfloor>log 2 y\<rfloor>))" | |
| 1377 | using assms `0 < x` | |
| 1378 | by (auto simp: algebra_simps) | |
| 1379 | finally have "truncate_up prec x \<le> 2 powr (log 2 y + real (int prec - \<lfloor>log 2 y\<rfloor>)) * 2 powr - real (int prec - \<lfloor>log 2 y\<rfloor>)" | |
| 1380 | by simp | |
| 1381 | also have "\<dots> = 2 powr (log 2 y + real (int prec - \<lfloor>log 2 y\<rfloor>) - real (int prec - \<lfloor>log 2 y\<rfloor>))" | |
| 1382 | by (subst powr_add[symmetric]) simp | |
| 1383 | also have "\<dots> = y" | |
| 1384 | using `0 < x` assms | |
| 1385 | by (simp add: powr_add) | |
| 1386 | also have "\<dots> \<le> truncate_up prec y" | |
| 1387 | by (rule truncate_up) | |
| 1388 | finally have ?thesis . | |
| 1389 |   } moreover {
 | |
| 1390 | assume "~ 0 < x" | |
| 1391 | hence ?thesis | |
| 1392 | using assms | |
| 1393 | by (auto intro!: truncate_up_le) | |
| 1394 | } ultimately show ?thesis | |
| 1395 | by blast | |
| 1396 | qed | |
| 1397 | ||
| 1398 | lemma truncate_up_nonpos: "x \<le> 0 \<Longrightarrow> truncate_up prec x \<le> 0" | |
| 1399 | by (auto simp: truncate_up_def round_up_def intro!: mult_nonpos_nonneg) | |
| 1400 | ||
| 1401 | lemma truncate_down_nonpos: "x \<le> 0 \<Longrightarrow> truncate_down prec x \<le> 0" | |
| 1402 | by (auto simp: truncate_down_def round_down_def intro!: mult_nonpos_nonneg | |
| 1403 | order_le_less_trans[of _ 0, OF mult_nonpos_nonneg]) | |
| 1404 | ||
| 1405 | lemma truncate_up_switch_sign_mono: | |
| 1406 | assumes "x \<le> 0" "0 \<le> y" | |
| 1407 | shows "truncate_up prec x \<le> truncate_up prec y" | |
| 1408 | proof - | |
| 1409 | note truncate_up_nonpos[OF `x \<le> 0`] | |
| 1410 | also note truncate_up_le[OF `0 \<le> y`] | |
| 1411 | finally show ?thesis . | |
| 1412 | qed | |
| 1413 | ||
| 1414 | lemma truncate_down_zeroprec_mono: | |
| 1415 | assumes "0 < x" "x \<le> y" | |
| 1416 | shows "truncate_down 0 x \<le> truncate_down 0 y" | |
| 1417 | proof - | |
| 1418 | have "x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1) = x * inverse (2 powr ((real \<lfloor>log 2 x\<rfloor> + 1)))" | |
| 1419 | by (simp add: powr_divide2[symmetric] powr_add powr_minus inverse_eq_divide) | |
| 1420 | also have "\<dots> = 2 powr (log 2 x - (real \<lfloor>log 2 x\<rfloor>) - 1)" | |
| 1421 | using `0 < x` | |
| 57862 | 1422 | by (auto simp: field_simps powr_add powr_divide2[symmetric]) | 
| 54784 | 1423 | also have "\<dots> < 2 powr 0" | 
| 1424 | using real_of_int_floor_add_one_gt | |
| 1425 | unfolding neg_less_iff_less | |
| 1426 | by (intro powr_less_mono) (auto simp: algebra_simps) | |
| 1427 | finally have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> < 1" | |
| 1428 | unfolding less_ceiling_eq real_of_int_minus real_of_one | |
| 1429 | by simp | |
| 1430 | moreover | |
| 1431 | have "0 \<le> \<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor>" | |
| 56536 | 1432 | using `x > 0` by auto | 
| 54784 | 1433 |   ultimately have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> \<in> {0 ..< 1}"
 | 
| 1434 | by simp | |
| 1435 |   also have "\<dots> \<subseteq> {0}" by auto
 | |
| 1436 | finally have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> = 0" by simp | |
| 1437 | with assms show ?thesis | |
| 56536 | 1438 | by (auto simp: truncate_down_def round_down_def) | 
| 54784 | 1439 | qed | 
| 1440 | ||
| 1441 | lemma truncate_down_nonneg: "0 \<le> y \<Longrightarrow> 0 \<le> truncate_down prec y" | |
| 56536 | 1442 | by (auto simp: truncate_down_def round_down_def) | 
| 54784 | 1443 | |
| 1444 | lemma truncate_down_zero: "truncate_down prec 0 = 0" | |
| 56536 | 1445 | by (auto simp: truncate_down_def round_down_def) | 
| 54784 | 1446 | |
| 1447 | lemma truncate_down_switch_sign_mono: | |
| 1448 | assumes "x \<le> 0" "0 \<le> y" | |
| 1449 | assumes "x \<le> y" | |
| 1450 | shows "truncate_down prec x \<le> truncate_down prec y" | |
| 1451 | proof - | |
| 1452 | note truncate_down_nonpos[OF `x \<le> 0`] | |
| 1453 | also note truncate_down_nonneg[OF `0 \<le> y`] | |
| 1454 | finally show ?thesis . | |
| 1455 | qed | |
| 1456 | ||
| 1457 | lemma truncate_up_uminus_truncate_down: | |
| 1458 | "truncate_up prec x = - truncate_down prec (- x)" | |
| 1459 | "truncate_up prec (-x) = - truncate_down prec x" | |
| 1460 | by (auto simp: truncate_up_def round_up_def truncate_down_def round_down_def ceiling_def) | |
| 1461 | ||
| 1462 | lemma truncate_down_uminus_truncate_up: | |
| 1463 | "truncate_down prec x = - truncate_up prec (- x)" | |
| 1464 | "truncate_down prec (-x) = - truncate_up prec x" | |
| 1465 | by (auto simp: truncate_up_def round_up_def truncate_down_def round_down_def ceiling_def) | |
| 1466 | ||
| 1467 | lemma truncate_down_nonneg_mono: | |
| 1468 | assumes "0 \<le> x" "x \<le> y" | |
| 1469 | shows "truncate_down prec x \<le> truncate_down prec y" | |
| 1470 | proof - | |
| 1471 |   {
 | |
| 1472 | assume "0 < x" "prec = 0" | |
| 1473 | with assms have ?thesis | |
| 1474 | by (simp add: truncate_down_zeroprec_mono) | |
| 1475 |   } moreover {
 | |
| 1476 | assume "~ 0 < x" | |
| 1477 | with assms have "x = 0" "0 \<le> y" by simp_all | |
| 1478 | hence ?thesis | |
| 1479 | by (auto simp add: truncate_down_zero intro!: truncate_down_nonneg) | |
| 1480 |   } moreover {
 | |
| 1481 | assume "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> = \<lfloor>log 2 \<bar>y\<bar>\<rfloor>" | |
| 1482 | hence ?thesis | |
| 1483 | using assms | |
| 1484 | by (auto simp: truncate_down_def round_down_def intro!: floor_mono) | |
| 1485 |   } moreover {
 | |
| 1486 | assume "0 < x" | |
| 1487 | hence "log 2 x \<le> log 2 y" "0 < y" "0 \<le> y" using assms by auto | |
| 1488 | moreover | |
| 1489 | assume "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>" | |
| 1490 | ultimately have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>" | |
| 1491 | unfolding atomize_conj abs_of_pos[OF `0 < x`] abs_of_pos[OF `0 < y`] | |
| 1492 | by (metis floor_less_cancel linorder_cases not_le) | |
| 1493 | assume "prec \<noteq> 0" hence [simp]: "prec \<ge> Suc 0" by auto | |
| 1494 | have "2 powr (prec - 1) \<le> y * 2 powr real (prec - 1) / (2 powr log 2 y)" | |
| 1495 | using `0 < y` | |
| 1496 | by simp | |
| 1497 | also have "\<dots> \<le> y * 2 powr real prec / (2 powr (real \<lfloor>log 2 y\<rfloor> + 1))" | |
| 1498 | using `0 \<le> y` `0 \<le> x` assms(2) | |
| 56544 | 1499 | by (auto intro!: powr_mono divide_left_mono | 
| 54784 | 1500 | simp: real_of_nat_diff powr_add | 
| 1501 | powr_divide2[symmetric]) | |
| 1502 | also have "\<dots> = y * 2 powr real prec / (2 powr real \<lfloor>log 2 y\<rfloor> * 2)" | |
| 1503 | by (auto simp: powr_add) | |
| 1504 | finally have "(2 ^ (prec - 1)) \<le> \<lfloor>y * 2 powr real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)\<rfloor>" | |
| 1505 | using `0 \<le> y` | |
| 1506 | by (auto simp: powr_divide2[symmetric] le_floor_eq powr_realpow) | |
| 1507 | hence "(2 ^ (prec - 1)) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1) \<le> truncate_down prec y" | |
| 1508 | by (auto simp: truncate_down_def round_down_def) | |
| 1509 | moreover | |
| 1510 |     {
 | |
| 1511 | have "x = 2 powr (log 2 \<bar>x\<bar>)" using `0 < x` by simp | |
| 1512 | also have "\<dots> \<le> (2 ^ (prec )) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1)" | |
| 1513 | using real_of_int_floor_add_one_ge[of "log 2 \<bar>x\<bar>"] | |
| 1514 | by (auto simp: powr_realpow[symmetric] powr_add[symmetric] algebra_simps) | |
| 1515 | also | |
| 1516 | have "2 powr - real (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) \<le> 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor>)" | |
| 1517 | using logless flogless `x > 0` `y > 0` | |
| 1518 | by (auto intro!: floor_mono) | |
| 1519 | finally have "x \<le> (2 ^ (prec - 1)) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)" | |
| 1520 | by (auto simp: powr_realpow[symmetric] powr_divide2[symmetric] assms real_of_nat_diff) | |
| 1521 | } ultimately have ?thesis | |
| 1522 | by (metis dual_order.trans truncate_down) | |
| 1523 | } ultimately show ?thesis by blast | |
| 1524 | qed | |
| 1525 | ||
| 1526 | lemma truncate_down_mono: "x \<le> y \<Longrightarrow> truncate_down p x \<le> truncate_down p y" | |
| 1527 | apply (cases "0 \<le> x") | |
| 1528 | apply (rule truncate_down_nonneg_mono, assumption+) | |
| 1529 | apply (simp add: truncate_down_uminus_truncate_up) | |
| 1530 | apply (cases "0 \<le> y") | |
| 1531 | apply (auto intro: truncate_up_nonneg_mono truncate_up_switch_sign_mono) | |
| 1532 | done | |
| 1533 | ||
| 1534 | lemma truncate_up_mono: "x \<le> y \<Longrightarrow> truncate_up p x \<le> truncate_up p y" | |
| 1535 | by (simp add: truncate_up_uminus_truncate_down truncate_down_mono) | |
| 1536 | ||
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1537 | lemma Float_le_zero_iff: "Float a b \<le> 0 \<longleftrightarrow> a \<le> 0" | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1538 | apply (auto simp: zero_float_def mult_le_0_iff) | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1539 | using powr_gt_zero[of 2 b] by simp | 
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1540 | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 1541 | lemma real_of_float_pprt[simp]: fixes a::float shows "real (pprt a) = pprt (real a)" | 
| 47600 | 1542 | 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 | 1543 | |
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 1544 | lemma real_of_float_nprt[simp]: fixes a::float shows "real (nprt a) = nprt (real a)" | 
| 47600 | 1545 | 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 | 1546 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
54784diff
changeset | 1547 | 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 | 1548 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1549 | 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 | 1550 | "int_floor_fl (Float m e) = (if 0 \<le> e then m * 2 ^ nat e else m div (2 ^ (nat (-e))))" | 
| 47600 | 1551 | by transfer (simp add: powr_int int_of_reals floor_divide_eq_div del: real_of_ints) | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 1552 | hide_fact (open) compute_int_floor_fl | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1553 | |
| 47600 | 1554 | lift_definition floor_fl :: "float \<Rightarrow> float" is "\<lambda>x. real (floor x)" by simp | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1555 | |
| 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1556 | 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 | 1557 | "floor_fl (Float m e) = (if 0 \<le> e then Float m e else Float (m div (2 ^ (nat (-e)))) 0)" | 
| 47600 | 1558 | by transfer (simp add: powr_int int_of_reals floor_divide_eq_div del: real_of_ints) | 
| 47621 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 hoelzl parents: 
47615diff
changeset | 1559 | hide_fact (open) compute_floor_fl | 
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1560 | |
| 47600 | 1561 | lemma floor_fl: "real (floor_fl x) \<le> real x" by transfer simp | 
| 1562 | ||
| 1563 | lemma int_floor_fl: "real (int_floor_fl x) \<le> real x" by transfer simp | |
| 29804 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 hoelzl parents: 
29667diff
changeset | 1564 | |
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1565 | lemma floor_pos_exp: "exponent (floor_fl x) \<ge> 0" | 
| 53381 | 1566 | proof (cases "floor_fl x = float_of 0") | 
| 1567 | case True | |
| 1568 | then show ?thesis by (simp add: floor_fl_def) | |
| 1569 | next | |
| 1570 | case False | |
| 1571 | have eq: "floor_fl x = Float \<lfloor>real x\<rfloor> 0" by transfer simp | |
| 1572 | obtain i where "\<lfloor>real x\<rfloor> = mantissa (floor_fl x) * 2 ^ i" "0 = exponent (floor_fl x) - int i" | |
| 1573 | by (rule denormalize_shift[OF eq[THEN eq_reflection] False]) | |
| 1574 | then show ?thesis by simp | |
| 1575 | qed | |
| 16782 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1576 | |
| 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 obua parents: diff
changeset | 1577 | end | 
| 47599 
400b158f1589
replace the float datatype by a type with unique representation
 hoelzl parents: 
47230diff
changeset | 1578 |