| author | blanchet | 
| Thu, 06 Dec 2012 11:25:10 +0100 | |
| changeset 50389 | ad0ac9112d2c | 
| parent 49834 | b27bbb021df1 | 
| child 51375 | d9e62d9c98de | 
| 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  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
9  | 
imports Complex_Main "~~/src/HOL/Library/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: 
47937 
diff
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: 
47937 
diff
changeset
 | 
13  | 
|
| 49834 | 14  | 
typedef float = float  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
15  | 
morphisms real_of_float float_of  | 
| 
49812
 
e3945ddcb9aa
eliminated some remaining uses of typedef with implicit set definition;
 
wenzelm 
parents: 
47937 
diff
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  | 
|
| 
47601
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
18  | 
defs (overloaded)  | 
| 
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
19  | 
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: 
47600 
diff
changeset
 | 
20  | 
|
| 
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
21  | 
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: 
47600 
diff
changeset
 | 
22  | 
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: 
47600 
diff
changeset
 | 
23  | 
|
| 
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: 
47780 
diff
changeset
 | 
24  | 
setup_lifting (no_code) type_definition_float'  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
25  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
26  | 
lemmas float_of_inject[simp]  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
27  | 
|
| 47600 | 28  | 
declare [[coercion "real :: float \<Rightarrow> real"]]  | 
29  | 
||
30  | 
lemma real_of_float_eq:  | 
|
31  | 
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: 
47230 
diff
changeset
 | 
32  | 
unfolding real_of_float_def real_of_float_inject ..  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
33  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
34  | 
lemma float_of_real[simp]: "float_of (real x) = x"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
35  | 
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: 
29667 
diff
changeset
 | 
36  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
37  | 
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: 
47230 
diff
changeset
 | 
38  | 
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
 | 
39  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
40  | 
subsection {* Real operations preserving the representation as floating point number *}
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
41  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
42  | 
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: 
47230 
diff
changeset
 | 
43  | 
by (auto simp: float_def)  | 
| 19765 | 44  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
45  | 
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: 
47230 
diff
changeset
 | 
46  | 
lemma one_float[simp]: "1 \<in> float" by (intro floatI[of 1 0]) simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
47  | 
lemma numeral_float[simp]: "numeral i \<in> float" by (intro floatI[of "numeral i" 0]) simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
48  | 
lemma neg_numeral_float[simp]: "neg_numeral i \<in> float" by (intro floatI[of "neg_numeral i" 0]) simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
49  | 
lemma real_of_int_float[simp]: "real (x :: int) \<in> float" by (intro floatI[of x 0]) simp  | 
| 47600 | 50  | 
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: 
47230 
diff
changeset
 | 
51  | 
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: 
47230 
diff
changeset
 | 
52  | 
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: 
47230 
diff
changeset
 | 
53  | 
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: 
47230 
diff
changeset
 | 
54  | 
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: 
47230 
diff
changeset
 | 
55  | 
lemma two_powr_numeral_float[simp]: "2 powr numeral i \<in> float" by (intro floatI[of 1 "numeral i"]) simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
56  | 
lemma two_powr_neg_numeral_float[simp]: "2 powr neg_numeral i \<in> float" by (intro floatI[of 1 "neg_numeral i"]) simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
57  | 
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: 
47230 
diff
changeset
 | 
58  | 
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: 
44766 
diff
changeset
 | 
59  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
60  | 
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: 
47230 
diff
changeset
 | 
61  | 
unfolding float_def  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
62  | 
proof (safe, simp)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
63  | 
fix e1 m1 e2 m2 :: int  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
64  | 
  { fix e1 m1 e2 m2 :: int assume "e1 \<le> e2"
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
65  | 
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: 
47230 
diff
changeset
 | 
66  | 
by (simp add: powr_realpow[symmetric] powr_divide2[symmetric] field_simps)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
67  | 
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: 
47230 
diff
changeset
 | 
68  | 
by blast }  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
69  | 
note * = this  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
70  | 
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: 
47230 
diff
changeset
 | 
71  | 
proof (cases e1 e2 rule: linorder_le_cases)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
72  | 
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: 
47230 
diff
changeset
 | 
73  | 
qed (rule *)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
74  | 
qed  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
75  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
76  | 
lemma uminus_float[simp]: "x \<in> float \<Longrightarrow> -x \<in> float"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
77  | 
apply (auto simp: float_def)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
78  | 
apply (rule_tac x="-x" in exI)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
79  | 
apply (rule_tac x="xa" in exI)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
80  | 
apply (simp add: field_simps)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
81  | 
done  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
82  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
83  | 
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: 
47230 
diff
changeset
 | 
84  | 
apply (auto simp: float_def)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
85  | 
apply (rule_tac x="x * xa" in exI)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
86  | 
apply (rule_tac x="xb + xc" in exI)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
87  | 
apply (simp add: powr_add)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
88  | 
done  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
89  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
90  | 
lemma minus_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: 
47230 
diff
changeset
 | 
91  | 
unfolding ab_diff_minus by (intro uminus_float plus_float)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
92  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
93  | 
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: 
47230 
diff
changeset
 | 
94  | 
by (cases x rule: linorder_cases[of 0]) auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
95  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
96  | 
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: 
47230 
diff
changeset
 | 
97  | 
by (cases x rule: linorder_cases[of 0]) (auto intro!: uminus_float)  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21256 
diff
changeset
 | 
98  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
99  | 
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: 
47230 
diff
changeset
 | 
100  | 
apply (auto simp add: float_def)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
101  | 
apply (rule_tac x="x" in exI)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
102  | 
apply (rule_tac x="xa - d" in exI)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
103  | 
apply (simp add: powr_realpow[symmetric] field_simps powr_add[symmetric])  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
104  | 
done  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
105  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
106  | 
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: 
47230 
diff
changeset
 | 
107  | 
apply (auto simp add: float_def)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
108  | 
apply (rule_tac x="x" in exI)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
109  | 
apply (rule_tac x="xa - d" in exI)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
110  | 
apply (simp add: powr_realpow[symmetric] field_simps powr_add[symmetric])  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
111  | 
done  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
112  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
113  | 
lemma div_numeral_Bit0_float[simp]:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
114  | 
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: 
47230 
diff
changeset
 | 
115  | 
proof -  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
116  | 
have "(x / numeral n) / 2^1 \<in> float"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
117  | 
by (intro x div_power_2_float)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
118  | 
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: 
47230 
diff
changeset
 | 
119  | 
by (induct n) auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
120  | 
finally show ?thesis .  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
121  | 
qed  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
122  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
123  | 
lemma div_neg_numeral_Bit0_float[simp]:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
124  | 
assumes x: "x / numeral n \<in> float" shows "x / (neg_numeral (Num.Bit0 n)) \<in> float"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
125  | 
proof -  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
126  | 
have "- (x / numeral (Num.Bit0 n)) \<in> float" using x by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
127  | 
also have "- (x / numeral (Num.Bit0 n)) = x / neg_numeral (Num.Bit0 n)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
128  | 
unfolding neg_numeral_def by (simp del: minus_numeral)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
129  | 
finally show ?thesis .  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
130  | 
qed  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
131  | 
|
| 47600 | 132  | 
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: 
47600 
diff
changeset
 | 
133  | 
declare Float.rep_eq[simp]  | 
| 
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
134  | 
|
| 47780 | 135  | 
lemma compute_real_of_float[code]:  | 
136  | 
"real_of_float (Float m e) = (if e \<ge> 0 then m * 2 ^ nat e else m / 2 ^ (nat (-e)))"  | 
|
137  | 
by (simp add: real_of_float_def[symmetric] powr_int)  | 
|
138  | 
||
| 
47601
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
139  | 
code_datatype Float  | 
| 47600 | 140  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
141  | 
subsection {* Arithmetic operations on floating point numbers *}
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
142  | 
|
| 47600 | 143  | 
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: 
47230 
diff
changeset
 | 
144  | 
begin  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
145  | 
|
| 47600 | 146  | 
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: 
47600 
diff
changeset
 | 
147  | 
declare zero_float.rep_eq[simp]  | 
| 47600 | 148  | 
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: 
47600 
diff
changeset
 | 
149  | 
declare one_float.rep_eq[simp]  | 
| 47600 | 150  | 
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: 
47600 
diff
changeset
 | 
151  | 
declare plus_float.rep_eq[simp]  | 
| 47600 | 152  | 
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: 
47600 
diff
changeset
 | 
153  | 
declare times_float.rep_eq[simp]  | 
| 47600 | 154  | 
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: 
47600 
diff
changeset
 | 
155  | 
declare minus_float.rep_eq[simp]  | 
| 47600 | 156  | 
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: 
47600 
diff
changeset
 | 
157  | 
declare uminus_float.rep_eq[simp]  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
158  | 
|
| 47600 | 159  | 
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: 
47600 
diff
changeset
 | 
160  | 
declare abs_float.rep_eq[simp]  | 
| 47600 | 161  | 
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: 
47600 
diff
changeset
 | 
162  | 
declare sgn_float.rep_eq[simp]  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
163  | 
|
| 47600 | 164  | 
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: 
47230 
diff
changeset
 | 
165  | 
|
| 47600 | 166  | 
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: 
47600 
diff
changeset
 | 
167  | 
declare less_eq_float.rep_eq[simp]  | 
| 47600 | 168  | 
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: 
47600 
diff
changeset
 | 
169  | 
declare less_float.rep_eq[simp]  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
170  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
171  | 
instance  | 
| 47600 | 172  | 
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: 
47230 
diff
changeset
 | 
