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