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