173  | 
end  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
174  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
175  | 
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: 
47230 
diff
changeset
 | 
176  | 
by (induct n) simp_all  | 
| 
45495
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
44766 
diff
changeset
 | 
177  | 
|
| 47600 | 178  | 
lemma fixes x y::float  | 
179  | 
shows real_of_float_min: "real (min x y) = min (real x) (real y)"  | 
|
180  | 
and real_of_float_max: "real (max x y) = max (real x) (real y)"  | 
|
181  | 
by (simp_all add: min_def max_def)  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
182  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
183  | 
instance float :: dense_linorder  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
184  | 
proof  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
185  | 
fix a b :: float  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
186  | 
show "\<exists>c. a < c"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
187  | 
apply (intro exI[of _ "a + 1"])  | 
| 47600 | 188  | 
apply transfer  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
189  | 
apply simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
190  | 
done  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
191  | 
show "\<exists>c. c < a"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
192  | 
apply (intro exI[of _ "a - 1"])  | 
| 47600 | 193  | 
apply transfer  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
194  | 
apply simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
195  | 
done  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
196  | 
assume "a < b"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
197  | 
then show "\<exists>c. a < c \<and> c < b"  | 
| 47600 | 198  | 
apply (intro exI[of _ "(a + b) * Float 1 -1"])  | 
199  | 
apply transfer  | 
|
200  | 
apply (simp add: powr_neg_numeral)  | 
|
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
201  | 
done  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
202  | 
qed  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
203  | 
|
| 47600 | 204  | 
instantiation float :: lattice_ab_group_add  | 
| 46573 | 205  | 
begin  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
206  | 
|
| 47600 | 207  | 
definition inf_float::"float\<Rightarrow>float\<Rightarrow>float"  | 
208  | 
where "inf_float a b = min a b"  | 
|
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
209  | 
|
| 47600 | 210  | 
definition sup_float::"float\<Rightarrow>float\<Rightarrow>float"  | 
211  | 
where "sup_float a b = max a b"  | 
|
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
212  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
213  | 
instance  | 
| 47600 | 214  | 
by default  | 
215  | 
(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: 
29667 
diff
changeset
 | 
216  | 
end  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
217  | 
|
| 47600 | 218  | 
lemma float_numeral[simp]: "real (numeral x :: float) = numeral x"  | 
219  | 
apply (induct x)  | 
|
220  | 
apply simp  | 
|
221  | 
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: 
47600 
diff
changeset
 | 
222  | 
plus_float.rep_eq one_float.rep_eq plus_float numeral_float one_float)  | 
| 47600 | 223  | 
done  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
224  | 
|
| 47600 | 225  | 
lemma transfer_numeral [transfer_rule]:  | 
226  | 
"fun_rel (op =) cr_float (numeral :: _ \<Rightarrow> real) (numeral :: _ \<Rightarrow> float)"  | 
|
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
227  | 
unfolding fun_rel_def cr_float_def by simp  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
228  | 
|
| 47600 | 229  | 
lemma float_neg_numeral[simp]: "real (neg_numeral x :: float) = neg_numeral x"  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
230  | 
by (simp add: minus_numeral[symmetric] del: minus_numeral)  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46670 
diff
changeset
 | 
231  | 
|
| 47600 | 232  | 
lemma transfer_neg_numeral [transfer_rule]:  | 
233  | 
"fun_rel (op =) cr_float (neg_numeral :: _ \<Rightarrow> real) (neg_numeral :: _ \<Rightarrow> float)"  | 
|
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
234  | 
unfolding fun_rel_def cr_float_def by simp  | 
| 47600 | 235  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
236  | 
lemma  | 
| 47600 | 237  | 
shows float_of_numeral[simp]: "numeral k = float_of (numeral k)"  | 
238  | 
and float_of_neg_numeral[simp]: "neg_numeral k = float_of (neg_numeral k)"  | 
|
239  | 
unfolding real_of_float_eq by simp_all  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46670 
diff
changeset
 | 
240  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
241  | 
subsection {* Represent floats as unique mantissa and exponent *}
 | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46670 
diff
changeset
 | 
242  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
243  | 
lemma int_induct_abs[case_names less]:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
244  | 
fixes j :: int  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
245  | 
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: 
47230 
diff
changeset
 | 
246  | 
shows "P j"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
247  | 
proof (induct "nat \<bar>j\<bar>" arbitrary: j rule: less_induct)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
248  | 
case less show ?case by (rule H[OF less]) simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
249  | 
qed  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
250  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
251  | 
lemma int_cancel_factors:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
252  | 
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: 
47230 
diff
changeset
 | 
253  | 
proof (induct n rule: int_induct_abs)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
254  | 
case (less n)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
255  | 
  { fix m assume n: "n \<noteq> 0" "n = m * r"
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
256  | 
then have "\<bar>m \<bar> < \<bar>n\<bar>"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
257  | 
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: 
47230 
diff
changeset
 | 
258  | 
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: 
47230 
diff
changeset
 | 
259  | 
mult_eq_0_iff zdvd_mult_cancel1)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
260  | 
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: 
47230 
diff
changeset
 | 
261  | 
then show ?case  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
262  | 
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: 
47230 
diff
changeset
 | 
263  | 
qed  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
264  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
265  | 
lemma mult_powr_eq_mult_powr_iff_asym:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
266  | 
fixes m1 m2 e1 e2 :: int  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
267  | 
assumes m1: "\<not> 2 dvd m1" and "e1 \<le> e2"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
268  | 
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: 
47230 
diff
changeset
 | 
269  | 
proof  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
270  | 
have "m1 \<noteq> 0" using m1 unfolding dvd_def by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
271  | 
assume eq: "m1 * 2 powr e1 = m2 * 2 powr e2"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
272  | 
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: 
47230 
diff
changeset
 | 
273  | 
by (simp add: powr_divide2[symmetric] field_simps)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
274  | 
also have "\<dots> = m2 * 2^nat (e2 - e1)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
275  | 
by (simp add: powr_realpow)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
276  | 
finally have m1_eq: "m1 = m2 * 2^nat (e2 - e1)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
277  | 
unfolding real_of_int_inject .  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
278  | 
with m1 have "m1 = m2"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
279  | 
by (cases "nat (e2 - e1)") (auto simp add: dvd_def)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
280  | 
then show "m1 = m2 \<and> e1 = e2"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
281  | 
using eq `m1 \<noteq> 0` by (simp add: powr_inj)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
282  | 
qed simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
283  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
284  | 
lemma mult_powr_eq_mult_powr_iff:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
285  | 
fixes m1 m2 e1 e2 :: int  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
286  | 
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: 
47230 
diff
changeset
 | 
287  | 
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: 
47230 
diff
changeset
 | 
288  | 
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: 
47230 
diff
changeset
 | 
289  | 
by (cases e1 e2 rule: linorder_le_cases) auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
290  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
291  | 
lemma floatE_normed:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
292  | 
assumes x: "x \<in> float"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
293  | 
obtains (zero) "x = 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
294  | 
| (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: 
47230 
diff
changeset
 | 
295  | 
proof atomize_elim  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
296  | 
  { assume "x \<noteq> 0"
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
297  | 
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: 
47230 
diff
changeset
 | 
298  | 
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: 
47230 
diff
changeset
 | 
299  | 
by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
300  | 
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: 
47230 
diff
changeset
 | 
301  | 
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: 
47230 
diff
changeset
 | 
302  | 
(simp add: powr_add powr_realpow) }  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
303  | 
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: 
47230 
diff
changeset
 | 
304  | 
by blast  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
305  | 
qed  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
306  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
307  | 
lemma float_normed_cases:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
308  | 
fixes f :: float  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
309  | 
obtains (zero) "f = 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
310  | 
| (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: 
47230 
diff
changeset
 | 
311  | 
proof (atomize_elim, induct f)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
312  | 
case (float_of y) then show ?case  | 
| 47600 | 313  | 
by (cases rule: floatE_normed) (auto simp: zero_float_def)  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
314  | 
qed  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
315  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
316  | 
definition mantissa :: "float \<Rightarrow> int" where  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
317  | 
"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: 
47230 
diff
changeset
 | 
318  | 
\<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: 
47230 
diff
changeset
 | 
319  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
320  | 
definition exponent :: "float \<Rightarrow> int" where  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
321  | 
"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: 
47230 
diff
changeset
 | 
322  | 
\<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: 
47230 
diff
changeset
 | 
323  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
324  | 
lemma  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
325  | 
shows exponent_0[simp]: "exponent (float_of 0) = 0" (is ?E)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
326  | 
and mantissa_0[simp]: "mantissa (float_of 0) = 0" (is ?M)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
327  | 
proof -  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
328  | 
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: 
47230 
diff
changeset
 | 
329  | 
then show ?E ?M  | 
| 47600 | 330  | 
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: 
29667 
diff
changeset
 | 
331  | 
qed  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
332  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
333  | 
lemma  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
334  | 
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: 
47230 
diff
changeset
 | 
335  | 
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: 
47230 
diff
changeset
 | 
336  | 
proof cases  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
337  | 
assume [simp]: "f \<noteq> (float_of 0)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
338  | 
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: 
47230 
diff
changeset
 | 
339  | 
proof (cases f rule: float_normed_cases)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
340  | 
case (powr m e)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
341  | 
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: 
47230 
diff
changeset
 | 
342  | 
\<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: 
47230 
diff
changeset
 | 
343  | 
by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
344  | 
then show ?thesis  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
345  | 
unfolding exponent_def mantissa_def  | 
| 47600 | 346  | 
by (rule someI2_ex) (simp add: zero_float_def)  | 
347  | 
qed (simp add: zero_float_def)  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
348  | 
then show ?E ?D by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
349  | 
qed simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
350  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
351  | 
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: 
47230 
diff
changeset
 | 
352  | 
using mantissa_not_dvd[of f] by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
353  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
354  | 
lemma  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
355  | 
fixes m e :: int  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
356  | 
defines "f \<equiv> float_of (m * 2 powr e)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
357  | 
assumes dvd: "\<not> 2 dvd m"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
358  | 
shows mantissa_float: "mantissa f = m" (is "?M")  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
359  | 
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: 
47230 
diff
changeset
 | 
360  | 
proof cases  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
361  | 
assume "m = 0" with dvd show "mantissa f = m" by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
362  | 
next  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
363  | 
assume "m \<noteq> 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
364  | 
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: 
47230 
diff
changeset
 | 
365  | 
from mantissa_exponent[of f]  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
366  | 
have "m * 2 powr e = mantissa f * 2 powr exponent f"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
367  | 
by (auto simp add: f_def)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
368  | 
then show "?M" "?E"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
369  | 
using mantissa_not_dvd[OF f_not_0] dvd  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
370  | 
by (auto simp: mult_powr_eq_mult_powr_iff)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
371  | 
qed  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
372  | 
|
| 47600 | 373  | 
subsection {* Compute arithmetic operations *}
 | 
374  | 
||
375  | 
lemma Float_mantissa_exponent: "Float (mantissa f) (exponent f) = f"  | 
|
376  | 
unfolding real_of_float_eq mantissa_exponent[of f] by simp  | 
|
377  | 
||
378  | 
lemma Float_cases[case_names Float, cases type: float]:  | 
|
379  | 
fixes f :: float  | 
|
380  | 
obtains (Float) m e :: int where "f = Float m e"  | 
|
381  | 
using Float_mantissa_exponent[symmetric]  | 
|
382  | 
by (atomize_elim) auto  | 
|
383  | 
||
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
384  | 
lemma denormalize_shift:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
385  | 
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: 
47230 
diff
changeset
 | 
386  | 
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: 
47230 
diff
changeset
 | 
387  | 
proof  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
388  | 
from mantissa_exponent[of f] f_def  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
389  | 
have "m * 2 powr e = mantissa f * 2 powr exponent f"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
390  | 
by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
391  | 
then have eq: "m = mantissa f * 2 powr (exponent f - e)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
392  | 
by (simp add: powr_divide2[symmetric] field_simps)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
393  | 
moreover  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
394  | 
have "e \<le> exponent f"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
395  | 
proof (rule ccontr)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
396  | 
assume "\<not> e \<le> exponent f"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
397  | 
then have pos: "exponent f < e" by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
398  | 
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: 
47230 
diff
changeset
 | 
399  | 
by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
400  | 
also have "\<dots> = 1 / 2^nat (e - exponent f)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
401  | 
using pos by (simp add: powr_realpow[symmetric] powr_divide2[symmetric])  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
402  | 
finally have "m * 2^nat (e - exponent f) = real (mantissa f)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
403  | 
using eq by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
404  | 
then have "mantissa f = m * 2^nat (e - exponent f)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
405  | 
unfolding real_of_int_inject by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
406  | 
with `exponent f < e` have "2 dvd mantissa f"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
407  | 
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: 
47230 
diff
changeset
 | 
408  | 
apply (cases "nat (e - exponent f)")  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
409  | 
apply auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
410  | 
done  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
411  | 
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: 
47230 
diff
changeset
 | 
412  | 
qed  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
413  | 
ultimately have "real m = mantissa f * 2^nat (exponent f - e)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
414  | 
by (simp add: powr_realpow[symmetric])  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
415  | 
with `e \<le> exponent f`  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
416  | 
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: 
47230 
diff
changeset
 | 
417  | 
unfolding real_of_int_inject by auto  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
418  | 
qed  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
419  | 
|
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
420  | 
lemma compute_float_zero[code_unfold, code]: "0 = Float 0 0"  | 
| 47600 | 421  | 
by transfer simp  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
422  | 
hide_fact (open) compute_float_zero  | 
| 47600 | 423  | 
|
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
424  | 
lemma compute_float_one[code_unfold, code]: "1 = Float 1 0"  | 
| 47600 | 425  | 
by transfer simp  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
426  | 
hide_fact (open) compute_float_one  | 
| 47600 | 427  | 
|
428  | 
definition normfloat :: "float \<Rightarrow> float" where  | 
|
429  | 
[simp]: "normfloat x = x"  | 
|
430  | 
||
431  | 
lemma compute_normfloat[code]: "normfloat (Float m e) =  | 
|
432  | 
(if m mod 2 = 0 \<and> m \<noteq> 0 then normfloat (Float (m div 2) (e + 1))  | 
|
433  | 
else if m = 0 then 0 else Float m e)"  | 
|
434  | 
unfolding normfloat_def  | 
|
435  | 
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: 
47615 
diff
changeset
 | 
436  | 
hide_fact (open) compute_normfloat  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
437  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
438  | 
lemma compute_float_numeral[code_abbrev]: "Float (numeral k) 0 = numeral k"  | 
| 47600 | 439  | 
by transfer simp  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
440  | 
hide_fact (open) compute_float_numeral  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
441  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
442  | 
lemma compute_float_neg_numeral[code_abbrev]: "Float (neg_numeral k) 0 = neg_numeral k"  | 
| 47600 | 443  | 
by transfer simp  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
444  | 
hide_fact (open) compute_float_neg_numeral  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
445  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
446  | 
lemma compute_float_uminus[code]: "- Float m1 e1 = Float (- m1) e1"  | 
| 47600 | 447  | 
by transfer simp  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
448  | 
hide_fact (open) compute_float_uminus  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
449  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
450  | 
lemma compute_float_times[code]: "Float m1 e1 * Float m2 e2 = Float (m1 * m2) (e1 + e2)"  | 
| 47600 | 451  | 
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: 
47615 
diff
changeset
 | 
452  | 
hide_fact (open) compute_float_times  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
453  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
454  | 
lemma compute_float_plus[code]: "Float m1 e1 + Float m2 e2 =  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
455  | 
(if e1 \<le> e2 then Float (m1 + m2 * 2^nat (e2 - e1)) e1  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
456  | 
else Float (m2 + m1 * 2^nat (e1 - e2)) e2)"  | 
| 47600 | 457  | 
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: 
47615 
diff
changeset
 | 
458  | 
hide_fact (open) compute_float_plus  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
459  | 
|
| 47600 | 460  | 
lemma compute_float_minus[code]: fixes f g::float shows "f - g = f + (-g)"  | 
461  | 
by simp  | 
|
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
462  | 
hide_fact (open) compute_float_minus  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
463  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
464  | 
lemma compute_float_sgn[code]: "sgn (Float m1 e1) = (if 0 < m1 then 1 else if m1 < 0 then -1 else 0)"  | 
| 47600 | 465  | 
by transfer (simp add: sgn_times)  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
466  | 
hide_fact (open) compute_float_sgn  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
467  | 
|
| 47600 | 468  | 
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: 
47230 
diff
changeset
 | 
469  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
470  | 
lemma compute_is_float_pos[code]: "is_float_pos (Float m e) \<longleftrightarrow> 0 < m"  | 
| 47600 | 471  | 
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: 
47615 
diff
changeset
 | 
472  | 
hide_fact (open) compute_is_float_pos  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
473  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
474  | 
lemma compute_float_less[code]: "a < b \<longleftrightarrow> is_float_pos (b - a)"  | 
| 47600 | 475  | 
by transfer (simp add: field_simps)  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
476  | 
hide_fact (open) compute_float_less  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
477  | 
|
| 47600 | 478  | 
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: 
47230 
diff
changeset
 | 
479  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
480  | 
lemma compute_is_float_nonneg[code]: "is_float_nonneg (Float m e) \<longleftrightarrow> 0 \<le> m"  | 
| 47600 | 481  | 
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: 
47615 
diff
changeset
 | 
482  | 
hide_fact (open) compute_is_float_nonneg  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
483  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
484  | 
lemma compute_float_le[code]: "a \<le> b \<longleftrightarrow> is_float_nonneg (b - a)"  | 
| 47600 | 485  | 
by transfer (simp add: field_simps)  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
486  | 
hide_fact (open) compute_float_le  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
487  | 
|
| 47600 | 488  | 
lift_definition is_float_zero :: "float \<Rightarrow> bool" is "op = 0 :: real \<Rightarrow> bool" by simp  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
489  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
490  | 
lemma compute_is_float_zero[code]: "is_float_zero (Float m e) \<longleftrightarrow> 0 = m"  | 
| 47600 | 491  | 
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: 
47615 
diff
changeset
 | 
492  | 
hide_fact (open) compute_is_float_zero  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
493  | 
|
| 47600 | 494  | 
lemma compute_float_abs[code]: "abs (Float m e) = Float (abs m) e"  | 
495  | 
by transfer (simp add: abs_mult)  | 
|
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
496  | 
hide_fact (open) compute_float_abs  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
497  | 
|
| 47600 | 498  | 
lemma compute_float_eq[code]: "equal_class.equal f g = is_float_zero (f - g)"  | 
499  | 
by transfer simp  | 
|
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
500  | 
hide_fact (open) compute_float_eq  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
501  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
502  | 
subsection {* Rounding Real numbers *}
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
503  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
504  | 
definition round_down :: "int \<Rightarrow> real \<Rightarrow> real" where  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
505  | 
"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: 
47230 
diff
changeset
 | 
506  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
507  | 
definition round_up :: "int \<Rightarrow> real \<Rightarrow> real" where  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
508  | 
"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: 
47230 
diff
changeset
 | 
509  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
510  | 
lemma round_down_float[simp]: "round_down prec x \<in> float"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
511  | 
unfolding round_down_def  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
512  | 
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: 
47230 
diff
changeset
 | 
513  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
514  | 
lemma round_up_float[simp]: "round_up prec x \<in> float"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
515  | 
unfolding round_up_def  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
516  | 
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: 
47230 
diff
changeset
 | 
517  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
518  | 
lemma round_up: "x \<le> round_up prec x"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
519  | 
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: 
47230 
diff
changeset
 | 
520  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
521  | 
lemma round_down: "round_down prec x \<le> x"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
522  | 
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: 
47230 
diff
changeset
 | 
523  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
524  | 
lemma round_up_0[simp]: "round_up p 0 = 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
525  | 
unfolding round_up_def by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
526  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
527  | 
lemma round_down_0[simp]: "round_down p 0 = 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
528  | 
unfolding round_down_def by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
529  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
530  | 
lemma round_up_diff_round_down:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
531  | 
"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: 
47230 
diff
changeset
 | 
532  | 
proof -  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
533  | 
have "round_up prec x - round_down prec x =  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
534  | 
(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: 
47230 
diff
changeset
 | 
535  | 
by (simp add: round_up_def round_down_def field_simps)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
536  | 
also have "\<dots> \<le> 1 * 2 powr -prec"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
537  | 
by (rule mult_mono)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
538  | 
(auto simp del: real_of_int_diff  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
539  | 
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: 
47230 
diff
changeset
 | 
540  | 
finally show ?thesis by simp  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
541  | 
qed  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
542  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
543  | 
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: 
47230 
diff
changeset
 | 
544  | 
unfolding round_down_def  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
545  | 
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: 
47230 
diff
changeset
 | 
546  | 
(simp add: powr_add[symmetric])  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
547  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
548  | 
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: 
47230 
diff
changeset
 | 
549  | 
unfolding round_up_def  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
550  | 
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: 
47230 
diff
changeset
 | 
551  | 
(simp add: powr_add[symmetric])  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
552  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
553  | 
subsection {* Rounding Floats *}
 | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
554  | 
|
| 47600 | 555  | 
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: 
47600 
diff
changeset
 | 
556  | 
declare float_up.rep_eq[simp]  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
557  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
558  | 
lemma float_up_correct:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
559  | 
  shows "real (float_up e f) - real f \<in> {0..2 powr -e}"
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
560  | 
unfolding atLeastAtMost_iff  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
561  | 
proof  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
562  | 
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: 
47230 
diff
changeset
 | 
563  | 
also have "\<dots> \<le> 2 powr -e" using round_up_diff_round_down by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
564  | 
finally show "real (float_up e f) - real f \<le> 2 powr real (- e)"  | 
| 47600 | 565  | 
by simp  | 
566  | 
qed (simp add: algebra_simps round_up)  | 
|
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
567  | 
|
| 47600 | 568  | 
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: 
47600 
diff
changeset
 | 
569  | 
declare float_down.rep_eq[simp]  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
570  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
571  | 
lemma float_down_correct:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
572  | 
  shows "real f - real (float_down e f) \<in> {0..2 powr -e}"
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
573  | 
unfolding atLeastAtMost_iff  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
574  | 
proof  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
575  | 
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: 
47230 
diff
changeset
 | 
576  | 
also have "\<dots> \<le> 2 powr -e" using round_up_diff_round_down by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
577  | 
finally show "real f - real (float_down e f) \<le> 2 powr real (- e)"  | 
| 47600 | 578  | 
by simp  | 
579  | 
qed (simp add: algebra_simps round_down)  | 
|
| 24301 | 580  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
581  | 
lemma compute_float_down[code]:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
582  | 
"float_down p (Float m e) =  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
583  | 
(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: 
47230 
diff
changeset
 | 
584  | 
proof cases  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
585  | 
assume "p + e < 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
586  | 
hence "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
587  | 
using powr_realpow[of 2 "nat (-(p + e))"] by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
588  | 
also have "... = 1 / 2 powr p / 2 powr e"  | 
| 47600 | 589  | 
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: 
47230 
diff
changeset
 | 
590  | 
finally show ?thesis  | 
| 47600 | 591  | 
using `p + e < 0`  | 
592  | 
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: 
47230 
diff
changeset
 | 
593  | 
next  | 
| 47600 | 594  | 
assume "\<not> p + e < 0"  | 
595  | 
then have r: "real e + real p = real (nat (e + p))" by simp  | 
|
596  | 
have r: "\<lfloor>(m * 2 powr e) * 2 powr real p\<rfloor> = (m * 2 powr e) * 2 powr real p"  | 
|
597  | 
by (auto intro: exI[where x="m*2^nat (e+p)"]  | 
|
598  | 
simp add: ac_simps powr_add[symmetric] r powr_realpow)  | 
|
599  | 
with `\<not> p + e < 0` show ?thesis  | 
|
600  | 
by transfer  | 
|
601  | 
(auto simp add: round_down_def field_simps powr_add powr_minus inverse_eq_divide)  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
602  | 
qed  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
603  | 
hide_fact (open) compute_float_down  | 
| 24301 | 604  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
605  | 
lemma ceil_divide_floor_conv:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
606  | 
assumes "b \<noteq> 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
607  | 
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: 
47230 
diff
changeset
 | 
608  | 
proof cases  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
609  | 
assume "\<not> b dvd a"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
610  | 
hence "a mod b \<noteq> 0" by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
611  | 
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: 
47230 
diff
changeset
 | 
612  | 
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: 
47230 
diff
changeset
 | 
613  | 
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: 
47230 
diff
changeset
 | 
614  | 
proof -  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
615  | 
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: 
47230 
diff
changeset
 | 
616  | 
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: 
47230 
diff
changeset
 | 
617  | 
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: 
47230 
diff
changeset
 | 
618  | 
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: 
47230 
diff
changeset
 | 
619  | 
qed  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
620  | 
thus ?thesis using `\<not> b dvd a` by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
621  | 
qed (simp add: ceiling_def real_of_int_minus[symmetric] divide_minus_left[symmetric]  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
622  | 
floor_divide_eq_div dvd_neg_div del: divide_minus_left real_of_int_minus)  | 
| 19765 | 623  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
624  | 
lemma compute_float_up[code]:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
625  | 
"float_up p (Float m e) =  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
626  | 
(let P = 2^nat (-(p + e)); r = m mod P in  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
627  | 
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: 
47230 
diff
changeset
 | 
628  | 
proof cases  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
629  | 
assume "p + e < 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
630  | 
hence "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
631  | 
using powr_realpow[of 2 "nat (-(p + e))"] by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
632  | 
also have "... = 1 / 2 powr p / 2 powr e"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
633  | 
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: 
47230 
diff
changeset
 | 
634  | 
finally have twopow_rewrite:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
635  | 
"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: 
47230 
diff
changeset
 | 
636  | 
with `p + e < 0` have powr_rewrite:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
637  | 
"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: 
47230 
diff
changeset
 | 
638  | 
unfolding powr_divide2 by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
639  | 
show ?thesis  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
640  | 
proof cases  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
641  | 
assume "2^nat (-(p + e)) dvd m"  | 
| 47615 | 642  | 
with `p + e < 0` twopow_rewrite show ?thesis  | 
| 47600 | 643  | 
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: 
47230 
diff
changeset
 | 
644  | 
next  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
645  | 
assume ndvd: "\<not> 2 ^ nat (- (p + e)) dvd m"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
646  | 
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: 
47230 
diff
changeset
 | 
647  | 
real m / real ((2::int) ^ nat (- (p + e)))"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
648  | 
by (simp add: field_simps)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
649  | 
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: 
47230 
diff
changeset
 | 
650  | 
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: 
47230 
diff
changeset
 | 
651  | 
using ndvd unfolding powr_rewrite one_div  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
652  | 
by (subst ceil_divide_floor_conv) (auto simp: field_simps)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
653  | 
thus ?thesis using `p + e < 0` twopow_rewrite  | 
| 47600 | 654  | 
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: 
29667 
diff
changeset
 | 
655  | 
qed  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
656  | 
next  | 
| 47600 | 657  | 
assume "\<not> p + e < 0"  | 
658  | 
then have r1: "real e + real p = real (nat (e + p))" by simp  | 
|
659  | 
have r: "\<lceil>(m * 2 powr e) * 2 powr real p\<rceil> = (m * 2 powr e) * 2 powr real p"  | 
|
660  | 
by (auto simp add: ac_simps powr_add[symmetric] r1 powr_realpow  | 
|
661  | 
intro: exI[where x="m*2^nat (e+p)"])  | 
|
662  | 
then show ?thesis using `\<not> p + e < 0`  | 
|
663  | 
by transfer  | 
|
664  | 
(simp add: round_up_def floor_divide_eq_div field_simps powr_add powr_minus inverse_eq_divide)  | 
|
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
665  | 
qed  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
666  | 
hide_fact (open) compute_float_up  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
667  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
668  | 
lemmas real_of_ints =  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
669  | 
real_of_int_zero  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
670  | 
real_of_one  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
671  | 
real_of_int_add  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
672  | 
real_of_int_minus  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
673  | 
real_of_int_diff  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
674  | 
real_of_int_mult  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
675  | 
real_of_int_power  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
676  | 
real_numeral  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
677  | 
lemmas real_of_nats =  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
678  | 
real_of_nat_zero  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
679  | 
real_of_nat_one  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
680  | 
real_of_nat_1  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
681  | 
real_of_nat_add  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
682  | 
real_of_nat_mult  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
683  | 
real_of_nat_power  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
684  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
685  | 
lemmas int_of_reals = real_of_ints[symmetric]  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
686  | 
lemmas nat_of_reals = real_of_nats[symmetric]  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
687  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
688  | 
lemma two_real_int: "(2::real) = real (2::int)" by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
689  | 
lemma two_real_nat: "(2::real) = real (2::nat)" by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
690  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
691  | 
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: 
47230 
diff
changeset
 | 
692  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
693  | 
subsection {* Compute bitlen of integers *}
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
694  | 
|
| 47600 | 695  | 
definition bitlen :: "int \<Rightarrow> int" where  | 
696  | 
"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: 
47230 
diff
changeset
 | 
697  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
698  | 
lemma bitlen_nonneg: "0 \<le> bitlen x"  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
699  | 
proof -  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
700  | 
  {
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
701  | 
assume "0 > x"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
702  | 
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: 
47230 
diff
changeset
 | 
703  | 
also have "... < log 2 (-x)" using `0 > x` by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
704  | 
finally have "-1 < log 2 (-x)" .  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
705  | 
} 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: 
47230 
diff
changeset
 | 
706  | 
qed  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
707  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
708  | 
lemma bitlen_bounds:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
709  | 
assumes "x > 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
710  | 
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: 
47230 
diff
changeset
 | 
711  | 
proof  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
712  | 
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: 
47230 
diff
changeset
 | 
713  | 
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: 
47230 
diff
changeset
 | 
714  | 
using real_nat_eq_real[of "floor (log 2 (real x))"]  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
715  | 
by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
716  | 
also have "... \<le> 2 powr log 2 (real x)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
717  | 
by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
718  | 
also have "... = real x"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
719  | 
using `0 < x` by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
720  | 
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: 
47230 
diff
changeset
 | 
721  | 
thus "2 ^ nat (bitlen x - 1) \<le> x" using `x > 0`  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
722  | 
by (simp add: bitlen_def)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
723  | 
next  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
724  | 
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: 
47230 
diff
changeset
 | 
725  | 
also have "... < 2 ^ nat (\<lfloor>log 2 (real x)\<rfloor> + 1)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
726  | 
apply (simp add: powr_realpow[symmetric])  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
727  | 
using `x > 0` by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
728  | 
finally show "x < 2 ^ nat (bitlen x)" using `x > 0`  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
729  | 
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: 
47230 
diff
changeset
 | 
730  | 
qed  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
731  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
732  | 
lemma bitlen_pow2[simp]:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
733  | 
assumes "b > 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
734  | 
shows "bitlen (b * 2 ^ c) = bitlen b + c"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
735  | 
proof -  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
736  | 
from assms have "b * 2 ^ c > 0" by (auto intro: mult_pos_pos)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
737  | 
thus ?thesis  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
738  | 
using floor_add[of "log 2 b" c] assms  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
739  | 
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: 
29667 
diff
changeset
 | 
740  | 
qed  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
741  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
742  | 
lemma bitlen_Float:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
743  | 
fixes m e  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
744  | 
defines "f \<equiv> Float m e"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
745  | 
shows "bitlen (\<bar>mantissa f\<bar>) + exponent f = (if m = 0 then 0 else bitlen \<bar>m\<bar> + e)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
746  | 
proof cases  | 
| 47600 | 747  | 
assume "m \<noteq> 0"  | 
748  | 
hence "f \<noteq> float_of 0"  | 
|
749  | 
unfolding real_of_float_eq by (simp add: f_def)  | 
|
750  | 
hence "mantissa f \<noteq> 0"  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
751  | 
by (simp add: mantissa_noteq_0)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
752  | 
moreover  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
753  | 
from f_def[THEN denormalize_shift, OF `f \<noteq> float_of 0`] guess i .  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
754  | 
ultimately show ?thesis by (simp add: abs_mult)  | 
| 47600 | 755  | 
qed (simp add: f_def bitlen_def Float_def)  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
756  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
757  | 
lemma compute_bitlen[code]:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
758  | 
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: 
47230 
diff
changeset
 | 
759  | 
proof -  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
760  | 
  { assume "2 \<le> x"
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
761  | 
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: 
47230 
diff
changeset
 | 
762  | 
by (simp add: log_mult zmod_zdiv_equality')  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
763  | 
also have "\<dots> = \<lfloor>log 2 (real x)\<rfloor>"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
764  | 
proof cases  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
765  | 
assume "x mod 2 = 0" then show ?thesis by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
766  | 
next  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
767  | 
def n \<equiv> "\<lfloor>log 2 (real x)\<rfloor>"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
768  | 
then have "0 \<le> n"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
769  | 
using `2 \<le> x` by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
770  | 
assume "x mod 2 \<noteq> 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
771  | 
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: 
47230 
diff
changeset
 | 
772  | 
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: 
47230 
diff
changeset
 | 
773  | 
moreover  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
774  | 
      { have "real (2^nat n :: int) = 2 powr (nat n)"
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
775  | 
by (simp add: powr_realpow)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
776  | 
also have "\<dots> \<le> 2 powr (log 2 x)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
777  | 
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: 
47230 
diff
changeset
 | 
778  | 
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: 
47230 
diff
changeset
 | 
779  | 
ultimately have "2^nat n \<le> x - 1" by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
780  | 
then have "2^nat n \<le> real (x - 1)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
781  | 
unfolding real_of_int_le_iff[symmetric] by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
782  | 
      { have "n = \<lfloor>log 2 (2^nat n)\<rfloor>"
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
783  | 
using `0 \<le> n` by (simp add: log_nat_power)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
784  | 
also have "\<dots> \<le> \<lfloor>log 2 (x - 1)\<rfloor>"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
785  | 
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: 
47230 
diff
changeset
 | 
786  | 
finally have "n \<le> \<lfloor>log 2 (x - 1)\<rfloor>" . }  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
787  | 
moreover have "\<lfloor>log 2 (x - 1)\<rfloor> \<le> n"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
788  | 
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: 
47230 
diff
changeset
 | 
789  | 
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: 
47230 
diff
changeset
 | 
790  | 
unfolding n_def `x mod 2 = 1` by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
791  | 
qed  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
792  | 
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: 
47230 
diff
changeset
 | 
793  | 
moreover  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
794  | 
  { assume "x < 2" "0 < x"
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
795  | 
then have "x = 1" by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
796  | 
then have "\<lfloor>log 2 (real x)\<rfloor> = 0" by simp }  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
797  | 
ultimately show ?thesis  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
798  | 
unfolding bitlen_def  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
799  | 
by (auto simp: pos_imp_zdiv_pos_iff not_le)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
800  | 
qed  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
801  | 
hide_fact (open) compute_bitlen  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
802  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
803  | 
lemma float_gt1_scale: assumes "1 \<le> Float m e"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
804  | 
shows "0 \<le> e + (bitlen m - 1)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
805  | 
proof -  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
806  | 
have "0 < Float m e" using assms by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
807  | 
hence "0 < m" using powr_gt_zero[of 2 e]  | 
| 47600 | 808  | 
by (auto simp: zero_less_mult_iff)  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
809  | 
hence "m \<noteq> 0" by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
810  | 
show ?thesis  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
811  | 
proof (cases "0 \<le> e")  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
812  | 
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: 
29667 
diff
changeset
 | 
813  | 
next  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
814  | 
have "(1::int) < 2" by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
815  | 
case False let ?S = "2^(nat (-e))"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
816  | 
have "inverse (2 ^ nat (- e)) = 2 powr e" using assms False powr_realpow[of 2 "nat (-e)"]  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
817  | 
by (auto simp: powr_minus field_simps inverse_eq_divide)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
818  | 
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: 
47230 
diff
changeset
 | 
819  | 
by (auto simp: powr_minus)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
820  | 
hence "1 * ?S \<le> real m * inverse ?S * ?S" by (rule mult_right_mono, auto)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
821  | 
hence "?S \<le> real m" unfolding mult_assoc by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
822  | 
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: 
47230 
diff
changeset
 | 
823  | 
from this bitlen_bounds[OF `0 < m`, THEN conjunct2]  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
824  | 
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: 
47230 
diff
changeset
 | 
825  | 
hence "-e < bitlen m" using False by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
826  | 
thus ?thesis by auto  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
827  | 
qed  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
828  | 
qed  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
829  | 
|
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
830  | 
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: 
29667 
diff
changeset
 | 
831  | 
proof -  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
832  | 
let ?B = "2^nat(bitlen m - 1)"  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
833  | 
|
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
834  | 
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: 
29667 
diff
changeset
 | 
835  | 
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: 
29667 
diff
changeset
 | 
836  | 
thus "1 \<le> real m / ?B" by auto  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
837  | 
|
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
838  | 
have "m \<noteq> 0" using assms by auto  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
839  | 
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
 | 
840  | 
|
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
841  | 
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: 
47230 
diff
changeset
 | 
842  | 
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: 
29667 
diff
changeset
 | 
843  | 
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: 
29667 
diff
changeset
 | 
844  | 
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: 
29667 
diff
changeset
 | 
845  | 
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: 
29667 
diff
changeset
 | 
846  | 
thus "real m / ?B < 2" by auto  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
847  | 
qed  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
848  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
849  | 
subsection {* Approximation of positive rationals *}
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
850  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
851  | 
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: 
47230 
diff
changeset
 | 
852  | 
by (simp add: zdiv_zmult2_eq)  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
853  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
854  | 
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: 
47230 
diff
changeset
 | 
855  | 
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: 
29667 
diff
changeset
 | 
856  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
857  | 
lemma real_div_nat_eq_floor_of_divide:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
858  | 
fixes a b::nat  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
859  | 
shows "a div b = real (floor (a/b))"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
860  | 
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: 
29667 
diff
changeset
 | 
861  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
862  | 
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: 
29667 
diff
changeset
 | 
863  | 
|
| 47600 | 864  | 
lift_definition lapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float"  | 
865  | 
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
 | 
866  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
867  | 
lemma compute_lapprox_posrat[code]:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
868  | 
fixes prec x y  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
869  | 
shows "lapprox_posrat prec x y =  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
870  | 
(let  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
871  | 
l = rat_precision prec x y;  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
872  | 
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: 
47230 
diff
changeset
 | 
873  | 
in normfloat (Float d (- l)))"  | 
| 47615 | 874  | 
unfolding div_mult_twopow_eq normfloat_def  | 
| 47600 | 875  | 
by transfer  | 
| 47615 | 876  | 
(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: 
47230 
diff
changeset
 | 
877  | 
del: two_powr_minus_int_float)  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
878  | 
hide_fact (open) compute_lapprox_posrat  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
879  | 
|
| 47600 | 880  | 
lift_definition rapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float"  | 
881  | 
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: 
29667 
diff
changeset
 | 
882  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
883  | 
(* TODO: optimize using zmod_zmult2_eq, pdivmod ? *)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
884  | 
lemma compute_rapprox_posrat[code]:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
885  | 
fixes prec x y  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
886  | 
defines "l \<equiv> rat_precision prec x y"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
887  | 
shows "rapprox_posrat prec x y = (let  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
888  | 
l = l ;  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
889  | 
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: 
47230 
diff
changeset
 | 
890  | 
d = fst X div snd X ;  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
891  | 
m = fst X mod snd X  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
892  | 
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: 
47230 
diff
changeset
 | 
893  | 
proof (cases "y = 0")  | 
| 47615 | 894  | 
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: 
47230 
diff
changeset
 | 
895  | 
next  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
896  | 
assume "y \<noteq> 0"  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
897  | 
show ?thesis  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
898  | 
proof (cases "0 \<le> l")  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
899  | 
assume "0 \<le> l"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
900  | 
def x' == "x * 2 ^ nat l"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
901  | 
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: 
47230 
diff
changeset
 | 
902  | 
moreover have "real x * 2 powr real l = real x'"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
903  | 
by (simp add: powr_realpow[symmetric] `0 \<le> l` x'_def)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
904  | 
ultimately show ?thesis  | 
| 47615 | 905  | 
unfolding normfloat_def  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
906  | 
using ceil_divide_floor_conv[of y x'] powr_realpow[of 2 "nat l"] `0 \<le> l` `y \<noteq> 0`  | 
| 47600 | 907  | 
l_def[symmetric, THEN meta_eq_to_obj_eq]  | 
908  | 
by transfer  | 
|
909  | 
(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: 
47230 
diff
changeset
 | 
910  | 
next  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
911  | 
assume "\<not> 0 \<le> l"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
912  | 
def y' == "y * 2 ^ nat (- l)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
913  | 
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: 
47230 
diff
changeset
 | 
914  | 
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: 
47230 
diff
changeset
 | 
915  | 
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: 
47230 
diff
changeset
 | 
916  | 
using `\<not> 0 \<le> l`  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
917  | 
by (simp add: powr_realpow[symmetric] powr_minus y'_def field_simps inverse_eq_divide)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
918  | 
ultimately show ?thesis  | 
| 47615 | 919  | 
unfolding normfloat_def  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
920  | 
using ceil_divide_floor_conv[of y' x] `\<not> 0 \<le> l` `y' \<noteq> 0` `y \<noteq> 0`  | 
| 47600 | 921  | 
l_def[symmetric, THEN meta_eq_to_obj_eq]  | 
922  | 
by transfer  | 
|
923  | 
(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: 
29667 
diff
changeset
 | 
924  | 
qed  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
925  | 
qed  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
926  | 
hide_fact (open) compute_rapprox_posrat  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
927  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
928  | 
lemma rat_precision_pos:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
929  | 
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: 
47230 
diff
changeset
 | 
930  | 
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: 
29667 
diff
changeset
 | 
931  | 
proof -  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
932  | 
  { 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: 
47230 
diff
changeset
 | 
933  | 
hence "bitlen (int x) < bitlen (int y)" using assms  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
934  | 
by (simp add: bitlen_def del: floor_add_one)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
935  | 
(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: 
47230 
diff
changeset
 | 
936  | 
thus ?thesis  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
937  | 
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: 
29667 
diff
changeset
 | 
938  | 
qed  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
939  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
940  | 
lemma power_aux: assumes "x > 0" shows "(2::int) ^ nat (x - 1) \<le> 2 ^ nat x - 1"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
941  | 
proof -  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
942  | 
def y \<equiv> "nat (x - 1)" moreover  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
943  | 
have "(2::int) ^ y \<le> (2 ^ (y + 1)) - 1" by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
944  | 
ultimately show ?thesis using assms by simp  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
945  | 
qed  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
946  | 
|
| 
47601
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
947  | 
lemma rapprox_posrat_less1:  | 
| 
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
948  | 
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: 
31021 
diff
changeset
 | 
949  | 
shows "real (rapprox_posrat n x y) < 1"  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
950  | 
proof -  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
951  | 
have powr1: "2 powr real (rat_precision n (int x) (int y)) =  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
952  | 
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: 
47230 
diff
changeset
 | 
953  | 
by (simp add: powr_realpow[symmetric])  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
954  | 
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: 
47230 
diff
changeset
 | 
955  | 
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: 
47230 
diff
changeset
 | 
956  | 
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: 
47230 
diff
changeset
 | 
957  | 
apply (rule mult_strict_right_mono) by (insert assms) auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
958  | 
also have "\<dots> = 2 powr real (rat_precision n (int x) (int y) - 1)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
959  | 
by (simp add: powr_add diff_def powr_neg_numeral)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
960  | 
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: 
47230 
diff
changeset
 | 
961  | 
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: 
47230 
diff
changeset
 | 
962  | 
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: 
47230 
diff
changeset
 | 
963  | 
unfolding int_of_reals real_of_int_le_iff  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
964  | 
using rat_precision_pos[OF assms] by (rule power_aux)  | 
| 47600 | 965  | 
finally show ?thesis  | 
| 
47601
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
966  | 
apply (transfer fixing: n x y)  | 
| 
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
967  | 
apply (simp add: round_up_def field_simps powr_minus inverse_eq_divide powr1)  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
968  | 
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: 
47600 
diff
changeset
 | 
969  | 
apply (simp add: ceiling_less_eq)  | 
| 47600 | 970  | 
done  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
971  | 
qed  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
972  | 
|
| 47600 | 973  | 
lift_definition lapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is  | 
974  | 
"\<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
 | 
975  | 
|
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
976  | 
lemma compute_lapprox_rat[code]:  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
977  | 
"lapprox_rat prec x y =  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
978  | 
(if y = 0 then 0  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
979  | 
else if 0 \<le> x then  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
980  | 
(if 0 < y then lapprox_posrat prec (nat x) (nat y)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
981  | 
else - (rapprox_posrat prec (nat x) (nat (-y))))  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
982  | 
else (if 0 < y  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
983  | 
then - (rapprox_posrat prec (nat (-x)) (nat y))  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
984  | 
else lapprox_posrat prec (nat (-x)) (nat (-y))))"  | 
| 
47601
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
985  | 
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: 
47615 
diff
changeset
 | 
986  | 
hide_fact (open) compute_lapprox_rat  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
987  | 
|
| 47600 | 988  | 
lift_definition rapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is  | 
989  | 
"\<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: 
47230 
diff
changeset
 | 
990  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
991  | 
lemma compute_rapprox_rat[code]:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
992  | 
"rapprox_rat prec x y =  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
993  | 
(if y = 0 then 0  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
994  | 
else if 0 \<le> x then  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
995  | 
(if 0 < y then rapprox_posrat prec (nat x) (nat y)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
996  | 
else - (lapprox_posrat prec (nat x) (nat (-y))))  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
997  | 
else (if 0 < y  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
998  | 
then - (lapprox_posrat prec (nat (-x)) (nat y))  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
999  | 
else rapprox_posrat prec (nat (-x)) (nat (-y))))"  | 
| 
47601
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
1000  | 
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: 
47615 
diff
changeset
 | 
1001  | 
hide_fact (open) compute_rapprox_rat  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1002  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1003  | 
subsection {* Division *}
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1004  | 
|
| 47600 | 1005  | 
lift_definition float_divl :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is  | 
1006  | 
"\<lambda>(prec::nat) a b. round_down (prec + \<lfloor> log 2 \<bar>b\<bar> \<rfloor> - \<lfloor> log 2 \<bar>a\<bar> \<rfloor>) (a / b)" by simp  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1007  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1008  | 
lemma compute_float_divl[code]:  | 
| 47600 | 1009  | 
"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: 
47230 
diff
changeset
 | 
1010  | 
proof cases  | 
| 
47601
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
1011  | 
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: 
47600 
diff
changeset
 | 
1012  | 
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: 
47600 
diff
changeset
 | 
1013  | 
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: 
47600 
diff
changeset
 | 
1014  | 
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: 
47600 
diff
changeset
 | 
1015  | 
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: 
47600 
diff
changeset
 | 
1016  | 
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: 
47600 
diff
changeset
 | 
1017  | 
by (simp add: field_simps powr_divide2[symmetric])  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1018  | 
|
| 
47601
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
1019  | 
show ?thesis  | 
| 
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
1020  | 
using not_0  | 
| 
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
1021  | 
by (transfer fixing: m1 s1 m2 s2 prec) (unfold eq1 eq2 round_down_shift, simp add: field_simps)  | 
| 47600 | 1022  | 
qed (transfer, auto)  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
1023  | 
hide_fact (open) compute_float_divl  | 
| 47600 | 1024  | 
|
1025  | 
lift_definition float_divr :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is  | 
|
1026  | 
"\<lambda>(prec::nat) a b. round_up (prec + \<lfloor> log 2 \<bar>b\<bar> \<rfloor> - \<lfloor> log 2 \<bar>a\<bar> \<rfloor>) (a / b)" by simp  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1027  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1028  | 
lemma compute_float_divr[code]:  | 
| 47600 | 1029  | 
"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: 
47230 
diff
changeset
 | 
1030  | 
proof cases  | 
| 
47601
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
1031  | 
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: 
47600 
diff
changeset
 | 
1032  | 
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: 
47600 
diff
changeset
 | 
1033  | 
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: 
47600 
diff
changeset
 | 
1034  | 
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: 
47600 
diff
changeset
 | 
1035  | 
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: 
47600 
diff
changeset
 | 
1036  | 
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: 
47600 
diff
changeset
 | 
1037  | 
by (simp add: field_simps powr_divide2[symmetric])  | 
| 47600 | 1038  | 
|
| 
47601
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
1039  | 
show ?thesis  | 
| 
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
1040  | 
using not_0  | 
| 
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
1041  | 
by (transfer fixing: m1 s1 m2 s2 prec) (unfold eq1 eq2 round_up_shift, simp add: field_simps)  | 
| 47600 | 1042  | 
qed (transfer, auto)  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
1043  | 
hide_fact (open) compute_float_divr  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
1044  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1045  | 
subsection {* Lemmas needed by Approximate *}
 | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1046  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1047  | 
lemma Float_num[simp]: shows  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1048  | 
"real (Float 1 0) = 1" and "real (Float 1 1) = 2" and "real (Float 1 2) = 4" and  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1049  | 
"real (Float 1 -1) = 1/2" and "real (Float 1 -2) = 1/4" and "real (Float 1 -3) = 1/8" and  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1050  | 
"real (Float -1 0) = -1" and "real (Float (number_of n) 0) = number_of n"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1051  | 
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: 
47230 
diff
changeset
 | 
1052  | 
using powr_realpow[of 2 2] powr_realpow[of 2 3]  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1053  | 
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: 
47230 
diff
changeset
 | 
1054  | 
by auto  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1055  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1056  | 
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: 
47230 
diff
changeset
 | 
1057  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1058  | 
lemma float_zero[simp]: "real (Float 0 e) = 0" by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1059  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1060  | 
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: 
47230 
diff
changeset
 | 
1061  | 
by arith  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
1062  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1063  | 
lemma lapprox_rat:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1064  | 
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: 
47230 
diff
changeset
 | 
1065  | 
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
 | 
1066  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1067  | 
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: 
47230 
diff
changeset
 | 
1068  | 
proof -  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1069  | 
from zmod_zdiv_equality'[of a b]  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1070  | 
have "a = b * (a div b) + a mod b" by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1071  | 
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: 
47230 
diff
changeset
 | 
1072  | 
using assms by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1073  | 
finally show ?thesis by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1074  | 
qed  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1075  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1076  | 
lemma lapprox_rat_nonneg:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1077  | 
fixes n x y  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1078  | 
defines "p == int n - ((bitlen \<bar>x\<bar>) - (bitlen \<bar>y\<bar>))"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1079  | 
assumes "0 \<le> x" "0 < y"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1080  | 
shows "0 \<le> real (lapprox_rat n x y)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1081  | 
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: 
47230 
diff
changeset
 | 
1082  | 
powr_int[of 2, simplified]  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1083  | 
by (auto simp add: inverse_eq_divide intro!: mult_nonneg_nonneg divide_nonneg_pos mult_pos_pos)  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
1084  | 
|
| 
31098
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
31021 
diff
changeset
 | 
1085  | 
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: 
47230 
diff
changeset
 | 
1086  | 
using round_up by (simp add: rapprox_rat_def)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1087  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1088  | 
lemma rapprox_rat_le1:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1089  | 
fixes n x y  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1090  | 
assumes xy: "0 \<le> x" "0 < y" "x \<le> y"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1091  | 
shows "real (rapprox_rat n x y) \<le> 1"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1092  | 
proof -  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1093  | 
have "bitlen \<bar>x\<bar> \<le> bitlen \<bar>y\<bar>"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1094  | 
using xy unfolding bitlen_def by (auto intro!: floor_mono)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1095  | 
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: 
47230 
diff
changeset
 | 
1096  | 
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: 
47230 
diff
changeset
 | 
1097  | 
\<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: 
47230 
diff
changeset
 | 
1098  | 
using xy by (auto intro!: ceiling_mono simp: field_simps)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1099  | 
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: 
47230 
diff
changeset
 | 
1100  | 
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: 
47230 
diff
changeset
 | 
1101  | 
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: 
47230 
diff
changeset
 | 
1102  | 
finally show ?thesis  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1103  | 
by (simp add: rapprox_rat_def round_up_def)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1104  | 
(simp add: powr_minus inverse_eq_divide)  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
1105  | 
qed  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
1106  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1107  | 
lemma rapprox_rat_nonneg_neg:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1108  | 
"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: 
47230 
diff
changeset
 | 
1109  | 
unfolding rapprox_rat_def round_up_def  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1110  | 
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
 | 
1111  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1112  | 
lemma rapprox_rat_neg:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1113  | 
"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: 
47230 
diff
changeset
 | 
1114  | 
unfolding rapprox_rat_def round_up_def  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1115  | 
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: 
29667 
diff
changeset
 | 
1116  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1117  | 
lemma rapprox_rat_nonpos_pos:  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1118  | 
"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: 
47230 
diff
changeset
 | 
1119  | 
unfolding rapprox_rat_def round_up_def  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1120  | 
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
 | 
1121  | 
|
| 
31098
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
31021 
diff
changeset
 | 
1122  | 
lemma float_divl: "real (float_divl prec x y) \<le> real x / real y"  | 
| 47600 | 1123  | 
by transfer (simp add: round_down)  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1124  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1125  | 
lemma float_divl_lower_bound:  | 
| 47600 | 1126  | 
"0 \<le> x \<Longrightarrow> 0 < y \<Longrightarrow> 0 \<le> real (float_divl prec x y)"  | 
1127  | 
by transfer (simp add: 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: 
47230 
diff
changeset
 | 
1128  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1129  | 
lemma exponent_1: "exponent 1 = 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1130  | 
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: 
47230 
diff
changeset
 | 
1131  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1132  | 
lemma mantissa_1: "mantissa 1 = 1"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1133  | 
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
 | 
1134  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1135  | 
lemma bitlen_1: "bitlen 1 = 1"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1136  | 
by (simp add: bitlen_def)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1137  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1138  | 
lemma mantissa_eq_zero_iff: "mantissa x = 0 \<longleftrightarrow> x = 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1139  | 
proof  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1140  | 
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: 
47230 
diff
changeset
 | 
1141  | 
show "x = 0" by (simp add: zero_float_def z)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1142  | 
qed (simp add: zero_float_def)  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
1143  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1144  | 
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: 
47230 
diff
changeset
 | 
1145  | 
proof (cases "x = 0", simp)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1146  | 
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: 
47230 
diff
changeset
 | 
1147  | 
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: 
47230 
diff
changeset
 | 
1148  | 
also have "mantissa x \<le> \<bar>mantissa x\<bar>" by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1149  | 
also have "... \<le> 2 powr (bitlen \<bar>mantissa x\<bar>)"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1150  | 
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: 
47230 
diff
changeset
 | 
1151  | 
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: 
47230 
diff
changeset
 | 
1152  | 
real_of_int_le_iff less_imp_le)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1153  | 
finally show ?thesis by (simp add: powr_add)  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
1154  | 
qed  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
1155  | 
|
| 41528 | 1156  | 
lemma float_divl_pos_less1_bound:  | 
| 47600 | 1157  | 
"0 < real x \<Longrightarrow> real x < 1 \<Longrightarrow> prec \<ge> 1 \<Longrightarrow> 1 \<le> real (float_divl prec 1 x)"  | 
1158  | 
proof transfer  | 
|
1159  | 
fix prec :: nat and x :: real assume x: "0 < x" "x < 1" "x \<in> float" and prec: "1 \<le> prec"  | 
|
1160  | 
def p \<equiv> "int prec + \<lfloor>log 2 \<bar>x\<bar>\<rfloor>"  | 
|
1161  | 
show "1 \<le> round_down (int prec + \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - \<lfloor>log 2 \<bar>1\<bar>\<rfloor>) (1 / x) "  | 
|
1162  | 
proof cases  | 
|
1163  | 
assume nonneg: "0 \<le> p"  | 
|
1164  | 
hence "2 powr real (p) = floor (real ((2::int) ^ nat p)) * floor (1::real)"  | 
|
1165  | 
by (simp add: powr_int del: real_of_int_power) simp  | 
|
1166  | 
also have "floor (1::real) \<le> floor (1 / x)" using x prec by simp  | 
|
1167  | 
also have "floor (real ((2::int) ^ nat p)) * floor (1 / x) \<le>  | 
|
1168  | 
floor (real ((2::int) ^ nat p) * (1 / x))"  | 
|
1169  | 
by (rule le_mult_floor) (auto simp: x prec less_imp_le)  | 
|
1170  | 
finally have "2 powr real p \<le> floor (2 powr nat p / x)" by (simp add: powr_realpow)  | 
|
1171  | 
thus ?thesis unfolding p_def[symmetric]  | 
|
1172  | 
using x prec nonneg by (simp add: powr_minus inverse_eq_divide round_down_def)  | 
|
1173  | 
next  | 
|
1174  | 
assume neg: "\<not> 0 \<le> p"  | 
|
1175  | 
||
1176  | 
have "x = 2 powr (log 2 x)"  | 
|
1177  | 
using x by simp  | 
|
1178  | 
also have "2 powr (log 2 x) \<le> 2 powr p"  | 
|
1179  | 
proof (rule powr_mono)  | 
|
1180  | 
have "log 2 x \<le> \<lceil>log 2 x\<rceil>"  | 
|
1181  | 
by simp  | 
|
1182  | 
also have "\<dots> \<le> \<lfloor>log 2 x\<rfloor> + 1"  | 
|
1183  | 
using ceiling_diff_floor_le_1[of "log 2 x"] by simp  | 
|
1184  | 
also have "\<dots> \<le> \<lfloor>log 2 x\<rfloor> + prec"  | 
|
1185  | 
using prec by simp  | 
|
1186  | 
finally show "log 2 x \<le> real p"  | 
|
1187  | 
using x by (simp add: p_def)  | 
|
1188  | 
qed simp  | 
|
1189  | 
finally have x_le: "x \<le> 2 powr p" .  | 
|
1190  | 
||
1191  | 
from neg have "2 powr real p \<le> 2 powr 0"  | 
|
1192  | 
by (intro powr_mono) auto  | 
|
1193  | 
also have "\<dots> \<le> \<lfloor>2 powr 0\<rfloor>" by simp  | 
|
1194  | 
also have "\<dots> \<le> \<lfloor>2 powr real p / x\<rfloor>" unfolding real_of_int_le_iff  | 
|
1195  | 
using x x_le by (intro floor_mono) (simp add: pos_le_divide_eq mult_pos_pos)  | 
|
1196  | 
finally show ?thesis  | 
|
1197  | 
using prec x unfolding p_def[symmetric]  | 
|
1198  | 
by (simp add: round_down_def powr_minus_divide pos_le_divide_eq mult_pos_pos)  | 
|
1199  | 
qed  | 
|
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
1200  | 
qed  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
1201  | 
|
| 
31098
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
31021 
diff
changeset
 | 
1202  | 
lemma float_divr: "real x / real y \<le> real (float_divr prec x y)"  | 
| 47600 | 1203  | 
using round_up by transfer simp  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
1204  | 
|
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
1205  | 
lemma float_divr_pos_less1_lower_bound: assumes "0 < x" and "x < 1" shows "1 \<le> float_divr prec 1 x"  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
1206  | 
proof -  | 
| 47600 | 1207  | 
have "1 \<le> 1 / real x" using `0 < x` and `x < 1` by auto  | 
| 
31098
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
31021 
diff
changeset
 | 
1208  | 
also have "\<dots> \<le> real (float_divr prec 1 x)" using float_divr[where x=1 and y=x] by auto  | 
| 47600 | 1209  | 
finally show ?thesis by auto  | 
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
1210  | 
qed  | 
| 
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
1211  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1212  | 
lemma float_divr_nonpos_pos_upper_bound:  | 
| 47600 | 1213  | 
"real x \<le> 0 \<Longrightarrow> 0 < real y \<Longrightarrow> real (float_divr prec x y) \<le> 0"  | 
1214  | 
by transfer (auto simp: field_simps mult_le_0_iff divide_le_0_iff round_up_def)  | 
|
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
1215  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1216  | 
lemma float_divr_nonneg_neg_upper_bound:  | 
| 47600 | 1217  | 
"0 \<le> real x \<Longrightarrow> real y < 0 \<Longrightarrow> real (float_divr prec x y) \<le> 0"  | 
1218  | 
by transfer (auto simp: field_simps mult_le_0_iff zero_le_mult_iff divide_le_0_iff round_up_def)  | 
|
1219  | 
||
1220  | 
lift_definition float_round_up :: "nat \<Rightarrow> float \<Rightarrow> float" is  | 
|
1221  | 
"\<lambda>(prec::nat) x. round_up (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) x" by simp  | 
|
1222  | 
||
1223  | 
lemma float_round_up: "real x \<le> real (float_round_up prec x)"  | 
|
1224  | 
using round_up by transfer simp  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1225  | 
|
| 47600 | 1226  | 
lift_definition float_round_down :: "nat \<Rightarrow> float \<Rightarrow> float" is  | 
1227  | 
"\<lambda>(prec::nat) x. round_down (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) x" by simp  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1228  | 
|
| 47600 | 1229  | 
lemma float_round_down: "real (float_round_down prec x) \<le> real x"  | 
1230  | 
using round_down by transfer simp  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1231  | 
|
| 47600 | 1232  | 
lemma floor_add2[simp]: "\<lfloor> real i + x \<rfloor> = i + \<lfloor> x \<rfloor>"  | 
1233  | 
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
 | 
1234  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1235  | 
lemma compute_float_round_down[code]:  | 
| 47600 | 1236  | 
"float_round_down prec (Float m e) = (let d = bitlen (abs m) - int prec in  | 
1237  | 
if 0 < d then let P = 2^nat d ; n = m div P in Float n (e + d)  | 
|
1238  | 
else Float m e)"  | 
|
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
1239  | 
using Float.compute_float_down[of "prec - bitlen \<bar>m\<bar> - e" m e, symmetric]  | 
| 
47601
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
1240  | 
by transfer (simp add: field_simps abs_mult log_mult bitlen_def cong del: if_weak_cong)  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
1241  | 
hide_fact (open) compute_float_round_down  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1242  | 
|
| 47600 | 1243  | 
lemma compute_float_round_up[code]:  | 
1244  | 
"float_round_up prec (Float m e) = (let d = (bitlen (abs m) - int prec) in  | 
|
1245  | 
if 0 < d then let P = 2^nat d ; n = m div P ; r = m mod P  | 
|
1246  | 
in Float (n + (if r = 0 then 0 else 1)) (e + d)  | 
|
1247  | 
else Float m e)"  | 
|
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
1248  | 
using Float.compute_float_up[of "prec - bitlen \<bar>m\<bar> - e" m e, symmetric]  | 
| 47600 | 1249  | 
unfolding Let_def  | 
| 
47601
 
050718fe6eee
use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
 
hoelzl 
parents: 
47600 
diff
changeset
 | 
1250  | 
by transfer (simp add: field_simps abs_mult log_mult bitlen_def cong del: if_weak_cong)  | 
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
1251  | 
hide_fact (open) compute_float_round_up  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
1252  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1253  | 
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: 
47230 
diff
changeset
 | 
1254  | 
apply (auto simp: zero_float_def mult_le_0_iff)  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1255  | 
using powr_gt_zero[of 2 b] by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1256  | 
|
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
1257  | 
lemma real_of_float_pprt[simp]: fixes a::float shows "real (pprt a) = pprt (real a)"  | 
| 47600 | 1258  | 
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: 
47230 
diff
changeset
 | 
1259  | 
|
| 
47621
 
4cf6011fb884
hide code generation facts in the Float theory, they are only exported for Approximation
 
hoelzl 
parents: 
47615 
diff
changeset
 | 
1260  | 
lemma real_of_float_nprt[simp]: fixes a::float shows "real (nprt a) = nprt (real a)"  | 
| 47600 | 1261  | 
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: 
47230 
diff
changeset
 | 
1262  | 
|
| 47600 | 1263  | 
lift_definition int_floor_fl :: "float \<Rightarrow> int" is floor by simp  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
1264  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1265  | 
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: 
47600 
diff
changeset
 | 
1266  | 
"int_floor_fl (Float m e) = (if 0 \<le> e then m * 2 ^ nat e else m div (2 ^ (nat (-e))))"  | 
| 47600 | 1267  | 
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: 
47615 
diff
changeset
 | 
1268  | 
hide_fact (open) compute_int_floor_fl  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1269  | 
|
| 47600 | 1270  | 
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: 
47230 
diff
changeset
 | 
1271  | 
|
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1272  | 
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: 
47600 
diff
changeset
 | 
1273  | 
"floor_fl (Float m e) = (if 0 \<le> e then Float m e else Float (m div (2 ^ (nat (-e)))) 0)"  | 
| 47600 | 1274  | 
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: 
47615 
diff
changeset
 | 
1275  | 
hide_fact (open) compute_floor_fl  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
1276  | 
|
| 47600 | 1277  | 
lemma floor_fl: "real (floor_fl x) \<le> real x" by transfer simp  | 
1278  | 
||
1279  | 
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: 
29667 
diff
changeset
 | 
1280  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1281  | 
lemma floor_pos_exp: "exponent (floor_fl x) \<ge> 0"  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1282  | 
proof cases  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1283  | 
assume nzero: "floor_fl x \<noteq> float_of 0"  | 
| 47600 | 1284  | 
have "floor_fl x = Float \<lfloor>real x\<rfloor> 0" by transfer simp  | 
1285  | 
from denormalize_shift[OF this[THEN eq_reflection] nzero] guess i . note i = this  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1286  | 
thus ?thesis by simp  | 
| 
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1287  | 
qed (simp add: floor_fl_def)  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
1288  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
1289  | 
end  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47230 
diff
changeset
 | 
1290  |