--- a/src/HOL/Decision_Procs/Approximation.thy Wed Apr 18 14:29:20 2012 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy Wed Apr 18 14:29:21 2012 +0200
@@ -54,23 +54,9 @@
case 0 thus ?case unfolding lb_0 ub_0 horner.simps by auto
next
case (Suc n)
- have "?lb (Suc n) j' \<le> ?horner (Suc n) j'" unfolding lb_Suc ub_Suc horner.simps real_of_float_sub diff_minus
- proof (rule add_mono)
- show "(lapprox_rat prec 1 (f j')) \<le> 1 / (f j')" using lapprox_rat[of prec 1 "f j'"] by auto
- from Suc[where j'="Suc j'", unfolded funpow.simps comp_def f_Suc, THEN conjunct2] `0 \<le> real x`
- show "- real (x * ub n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) x) \<le>
- - (x * horner F G n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) x)"
- unfolding real_of_float_mult neg_le_iff_le by (rule mult_left_mono)
- qed
- moreover have "?horner (Suc n) j' \<le> ?ub (Suc n) j'" unfolding ub_Suc horner.simps real_of_float_sub diff_minus
- proof (rule add_mono)
- show "1 / (f j') \<le> (rapprox_rat prec 1 (f j'))" using rapprox_rat[of 1 "f j'" prec] by auto
- from Suc[where j'="Suc j'", unfolded funpow.simps comp_def f_Suc, THEN conjunct1] `0 \<le> real x`
- show "- (x * horner F G n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) x) \<le>
- - real (x * lb n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) x)"
- unfolding real_of_float_mult neg_le_iff_le by (rule mult_left_mono)
- qed
- ultimately show ?case by blast
+ thus ?case using lapprox_rat[of prec 1 "f j'"] using rapprox_rat[of 1 "f j'" prec]
+ Suc[where j'="Suc j'"] `0 \<le> real x`
+ by (auto intro!: add_mono mult_left_mono simp add: lb_Suc ub_Suc field_simps f_Suc)
qed
subsection "Theorems for floating point functions implementing the horner scheme"
@@ -106,24 +92,10 @@
shows "(lb n ((F ^^ j') s) (f j') x) \<le> (\<Sum>j=0..<n. (1 / (f (j' + j))) * real x ^ j)" (is "?lb") and
"(\<Sum>j=0..<n. (1 / (f (j' + j))) * real x ^ j) \<le> (ub n ((F ^^ j') s) (f j') x)" (is "?ub")
proof -
- { fix x y z :: float have "x - y * z = x + - y * z"
- by (cases x, cases y, cases z, simp add: plus_float.simps minus_float_def times_float.simps algebra_simps)
- } note diff_mult_minus = this
-
- { fix x :: float have "- (- x) = x" by (cases x) auto } note minus_minus = this
-
- have move_minus: "(-x) = -1 * real x" by auto (* coercion "inside" is necessary *)
-
+ { fix x y z :: float have "x - y * z = x + - y * z" by simp } note diff_mult_minus = this
have sum_eq: "(\<Sum>j=0..<n. (1 / (f (j' + j))) * real x ^ j) =
(\<Sum>j = 0..<n. -1 ^ j * (1 / (f (j' + j))) * real (- x) ^ j)"
- proof (rule setsum_cong, simp)
- fix j assume "j \<in> {0 ..< n}"
- show "1 / (f (j' + j)) * real x ^ j = -1 ^ j * (1 / (f (j' + j))) * real (- x) ^ j"
- unfolding move_minus power_mult_distrib mult_assoc[symmetric]
- unfolding mult_commute unfolding mult_assoc[of "-1 ^ j", symmetric] power_mult_distrib[symmetric]
- by auto
- qed
-
+ by (auto simp add: field_simps power_mult_distrib[symmetric])
have "0 \<le> real (-x)" using assms by auto
from horner_bounds[where G=G and F=F and f=f and s=s and prec=prec
and lb="\<lambda> n i k x. lb n i k (-x)" and ub="\<lambda> n i k x. ub n i k (-x)", unfolded lb_Suc ub_Suc diff_mult_minus,
@@ -177,9 +149,11 @@
show ?thesis
proof (cases "0 < l")
case True hence "odd n \<or> 0 < l" and "0 \<le> real l" unfolding less_float_def by auto
- have u1: "u1 = u ^ n" and l1: "l1 = l ^ n" using assms unfolding float_power_bnds_def if_P[OF `odd n \<or> 0 < l`] by auto
- have "real l ^ n \<le> x ^ n" and "x ^ n \<le> real u ^ n " using `0 \<le> real l` and assms unfolding atLeastAtMost_iff using power_mono[of l x] power_mono[of x u] by auto
- thus ?thesis using assms `0 < l` unfolding atLeastAtMost_iff l1 u1 float_power less_float_def by auto
+ have u1: "u1 = u ^ n" and l1: "l1 = l ^ n" using assms
+ unfolding float_power_bnds_def if_P[OF `odd n \<or> 0 < l`] by auto
+ have "real l ^ n \<le> x ^ n" and "x ^ n \<le> real u ^ n " using `0 \<le> real l` assms
+ by (auto simp: power_mono)
+ thus ?thesis using assms `0 < l` unfolding l1 u1 by auto
next
case False hence P: "\<not> (odd n \<or> 0 < l)" using `even n` by auto
show ?thesis
@@ -188,25 +162,25 @@
hence "real u ^ n \<le> x ^ n" and "x ^ n \<le> real l ^ n" using power_mono[of "-x" "-real l" n] power_mono[of "-real u" "-x" n]
unfolding power_minus_even[OF `even n`] by auto
moreover have u1: "u1 = l ^ n" and l1: "l1 = u ^ n" using assms unfolding float_power_bnds_def if_not_P[OF P] if_P[OF True] by auto
- ultimately show ?thesis using float_power by auto
+ ultimately show ?thesis by auto
next
case False
have "\<bar>x\<bar> \<le> real (max (-l) u)"
proof (cases "-l \<le> u")
- case True thus ?thesis unfolding max_def if_P[OF True] using assms unfolding le_float_def by auto
+ case True thus ?thesis unfolding max_def if_P[OF True] using assms unfolding less_eq_float_def by auto
next
- case False thus ?thesis unfolding max_def if_not_P[OF False] using assms unfolding le_float_def by auto
+ case False thus ?thesis unfolding max_def if_not_P[OF False] using assms unfolding less_eq_float_def by auto
qed
hence x_abs: "\<bar>x\<bar> \<le> \<bar>real (max (-l) u)\<bar>" by auto
have u1: "u1 = (max (-l) u) ^ n" and l1: "l1 = 0" using assms unfolding float_power_bnds_def if_not_P[OF P] if_not_P[OF False] by auto
- show ?thesis unfolding atLeastAtMost_iff l1 u1 float_power using zero_le_even_power[OF `even n`] power_mono_even[OF `even n` x_abs] by auto
+ show ?thesis unfolding atLeastAtMost_iff l1 u1 using zero_le_even_power[OF `even n`] power_mono_even[OF `even n` x_abs] by auto
qed
qed
next
case False hence "odd n \<or> 0 < l" by auto
have u1: "u1 = u ^ n" and l1: "l1 = l ^ n" using assms unfolding float_power_bnds_def if_P[OF `odd n \<or> 0 < l`] by auto
have "real l ^ n \<le> x ^ n" and "x ^ n \<le> real u ^ n " using assms unfolding atLeastAtMost_iff using power_mono_odd[OF False] by auto
- thus ?thesis unfolding atLeastAtMost_iff l1 u1 float_power less_float_def by auto
+ thus ?thesis unfolding atLeastAtMost_iff l1 u1 less_float_def by auto
qed
lemma bnds_power: "\<forall> (x::real) l u. (l1, u1) = float_power_bnds n l u \<and> x \<in> {l .. u} \<longrightarrow> l1 \<le> x ^ n \<and> x ^ n \<le> u1"
@@ -222,10 +196,17 @@
*}
fun sqrt_iteration :: "nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" where
-"sqrt_iteration prec 0 (Float m e) = Float 1 ((e + bitlen m) div 2 + 1)" |
+"sqrt_iteration prec 0 x = Float 1 ((bitlen \<bar>mantissa x\<bar> + exponent x) div 2 + 1)" |
"sqrt_iteration prec (Suc m) x = (let y = sqrt_iteration prec m x
in Float 1 -1 * (y + float_divr prec x y))"
+lemma compute_sqrt_iteration_base[code]:
+ shows "sqrt_iteration prec n (Float m e) =
+ (if n = 0 then Float 1 ((if m = 0 then 0 else bitlen \<bar>m\<bar> + e) div 2 + 1)
+ else (let y = sqrt_iteration prec (n - 1) (Float m e) in
+ Float 1 -1 * (y + float_divr prec (Float m e) y)))"
+ using bitlen_Float by (cases n) simp_all
+
function ub_sqrt lb_sqrt :: "nat \<Rightarrow> float \<Rightarrow> float" where
"ub_sqrt prec x = (if 0 < x then (sqrt_iteration prec prec x)
else if x < 0 then - lb_sqrt prec (- x)
@@ -259,23 +240,23 @@
show ?case
proof (cases x)
case (Float m e)
- hence "0 < m" using float_pos_m_pos[unfolded less_float_def] assms by auto
+ hence "0 < m" using assms powr_gt_zero[of 2 e] by (auto simp: sign_simps)
hence "0 < sqrt m" by auto
- have int_nat_bl: "(nat (bitlen m)) = bitlen m" using bitlen_ge0 by auto
-
- have "x = (m / 2^nat (bitlen m)) * pow2 (e + (nat (bitlen m)))"
- unfolding pow2_add pow2_int Float real_of_float_simp by auto
- also have "\<dots> < 1 * pow2 (e + nat (bitlen m))"
+ have int_nat_bl: "(nat (bitlen m)) = bitlen m" using bitlen_nonneg by auto
+
+ have "x = (m / 2^nat (bitlen m)) * 2 powr (e + (nat (bitlen m)))"
+ unfolding Float by (auto simp: powr_realpow[symmetric] field_simps powr_add)
+ also have "\<dots> < 1 * 2 powr (e + nat (bitlen m))"
proof (rule mult_strict_right_mono, auto)
show "real m < 2^nat (bitlen m)" using bitlen_bounds[OF `0 < m`, THEN conjunct2]
unfolding real_of_int_less_iff[of m, symmetric] by auto
qed
- finally have "sqrt x < sqrt (pow2 (e + bitlen m))" unfolding int_nat_bl by auto
- also have "\<dots> \<le> pow2 ((e + bitlen m) div 2 + 1)"
+ finally have "sqrt x < sqrt (2 powr (e + bitlen m))" unfolding int_nat_bl by auto
+ also have "\<dots> \<le> 2 powr ((e + bitlen m) div 2 + 1)"
proof -
let ?E = "e + bitlen m"
- have E_mod_pow: "pow2 (?E mod 2) < 4"
+ have E_mod_pow: "2 powr (?E mod 2) < 4"
proof (cases "?E mod 2 = 1")
case True thus ?thesis by auto
next
@@ -287,21 +268,23 @@
from xt1(5)[OF `0 \<le> ?E mod 2` this]
show ?thesis by auto
qed
- hence "sqrt (pow2 (?E mod 2)) < sqrt (2 * 2)" by auto
- hence E_mod_pow: "sqrt (pow2 (?E mod 2)) < 2" unfolding real_sqrt_abs2 by auto
-
- have E_eq: "pow2 ?E = pow2 (?E div 2 + ?E div 2 + ?E mod 2)" by auto
- have "sqrt (pow2 ?E) = sqrt (pow2 (?E div 2) * pow2 (?E div 2) * pow2 (?E mod 2))"
- unfolding E_eq unfolding pow2_add ..
- also have "\<dots> = pow2 (?E div 2) * sqrt (pow2 (?E mod 2))"
- unfolding real_sqrt_mult[of _ "pow2 (?E mod 2)"] real_sqrt_abs2 by auto
- also have "\<dots> < pow2 (?E div 2) * 2"
+ hence "sqrt (2 powr (?E mod 2)) < sqrt (2 * 2)" by auto
+ hence E_mod_pow: "sqrt (2 powr (?E mod 2)) < 2" unfolding real_sqrt_abs2 by auto
+
+ have E_eq: "2 powr ?E = 2 powr (?E div 2 + ?E div 2 + ?E mod 2)" by auto
+ have "sqrt (2 powr ?E) = sqrt (2 powr (?E div 2) * 2 powr (?E div 2) * 2 powr (?E mod 2))"
+ unfolding E_eq unfolding powr_add[symmetric] by (simp add: int_of_reals del: real_of_ints)
+ also have "\<dots> = 2 powr (?E div 2) * sqrt (2 powr (?E mod 2))"
+ unfolding real_sqrt_mult[of _ "2 powr (?E mod 2)"] real_sqrt_abs2 by auto
+ also have "\<dots> < 2 powr (?E div 2) * 2 powr 1"
by (rule mult_strict_left_mono, auto intro: E_mod_pow)
- also have "\<dots> = pow2 (?E div 2 + 1)" unfolding add_commute[of _ 1] pow2_add1 by auto
+ also have "\<dots> = 2 powr (?E div 2 + 1)" unfolding add_commute[of _ 1] powr_add[symmetric]
+ by simp
finally show ?thesis by auto
qed
- finally show ?thesis
- unfolding Float sqrt_iteration.simps real_of_float_simp by auto
+ finally show ?thesis using `0 < m`
+ unfolding Float
+ by (subst compute_sqrt_iteration_base) (simp add: ac_simps real_Float del: Float_def)
qed
next
case (Suc n)
@@ -310,8 +293,8 @@
also have "\<dots> < real ?b" using Suc .
finally have "sqrt x < (?b + x / ?b)/2" using sqrt_ub_pos_pos_1[OF Suc _ `0 < real x`] by auto
also have "\<dots> \<le> (?b + (float_divr prec x ?b))/2" by (rule divide_right_mono, auto simp add: float_divr)
- also have "\<dots> = (Float 1 -1) * (?b + (float_divr prec x ?b))" by auto
- finally show ?case unfolding sqrt_iteration.simps Let_def real_of_float_mult real_of_float_add right_distrib .
+ also have "\<dots> = (Float 1 -1) * (?b + (float_divr prec x ?b))" by simp
+ finally show ?case unfolding sqrt_iteration.simps Let_def right_distrib .
qed
lemma sqrt_iteration_lower_bound: assumes "0 < real x"
@@ -325,9 +308,9 @@
lemma lb_sqrt_lower_bound: assumes "0 \<le> real x"
shows "0 \<le> real (lb_sqrt prec x)"
proof (cases "0 < x")
- case True hence "0 < real x" and "0 \<le> x" using `0 \<le> real x` unfolding less_float_def le_float_def by auto
+ case True hence "0 < real x" and "0 \<le> x" using `0 \<le> real x` unfolding less_float_def less_eq_float_def by auto
hence "0 < sqrt_iteration prec prec x" unfolding less_float_def using sqrt_iteration_lower_bound by auto
- hence "0 \<le> real (float_divl prec x (sqrt_iteration prec prec x))" using float_divl_lower_bound[OF `0 \<le> x`] unfolding le_float_def by auto
+ hence "0 \<le> real (float_divl prec x (sqrt_iteration prec prec x))" using float_divl_lower_bound[OF `0 \<le> x`] unfolding less_eq_float_def by auto
thus ?thesis unfolding lb_sqrt.simps using True by auto
next
case False with `0 \<le> real x` have "real x = 0" unfolding less_float_def by auto
@@ -446,7 +429,7 @@
{ have "(x * lb_arctan_horner prec n 1 (x*x)) \<le> ?S n"
using bounds(1) `0 \<le> real x`
- unfolding real_of_float_mult power_add power_one_right mult_assoc[symmetric] setsum_left_distrib[symmetric]
+ unfolding power_add power_one_right mult_assoc[symmetric] setsum_left_distrib[symmetric]
unfolding mult_commute[where 'a=real] mult_commute[of _ "2::nat"] power_mult power2_eq_square[of "real x"]
by (auto intro!: mult_left_mono)
also have "\<dots> \<le> arctan x" using arctan_bounds ..
@@ -455,7 +438,7 @@
{ have "arctan x \<le> ?S (Suc n)" using arctan_bounds ..
also have "\<dots> \<le> (x * ub_arctan_horner prec (Suc n) 1 (x*x))"
using bounds(2)[of "Suc n"] `0 \<le> real x`
- unfolding real_of_float_mult power_add power_one_right mult_assoc[symmetric] setsum_left_distrib[symmetric]
+ unfolding power_add power_one_right mult_assoc[symmetric] setsum_left_distrib[symmetric]
unfolding mult_commute[where 'a=real] mult_commute[of _ "2::nat"] power_mult power2_eq_square[of "real x"]
by (auto intro!: mult_left_mono)
finally have "arctan x \<le> (x * ub_arctan_horner prec (Suc n) 1 (x*x))" . }
@@ -512,8 +495,8 @@
have "1 div k = 0" using div_pos_pos_trivial[OF _ `1 < k`] by auto
have "0 \<le> real ?k" by (rule order_trans[OF _ rapprox_rat], auto simp add: `0 \<le> k`)
- have "real ?k \<le> 1" unfolding rapprox_rat.simps(2)[OF zero_le_one `0 < k`]
- by (rule rapprox_posrat_le1, auto simp add: `0 < k` `1 \<le> k`)
+ have "real ?k \<le> 1"
+ by (rule rapprox_rat_le1, auto simp add: `0 < k` `1 \<le> k`)
have "1 / k \<le> ?k" using rapprox_rat[where x=1 and y=k] by auto
hence "arctan (1 / k) \<le> arctan ?k" by (rule arctan_monotone')
@@ -526,8 +509,7 @@
let ?k = "lapprox_rat prec 1 k"
have "1 div k = 0" using div_pos_pos_trivial[OF _ `1 < k`] by auto
have "1 / k \<le> 1" using `1 < k` by auto
-
- have "\<And>n. 0 \<le> real ?k" using lapprox_rat_bottom[where x=1 and y=k, OF zero_le_one `0 < k`] by (auto simp add: `1 div k = 0`)
+ have "\<And>n. 0 \<le> real ?k" using lapprox_rat_nonneg[where x=1 and y=k, OF zero_le_one `0 < k`] by (auto simp add: `1 div k = 0`)
have "\<And>n. real ?k \<le> 1" using lapprox_rat by (rule order_trans, auto simp add: `1 / k \<le> 1`)
have "?k \<le> 1 / k" using lapprox_rat[where x=1 and y=k] by auto
@@ -539,14 +521,14 @@
} note lb_arctan = this
have "pi \<le> ub_pi n"
- unfolding ub_pi_def machin_pi Let_def real_of_float_mult real_of_float_sub unfolding Float_num
- using lb_arctan[of 239] ub_arctan[of 5]
- by (auto intro!: mult_left_mono add_mono simp add: diff_minus simp del: lapprox_rat.simps rapprox_rat.simps)
+ unfolding ub_pi_def machin_pi Let_def unfolding Float_num
+ using lb_arctan[of 239] ub_arctan[of 5] powr_realpow[of 2 2]
+ by (auto intro!: mult_left_mono add_mono simp add: diff_minus)
moreover
have "lb_pi n \<le> pi"
- unfolding lb_pi_def machin_pi Let_def real_of_float_mult real_of_float_sub Float_num
- using lb_arctan[of 5] ub_arctan[of 239]
- by (auto intro!: mult_left_mono add_mono simp add: diff_minus simp del: lapprox_rat.simps rapprox_rat.simps)
+ unfolding lb_pi_def machin_pi Let_def Float_num
+ using lb_arctan[of 5] ub_arctan[of 239] powr_realpow[of 2 2]
+ by (auto intro!: mult_left_mono add_mono simp add: diff_minus)
ultimately show ?thesis by auto
qed
@@ -579,17 +561,17 @@
lemma lb_arctan_bound': assumes "0 \<le> real x"
shows "lb_arctan prec x \<le> arctan x"
proof -
- have "\<not> x < 0" and "0 \<le> x" unfolding less_float_def le_float_def using `0 \<le> real x` by auto
+ have "\<not> x < 0" and "0 \<le> x" unfolding less_float_def less_eq_float_def using `0 \<le> real x` by auto
let "?ub_horner x" = "x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (x * x)"
and "?lb_horner x" = "x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (x * x)"
show ?thesis
proof (cases "x \<le> Float 1 -1")
- case True hence "real x \<le> 1" unfolding le_float_def Float_num by auto
+ case True hence "real x \<le> 1" unfolding less_eq_float_def Float_num by auto
show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_P[OF True]
using arctan_0_1_bounds[OF `0 \<le> real x` `real x \<le> 1`] by auto
next
- case False hence "0 < real x" unfolding le_float_def Float_num by auto
+ case False hence "0 < real x" unfolding less_eq_float_def Float_num by auto
let ?R = "1 + sqrt (1 + real x * real x)"
let ?fR = "1 + ub_sqrt prec (1 + x * x)"
let ?DIV = "float_divl prec x ?fR"
@@ -621,9 +603,9 @@
moreover have "?DIV \<le> real x / ?fR" by (rule float_divl)
ultimately have "real ?DIV \<le> 1" unfolding divide_le_eq_1_pos[OF `0 < real ?fR`, symmetric] by auto
- have "0 \<le> real ?DIV" using float_divl_lower_bound[OF `0 \<le> x` `0 < ?fR`] unfolding le_float_def by auto
-
- have "(Float 1 1 * ?lb_horner ?DIV) \<le> 2 * arctan (float_divl prec x ?fR)" unfolding real_of_float_mult[of "Float 1 1"] Float_num
+ have "0 \<le> real ?DIV" using float_divl_lower_bound[OF `0 \<le> x` `0 < ?fR`] unfolding less_eq_float_def by auto
+
+ have "(Float 1 1 * ?lb_horner ?DIV) \<le> 2 * arctan (float_divl prec x ?fR)"
using arctan_0_1_bounds[OF `0 \<le> real ?DIV` `real ?DIV \<le> 1`] by auto
also have "\<dots> \<le> 2 * arctan (x / ?R)"
using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono)
@@ -631,7 +613,7 @@
finally show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_P[OF True] .
next
case False
- hence "2 < real x" unfolding le_float_def Float_num by auto
+ hence "2 < real x" unfolding less_eq_float_def Float_num by auto
hence "1 \<le> real x" by auto
let "?invx" = "float_divr prec 1 x"
@@ -649,13 +631,14 @@
have "1 / x \<noteq> 0" and "0 < 1 / x" using `0 < real x` by auto
- have "arctan (1 / x) \<le> arctan ?invx" unfolding real_of_float_1[symmetric] by (rule arctan_monotone', rule float_divr)
+ have "arctan (1 / x) \<le> arctan ?invx" unfolding real_of_float_one[symmetric] by (rule arctan_monotone', rule float_divr)
also have "\<dots> \<le> (?ub_horner ?invx)" using arctan_0_1_bounds[OF `0 \<le> real ?invx` `real ?invx \<le> 1`] by auto
finally have "pi / 2 - (?ub_horner ?invx) \<le> arctan x"
using `0 \<le> arctan x` arctan_inverse[OF `1 / x \<noteq> 0`]
unfolding real_sgn_pos[OF `0 < 1 / real x`] le_diff_eq by auto
moreover
- have "lb_pi prec * Float 1 -1 \<le> pi / 2" unfolding real_of_float_mult Float_num times_divide_eq_right mult_1_left using pi_boundaries by auto
+ have "lb_pi prec * Float 1 -1 \<le> pi / 2"
+ unfolding Float_num times_divide_eq_right mult_1_left using pi_boundaries by simp
ultimately
show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_not_P[OF `\<not> x \<le> Float 1 1`] if_not_P[OF False]
by auto
@@ -667,18 +650,18 @@
lemma ub_arctan_bound': assumes "0 \<le> real x"
shows "arctan x \<le> ub_arctan prec x"
proof -
- have "\<not> x < 0" and "0 \<le> x" unfolding less_float_def le_float_def using `0 \<le> real x` by auto
+ have "\<not> x < 0" and "0 \<le> x" unfolding less_float_def less_eq_float_def using `0 \<le> real x` by auto
let "?ub_horner x" = "x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (x * x)"
and "?lb_horner x" = "x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (x * x)"
show ?thesis
proof (cases "x \<le> Float 1 -1")
- case True hence "real x \<le> 1" unfolding le_float_def Float_num by auto
+ case True hence "real x \<le> 1" unfolding less_eq_float_def Float_num by auto
show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_P[OF True]
using arctan_0_1_bounds[OF `0 \<le> real x` `real x \<le> 1`] by auto
next
- case False hence "0 < real x" unfolding le_float_def Float_num by auto
+ case False hence "0 < real x" unfolding less_eq_float_def Float_num by auto
let ?R = "1 + sqrt (1 + real x * real x)"
let ?fR = "1 + lb_sqrt prec (1 + x * x)"
let ?DIV = "float_divr prec x ?fR"
@@ -691,7 +674,7 @@
have "lb_sqrt prec (1 + x * x) \<le> sqrt (1 + x * x)"
using bnds_sqrt'[of "1 + x * x"] by auto
hence "?fR \<le> ?R" by auto
- have "0 < real ?fR" unfolding real_of_float_add real_of_float_1 by (rule order_less_le_trans[OF zero_less_one], auto simp add: lb_sqrt_lower_bound[OF `0 \<le> real (1 + x*x)`])
+ have "0 < real ?fR" by (rule order_less_le_trans[OF zero_less_one], auto simp add: lb_sqrt_lower_bound[OF `0 \<le> real (1 + x*x)`])
have monotone: "x / ?R \<le> (float_divr prec x ?fR)"
proof -
@@ -707,7 +690,7 @@
show ?thesis
proof (cases "?DIV > 1")
case True
- have "pi / 2 \<le> ub_pi prec * Float 1 -1" unfolding real_of_float_mult Float_num times_divide_eq_right mult_1_left using pi_boundaries by auto
+ have "pi / 2 \<le> ub_pi prec * Float 1 -1" unfolding Float_num times_divide_eq_right mult_1_left using pi_boundaries by auto
from order_less_le_trans[OF arctan_ubound this, THEN less_imp_le]
show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_P[OF `x \<le> Float 1 1`] if_P[OF True] .
next
@@ -720,13 +703,13 @@
have "arctan x = 2 * arctan (x / ?R)" using arctan_half unfolding numeral_2_eq_2 power_Suc2 power_0 mult_1_left .
also have "\<dots> \<le> 2 * arctan (?DIV)"
using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono)
- also have "\<dots> \<le> (Float 1 1 * ?ub_horner ?DIV)" unfolding real_of_float_mult[of "Float 1 1"] Float_num
+ also have "\<dots> \<le> (Float 1 1 * ?ub_horner ?DIV)" unfolding Float_num
using arctan_0_1_bounds[OF `0 \<le> real ?DIV` `real ?DIV \<le> 1`] by auto
finally show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x \<le> Float 1 -1`] if_P[OF `x \<le> Float 1 1`] if_not_P[OF False] .
qed
next
case False
- hence "2 < real x" unfolding le_float_def Float_num by auto
+ hence "2 < real x" unfolding less_eq_float_def Float_num by auto
hence "1 \<le> real x" by auto
hence "0 < real x" by auto
hence "0 < x" unfolding less_float_def by auto
@@ -735,17 +718,17 @@
have "0 \<le> arctan x" using arctan_monotone'[OF `0 \<le> real x`] using arctan_tan[of 0, unfolded tan_zero] by auto
have "real ?invx \<le> 1" unfolding less_float_def by (rule order_trans[OF float_divl], auto simp add: `1 \<le> real x` divide_le_eq_1_pos[OF `0 < real x`])
- have "0 \<le> real ?invx" unfolding real_of_float_0[symmetric] by (rule float_divl_lower_bound[unfolded le_float_def], auto simp add: `0 < x`)
+ have "0 \<le> real ?invx" by (rule float_divl_lower_bound[unfolded less_eq_float_def], auto simp add: `0 < x`)
have "1 / x \<noteq> 0" and "0 < 1 / x" using `0 < real x` by auto
have "(?lb_horner ?invx) \<le> arctan (?invx)" using arctan_0_1_bounds[OF `0 \<le> real ?invx` `real ?invx \<le> 1`] by auto
- also have "\<dots> \<le> arctan (1 / x)" unfolding real_of_float_1[symmetric] by (rule arctan_monotone', rule float_divl)
+ also have "\<dots> \<le> arctan (1 / x)" unfolding real_of_float_one[symmetric] by (rule arctan_monotone', rule float_divl)
finally have "arctan x \<le> pi / 2 - (?lb_horner ?invx)"
using `0 \<le> arctan x` arctan_inverse[OF `1 / x \<noteq> 0`]
unfolding real_sgn_pos[OF `0 < 1 / x`] le_diff_eq by auto
moreover
- have "pi / 2 \<le> ub_pi prec * Float 1 -1" unfolding real_of_float_mult Float_num times_divide_eq_right mult_1_right using pi_boundaries by auto
+ have "pi / 2 \<le> ub_pi prec * Float 1 -1" unfolding Float_num times_divide_eq_right mult_1_right using pi_boundaries by auto
ultimately
show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF `\<not> x < 0`]if_not_P[OF `\<not> x \<le> Float 1 -1`] if_not_P[OF False]
by auto
@@ -756,15 +739,16 @@
lemma arctan_boundaries:
"arctan x \<in> {(lb_arctan prec x) .. (ub_arctan prec x)}"
proof (cases "0 \<le> x")
- case True hence "0 \<le> real x" unfolding le_float_def by auto
+ case True hence "0 \<le> real x" unfolding less_eq_float_def by auto
show ?thesis using ub_arctan_bound'[OF `0 \<le> real x`] lb_arctan_bound'[OF `0 \<le> real x`] unfolding atLeastAtMost_iff by auto
next
let ?mx = "-x"
- case False hence "x < 0" and "0 \<le> real ?mx" unfolding le_float_def less_float_def by auto
+ case False hence "x < 0" and "0 \<le> real ?mx" unfolding less_eq_float_def less_float_def by auto
hence bounds: "lb_arctan prec ?mx \<le> arctan ?mx \<and> arctan ?mx \<le> ub_arctan prec ?mx"
using ub_arctan_bound'[OF `0 \<le> real ?mx`] lb_arctan_bound'[OF `0 \<le> real ?mx`] by auto
show ?thesis unfolding real_of_float_minus arctan_minus lb_arctan.simps[where x=x] ub_arctan.simps[where x=x] Let_def if_P[OF `x < 0`]
- unfolding atLeastAtMost_iff using bounds[unfolded real_of_float_minus arctan_minus] by auto
+ unfolding atLeastAtMost_iff using bounds[unfolded real_of_float_minus arctan_minus]
+ by (simp add: arctan_minus)
qed
lemma bnds_arctan: "\<forall> (x::real) lx ux. (l, u) = (lb_arctan prec lx, ub_arctan prec ux) \<and> x \<in> {lx .. ux} \<longrightarrow> l \<le> arctan x \<and> arctan x \<le> u"
@@ -800,7 +784,7 @@
shows "(lb_sin_cos_aux prec n 1 1 (x * x)) \<le> (\<Sum> i=0..<n. -1^i * (1/real (fact (2 * i))) * x ^(2 * i))" (is "?lb")
and "(\<Sum> i=0..<n. -1^i * (1/real (fact (2 * i))) * x^(2 * i)) \<le> (ub_sin_cos_aux prec n 1 1 (x * x))" (is "?ub")
proof -
- have "0 \<le> real (x * x)" unfolding real_of_float_mult by auto
+ have "0 \<le> real (x * x)" by auto
let "?f n" = "fact (2 * n)"
{ fix n
@@ -818,7 +802,7 @@
proof (cases "real x = 0")
case False hence "real x \<noteq> 0" by auto
hence "0 < x" and "0 < real x" using `0 \<le> real x` unfolding less_float_def by auto
- have "0 < x * x" using `0 < x` unfolding less_float_def real_of_float_mult real_of_float_0
+ have "0 < x * x" using `0 < x` unfolding less_float_def real_of_float_zero
using mult_pos_pos[where a="real x" and b="real x"] by auto
{ fix x n have "(\<Sum> i=0..<n. -1^i * (1/real (fact (2 * i))) * x ^ (2 * i))
@@ -893,7 +877,7 @@
hence "get_even n = 0" by auto
have "- (pi / 2) \<le> x" by (rule order_trans[OF _ `0 < real x`[THEN less_imp_le]], auto)
with `x \<le> pi / 2`
- show ?thesis unfolding `get_even n = 0` lb_sin_cos_aux.simps real_of_float_minus real_of_float_0 using cos_ge_zero by auto
+ show ?thesis unfolding `get_even n = 0` lb_sin_cos_aux.simps real_of_float_minus real_of_float_zero using cos_ge_zero by auto
qed
ultimately show ?thesis by auto
next
@@ -901,7 +885,9 @@
show ?thesis
proof (cases "n = 0")
case True
- thus ?thesis unfolding `n = 0` get_even_def get_odd_def using `real x = 0` lapprox_rat[where x="-1" and y=1] by auto
+ thus ?thesis unfolding `n = 0` get_even_def get_odd_def
+ using `real x = 0` lapprox_rat[where x="-1" and y=1]
+ by (auto simp: compute_lapprox_rat compute_rapprox_rat)
next
case False with not0_implies_Suc obtain m where "n = Suc m" by blast
thus ?thesis unfolding `n = Suc m` get_even_def get_odd_def using `real x = 0` rapprox_rat[where x=1 and y=1] lapprox_rat[where x=1 and y=1] by (cases "even (Suc m)", auto)
@@ -912,7 +898,7 @@
shows "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \<le> (\<Sum> i=0..<n. -1^i * (1/real (fact (2 * i + 1))) * x^(2 * i + 1))" (is "?lb")
and "(\<Sum> i=0..<n. -1^i * (1/real (fact (2 * i + 1))) * x^(2 * i + 1)) \<le> (x * ub_sin_cos_aux prec n 2 1 (x * x))" (is "?ub")
proof -
- have "0 \<le> real (x * x)" unfolding real_of_float_mult by auto
+ have "0 \<le> real (x * x)" by auto
let "?f n" = "fact (2 * n + 1)"
{ fix n
@@ -922,7 +908,7 @@
from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0,
OF `0 \<le> real (x * x)` f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps]
- show "?lb" and "?ub" using `0 \<le> real x` unfolding real_of_float_mult
+ show "?lb" and "?ub" using `0 \<le> real x`
unfolding power_add power_one_right mult_assoc[symmetric] setsum_left_distrib[symmetric]
unfolding mult_commute[where 'a=real]
by (auto intro!: mult_left_mono simp add: power_mult power2_eq_square[of "real x"])
@@ -933,7 +919,7 @@
proof (cases "real x = 0")
case False hence "real x \<noteq> 0" by auto
hence "0 < x" and "0 < real x" using `0 \<le> real x` unfolding less_float_def by auto
- have "0 < x * x" using `0 < x` unfolding less_float_def real_of_float_mult real_of_float_0
+ have "0 < x * x" using `0 < x` unfolding less_float_def real_of_float_zero
using mult_pos_pos[where a="real x" and b="real x"] by auto
{ fix x n have "(\<Sum> j = 0 ..< n. -1 ^ (((2 * j + 1) - Suc 0) div 2) / (real (fact (2 * j + 1))) * x ^(2 * j + 1))
@@ -1006,7 +992,7 @@
case False
hence "get_even n = 0" by auto
with `x \<le> pi / 2` `0 \<le> real x`
- show ?thesis unfolding `get_even n = 0` ub_sin_cos_aux.simps real_of_float_minus real_of_float_0 using sin_ge_zero by auto
+ show ?thesis unfolding `get_even n = 0` ub_sin_cos_aux.simps real_of_float_minus using sin_ge_zero by auto
qed
ultimately show ?thesis by auto
next
@@ -1065,7 +1051,7 @@
case False
{ fix y x :: float let ?x2 = "(x * Float 1 -1)"
assume "y \<le> cos ?x2" and "-pi \<le> x" and "x \<le> pi"
- hence "- (pi / 2) \<le> ?x2" and "?x2 \<le> pi / 2" using pi_ge_two unfolding real_of_float_mult Float_num by auto
+ hence "- (pi / 2) \<le> ?x2" and "?x2 \<le> pi / 2" using pi_ge_two unfolding Float_num by auto
hence "0 \<le> cos ?x2" by (rule cos_ge_zero)
have "(?lb_half y) \<le> cos x"
@@ -1077,14 +1063,14 @@
from mult_mono[OF `y \<le> cos ?x2` `y \<le> cos ?x2` `0 \<le> cos ?x2` this]
have "real y * real y \<le> cos ?x2 * cos ?x2" .
hence "2 * real y * real y \<le> 2 * cos ?x2 * cos ?x2" by auto
- hence "2 * real y * real y - 1 \<le> 2 * cos (x / 2) * cos (x / 2) - 1" unfolding Float_num real_of_float_mult by auto
- thus ?thesis unfolding if_not_P[OF False] x_half Float_num real_of_float_mult real_of_float_sub by auto
+ hence "2 * real y * real y - 1 \<le> 2 * cos (x / 2) * cos (x / 2) - 1" unfolding Float_num by auto
+ thus ?thesis unfolding if_not_P[OF False] x_half Float_num by auto
qed
} note lb_half = this
{ fix y x :: float let ?x2 = "(x * Float 1 -1)"
assume ub: "cos ?x2 \<le> y" and "- pi \<le> x" and "x \<le> pi"
- hence "- (pi / 2) \<le> ?x2" and "?x2 \<le> pi / 2" using pi_ge_two unfolding real_of_float_mult Float_num by auto
+ hence "- (pi / 2) \<le> ?x2" and "?x2 \<le> pi / 2" using pi_ge_two unfolding Float_num by auto
hence "0 \<le> cos ?x2" by (rule cos_ge_zero)
have "cos x \<le> (?ub_half y)"
@@ -1093,8 +1079,8 @@
from mult_mono[OF ub ub this `0 \<le> cos ?x2`]
have "cos ?x2 * cos ?x2 \<le> real y * real y" .
hence "2 * cos ?x2 * cos ?x2 \<le> 2 * real y * real y" by auto
- hence "2 * cos (x / 2) * cos (x / 2) - 1 \<le> 2 * real y * real y - 1" unfolding Float_num real_of_float_mult by auto
- thus ?thesis unfolding x_half real_of_float_mult Float_num real_of_float_sub by auto
+ hence "2 * cos (x / 2) * cos (x / 2) - 1 \<le> 2 * real y * real y - 1" unfolding Float_num by auto
+ thus ?thesis unfolding x_half Float_num by auto
qed
} note ub_half = this
@@ -1106,7 +1092,7 @@
show ?thesis
proof (cases "x < 1")
case True hence "real x \<le> 1" unfolding less_float_def by auto
- have "0 \<le> real ?x2" and "?x2 \<le> pi / 2" using pi_ge_two `0 \<le> real x` unfolding real_of_float_mult Float_num using assms by auto
+ have "0 \<le> real ?x2" and "?x2 \<le> pi / 2" using pi_ge_two `0 \<le> real x` using assms by auto
from cos_boundaries[OF this]
have lb: "(?lb_horner ?x2) \<le> ?cos ?x2" and ub: "?cos ?x2 \<le> (?ub_horner ?x2)" by auto
@@ -1123,21 +1109,21 @@
ultimately show ?thesis by auto
next
case False
- have "0 \<le> real ?x4" and "?x4 \<le> pi / 2" using pi_ge_two `0 \<le> real x` `x \<le> pi` unfolding real_of_float_mult Float_num by auto
+ have "0 \<le> real ?x4" and "?x4 \<le> pi / 2" using pi_ge_two `0 \<le> real x` `x \<le> pi` unfolding Float_num by auto
from cos_boundaries[OF this]
have lb: "(?lb_horner ?x4) \<le> ?cos ?x4" and ub: "?cos ?x4 \<le> (?ub_horner ?x4)" by auto
- have eq_4: "?x2 * Float 1 -1 = x * Float 1 -2" by (cases x, auto simp add: times_float.simps)
+ have eq_4: "?x2 * Float 1 -1 = x * Float 1 -2" by simp
have "(?lb x) \<le> ?cos x"
proof -
- have "-pi \<le> ?x2" and "?x2 \<le> pi" unfolding real_of_float_mult Float_num using pi_ge_two `0 \<le> real x` `x \<le> pi` by auto
+ have "-pi \<le> ?x2" and "?x2 \<le> pi" using pi_ge_two `0 \<le> real x` `x \<le> pi` by auto
from lb_half[OF lb_half[OF lb this] `-pi \<le> x` `x \<le> pi`, unfolded eq_4]
show ?thesis unfolding lb_cos_def[where x=x] if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x < Float 1 -1`] if_not_P[OF `\<not> x < 1`] Let_def .
qed
moreover have "?cos x \<le> (?ub x)"
proof -
- have "-pi \<le> ?x2" and "?x2 \<le> pi" unfolding real_of_float_mult Float_num using pi_ge_two `0 \<le> real x` ` x \<le> pi` by auto
+ have "-pi \<le> ?x2" and "?x2 \<le> pi" using pi_ge_two `0 \<le> real x` ` x \<le> pi` by auto
from ub_half[OF ub_half[OF ub this] `-pi \<le> x` `x \<le> pi`, unfolded eq_4]
show ?thesis unfolding ub_cos_def[where x=x] if_not_P[OF `\<not> x < 0`] if_not_P[OF `\<not> x < Float 1 -1`] if_not_P[OF `\<not> x < 1`] Let_def .
qed
@@ -1155,8 +1141,8 @@
definition bnds_cos :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float * float" where
"bnds_cos prec lx ux = (let
- lpi = round_down prec (lb_pi prec) ;
- upi = round_up prec (ub_pi prec) ;
+ lpi = float_round_down prec (lb_pi prec) ;
+ upi = float_round_up prec (ub_pi prec) ;
k = floor_fl (float_divr prec (lx + lpi) (2 * lpi)) ;
lx = lx - k * 2 * (if k < 0 then lpi else upi) ;
ux = ux - k * 2 * (if k < 0 then upi else lpi)
@@ -1169,14 +1155,7 @@
lemma floor_int:
obtains k :: int where "real k = (floor_fl f)"
-proof -
- assume *: "\<And> k :: int. real k = (floor_fl f) \<Longrightarrow> thesis"
- obtain m e where fl: "Float m e = floor_fl f" by (cases "floor_fl f", auto)
- from floor_pos_exp[OF this]
- have "real (m* 2^(nat e)) = (floor_fl f)"
- by (auto simp add: fl[symmetric] real_of_float_def pow2_def)
- from *[OF this] show thesis by blast
-qed
+ by (simp add: floor_fl_def)
lemma cos_periodic_nat[simp]: fixes n :: nat shows "cos (x + n * (2 * pi)) = cos x"
proof (induct n arbitrary: x)
@@ -1203,8 +1182,8 @@
fix x :: real fix lx ux
assume bnds: "(l, u) = bnds_cos prec lx ux" and x: "x \<in> {lx .. ux}"
- let ?lpi = "round_down prec (lb_pi prec)"
- let ?upi = "round_up prec (ub_pi prec)"
+ let ?lpi = "float_round_down prec (lb_pi prec)"
+ let ?upi = "float_round_up prec (ub_pi prec)"
let ?k = "floor_fl (float_divr prec (lx + ?lpi) (2 * ?lpi))"
let ?lx = "lx - ?k * 2 * (if ?k < 0 then ?lpi else ?upi)"
let ?ux = "ux - ?k * 2 * (if ?k < 0 then ?upi else ?lpi)"
@@ -1212,24 +1191,27 @@
obtain k :: int where k: "k = real ?k" using floor_int .
have upi: "pi \<le> ?upi" and lpi: "?lpi \<le> pi"
- using round_up[of "ub_pi prec" prec] pi_boundaries[of prec]
- round_down[of prec "lb_pi prec"] by auto
+ using float_round_up[of "ub_pi prec" prec] pi_boundaries[of prec]
+ float_round_down[of prec "lb_pi prec"] by auto
hence "?lx \<le> x - k * (2 * pi) \<and> x - k * (2 * pi) \<le> ?ux"
- using x by (cases "k = 0") (auto intro!: add_mono
- simp add: diff_minus k[symmetric] less_float_def)
+ using x unfolding k[symmetric]
+ by (cases "k = 0")
+ (auto intro!: add_mono
+ simp add: diff_minus k[symmetric] less_float_def
+ simp del: float_of_numeral)
note lx = this[THEN conjunct1] and ux = this[THEN conjunct2]
hence lx_less_ux: "?lx \<le> real ?ux" by (rule order_trans)
{ assume "- ?lpi \<le> ?lx" and x_le_0: "x - k * (2 * pi) \<le> 0"
with lpi[THEN le_imp_neg_le] lx
have pi_lx: "- pi \<le> ?lx" and lx_0: "real ?lx \<le> 0"
- by (simp_all add: le_float_def)
+ by (simp_all add: less_eq_float_def)
have "(lb_cos prec (- ?lx)) \<le> cos (real (- ?lx))"
using lb_cos_minus[OF pi_lx lx_0] by simp
also have "\<dots> \<le> cos (x + (-k) * (2 * pi))"
using cos_monotone_minus_pi_0'[OF pi_lx lx x_le_0]
- by (simp only: real_of_float_minus real_of_int_minus
+ by (simp only: real_of_float_uminus real_of_int_minus
cos_minus diff_minus mult_minus_left)
finally have "(lb_cos prec (- ?lx)) \<le> cos x"
unfolding cos_periodic_int . }
@@ -1238,11 +1220,11 @@
{ assume "0 \<le> ?lx" and pi_x: "x - k * (2 * pi) \<le> pi"
with lx
have pi_lx: "?lx \<le> pi" and lx_0: "0 \<le> real ?lx"
- by (auto simp add: le_float_def)
+ by (auto simp add: less_eq_float_def)
have "cos (x + (-k) * (2 * pi)) \<le> cos ?lx"
using cos_monotone_0_pi'[OF lx_0 lx pi_x]
- by (simp only: real_of_float_minus real_of_int_minus
+ by (simp only: real_of_int_minus
cos_minus diff_minus mult_minus_left)
also have "\<dots> \<le> (ub_cos prec ?lx)"
using lb_cos[OF lx_0 pi_lx] by simp
@@ -1253,11 +1235,11 @@
{ assume pi_x: "- pi \<le> x - k * (2 * pi)" and "?ux \<le> 0"
with ux
have pi_ux: "- pi \<le> ?ux" and ux_0: "real ?ux \<le> 0"
- by (simp_all add: le_float_def)
+ by (simp_all add: less_eq_float_def)
have "cos (x + (-k) * (2 * pi)) \<le> cos (real (- ?ux))"
using cos_monotone_minus_pi_0'[OF pi_x ux ux_0]
- by (simp only: real_of_float_minus real_of_int_minus
+ by (simp only: real_of_float_uminus real_of_int_minus
cos_minus diff_minus mult_minus_left)
also have "\<dots> \<le> (ub_cos prec (- ?ux))"
using lb_cos_minus[OF pi_ux ux_0, of prec] by simp
@@ -1268,13 +1250,13 @@
{ assume "?ux \<le> ?lpi" and x_ge_0: "0 \<le> x - k * (2 * pi)"
with lpi ux
have pi_ux: "?ux \<le> pi" and ux_0: "0 \<le> real ?ux"
- by (simp_all add: le_float_def)
+ by (simp_all add: less_eq_float_def)
have "(lb_cos prec ?ux) \<le> cos ?ux"
using lb_cos[OF ux_0 pi_ux] by simp
also have "\<dots> \<le> cos (x + (-k) * (2 * pi))"
using cos_monotone_0_pi'[OF x_ge_0 ux pi_ux]
- by (simp only: real_of_float_minus real_of_int_minus
+ by (simp only: real_of_int_minus
cos_minus diff_minus mult_minus_left)
finally have "(lb_cos prec ?ux) \<le> cos x"
unfolding cos_periodic_int . }
@@ -1290,7 +1272,7 @@
from True lpi[THEN le_imp_neg_le] lx ux
have "- pi \<le> x - k * (2 * pi)"
and "x - k * (2 * pi) \<le> 0"
- by (auto simp add: le_float_def)
+ by (auto simp add: less_eq_float_def)
with True negative_ux negative_lx
show ?thesis unfolding l u by simp
next case False note 1 = this show ?thesis
@@ -1303,7 +1285,7 @@
from True lpi lx ux
have "0 \<le> x - k * (2 * pi)"
and "x - k * (2 * pi) \<le> pi"
- by (auto simp add: le_float_def)
+ by (auto simp add: less_eq_float_def)
with True positive_ux positive_lx
show ?thesis unfolding l u by simp
next case False note 2 = this show ?thesis
@@ -1314,7 +1296,8 @@
by (auto simp add: bnds_cos_def Let_def)
show ?thesis unfolding u l using negative_lx positive_ux Cond
- by (cases "x - k * (2 * pi) < 0", simp_all add: real_of_float_min)
+ by (cases "x - k * (2 * pi) < 0") (auto simp add: real_of_float_min)
+
next case False note 3 = this show ?thesis
proof (cases "0 \<le> ?lx \<and> ?ux \<le> 2 * ?lpi")
case True note Cond = this with bnds 1 2 3
@@ -1331,28 +1314,27 @@
case False hence "pi \<le> x - k * (2 * pi)" by simp
hence pi_x: "- pi \<le> x - k * (2 * pi) - 2 * pi" by simp
- have "?ux \<le> 2 * pi" using Cond lpi by (auto simp add: le_float_def)
+ have "?ux \<le> 2 * pi" using Cond lpi by (auto simp add: less_eq_float_def)
hence "x - k * (2 * pi) - 2 * pi \<le> 0" using ux by simp
have ux_0: "real (?ux - 2 * ?lpi) \<le> 0"
- using Cond by (auto simp add: le_float_def)
+ using Cond by (auto simp add: less_eq_float_def)
from 2 and Cond have "\<not> ?ux \<le> ?lpi" by auto
- hence "- ?lpi \<le> ?ux - 2 * ?lpi" by (auto simp add: le_float_def)
+ hence "- ?lpi \<le> ?ux - 2 * ?lpi" by (auto simp add: less_eq_float_def)
hence pi_ux: "- pi \<le> (?ux - 2 * ?lpi)"
- using lpi[THEN le_imp_neg_le] by (auto simp add: le_float_def)
+ using lpi[THEN le_imp_neg_le] by (auto simp add: less_eq_float_def)
have x_le_ux: "x - k * (2 * pi) - 2 * pi \<le> (?ux - 2 * ?lpi)"
using ux lpi by auto
-
have "cos x = cos (x + (-k) * (2 * pi) + (-1::int) * (2 * pi))"
unfolding cos_periodic_int ..
also have "\<dots> \<le> cos ((?ux - 2 * ?lpi))"
using cos_monotone_minus_pi_0'[OF pi_x x_le_ux ux_0]
- by (simp only: real_of_float_minus real_of_int_minus real_of_one
- minus_one [symmetric] diff_minus mult_minus_left mult_1_left)
+ by (simp only: real_of_float_minus real_of_int_minus real_of_one minus_one[symmetric]
+ diff_minus mult_minus_left mult_1_left)
also have "\<dots> = cos ((- (?ux - 2 * ?lpi)))"
- unfolding real_of_float_minus cos_minus ..
+ unfolding real_of_float_uminus cos_minus ..
also have "\<dots> \<le> (ub_cos prec (- (?ux - 2 * ?lpi)))"
using lb_cos_minus[OF pi_ux ux_0] by simp
finally show ?thesis unfolding u by (simp add: real_of_float_max)
@@ -1363,7 +1345,7 @@
case True note Cond = this with bnds 1 2 3 4
have l: "l = Float -1 0"
and u: "u = max (ub_cos prec (?lx + 2 * ?lpi)) (ub_cos prec (-?ux))"
- by (auto simp add: bnds_cos_def Let_def)
+ by (auto simp add: bnds_cos_def Let_def simp del: neg_numeral_float_Float)
have "cos x \<le> u"
proof (cases "-pi < x - k * (2 * pi)")
@@ -1374,17 +1356,17 @@
case False hence "x - k * (2 * pi) \<le> -pi" by simp
hence pi_x: "x - k * (2 * pi) + 2 * pi \<le> pi" by simp
- have "-2 * pi \<le> ?lx" using Cond lpi by (auto simp add: le_float_def)
+ have "-2 * pi \<le> ?lx" using Cond lpi by (auto simp add: less_eq_float_def)
hence "0 \<le> x - k * (2 * pi) + 2 * pi" using lx by simp
have lx_0: "0 \<le> real (?lx + 2 * ?lpi)"
- using Cond lpi by (auto simp add: le_float_def)
+ using Cond lpi by (auto simp add: less_eq_float_def)
from 1 and Cond have "\<not> -?lpi \<le> ?lx" by auto
- hence "?lx + 2 * ?lpi \<le> ?lpi" by (auto simp add: le_float_def)
+ hence "?lx + 2 * ?lpi \<le> ?lpi" by (auto simp add: less_eq_float_def)
hence pi_lx: "(?lx + 2 * ?lpi) \<le> pi"
- using lpi[THEN le_imp_neg_le] by (auto simp add: le_float_def)
+ using lpi[THEN le_imp_neg_le] by (auto simp add: less_eq_float_def)
have lx_le_x: "(?lx + 2 * ?lpi) \<le> x - k * (2 * pi) + 2 * pi"
using lx lpi by auto
@@ -1394,7 +1376,7 @@
also have "\<dots> \<le> cos ((?lx + 2 * ?lpi))"
using cos_monotone_0_pi'[OF lx_0 lx_le_x pi_x]
by (simp only: real_of_float_minus real_of_int_minus real_of_one
- minus_one [symmetric] diff_minus mult_minus_left mult_1_left)
+ minus_one[symmetric] diff_minus mult_minus_left mult_1_left)
also have "\<dots> \<le> (ub_cos prec (?lx + 2 * ?lpi))"
using lb_cos[OF lx_0 pi_lx] by simp
finally show ?thesis unfolding u by (simp add: real_of_float_max)
@@ -1467,11 +1449,10 @@
"lb_exp prec x = (if 0 < x then float_divl prec 1 (ub_exp prec (-x))
else let
horner = (\<lambda> x. let y = lb_exp_horner prec (get_even (prec + 2)) 1 1 x in if y \<le> 0 then Float 1 -2 else y)
- in if x < - 1 then (case floor_fl x of (Float m e) \<Rightarrow> (horner (float_divl prec x (- Float m e))) ^ (nat (-m) * 2 ^ nat e))
+ in if x < - 1 then (horner (float_divl prec x (- floor_fl x))) ^ nat (- int_floor_fl x)
else horner x)" |
"ub_exp prec x = (if 0 < x then float_divr prec 1 (lb_exp prec (-x))
- else if x < - 1 then (case floor_fl x of (Float m e) \<Rightarrow>
- (ub_exp_horner prec (get_odd (prec + 2)) 1 1 (float_divr prec x (- Float m e))) ^ (nat (-m) * 2 ^ nat e))
+ else if x < - 1 then ub_exp_horner prec (get_odd (prec + 2)) 1 1 (float_divr prec x (- floor_fl x)) ^ (nat (- int_floor_fl x))
else ub_exp_horner prec (get_odd (prec + 2)) 1 1 x)"
by pat_completeness auto
termination by (relation "measure (\<lambda> v. let (prec, x) = sum_case id id v in (if 0 < x then 1 else 0))", auto simp add: less_float_def)
@@ -1483,24 +1464,24 @@
have "1 / 4 = (Float 1 -2)" unfolding Float_num by auto
also have "\<dots> \<le> lb_exp_horner 1 (get_even 4) 1 1 (- 1)"
unfolding get_even_def eq4
- by (auto simp add: lapprox_posrat_def rapprox_posrat_def normfloat.simps)
+ by (auto simp add: compute_lapprox_rat compute_rapprox_rat compute_lapprox_posrat compute_rapprox_posrat rat_precision_def compute_bitlen)
also have "\<dots> \<le> exp (- 1 :: float)" using bnds_exp_horner[where x="- 1"] by auto
- finally show ?thesis unfolding real_of_float_minus real_of_float_1 .
+ finally show ?thesis unfolding real_of_float_minus real_of_float_one by simp
qed
lemma lb_exp_pos: assumes "\<not> 0 < x" shows "0 < lb_exp prec x"
proof -
let "?lb_horner x" = "lb_exp_horner prec (get_even (prec + 2)) 1 1 x"
let "?horner x" = "let y = ?lb_horner x in if y \<le> 0 then Float 1 -2 else y"
- have pos_horner: "\<And> x. 0 < ?horner x" unfolding Let_def by (cases "?lb_horner x \<le> 0", auto simp add: le_float_def less_float_def)
+ have pos_horner: "\<And> x. 0 < ?horner x" unfolding Let_def by (cases "?lb_horner x \<le> 0", auto simp add: less_eq_float_def less_float_def)
moreover { fix x :: float fix num :: nat
- have "0 < real (?horner x) ^ num" using `0 < ?horner x`[unfolded less_float_def real_of_float_0] by (rule zero_less_power)
- also have "\<dots> = (?horner x) ^ num" using float_power by auto
+ have "0 < real (?horner x) ^ num" using `0 < ?horner x`[unfolded less_float_def real_of_float_zero] by (rule zero_less_power)
+ also have "\<dots> = (?horner x) ^ num" by auto
finally have "0 < real ((?horner x) ^ num)" .
}
ultimately show ?thesis
unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] Let_def
- by (cases "floor_fl x", cases "x < - 1", auto simp add: float_power le_float_def less_float_def)
+ by (cases "floor_fl x", cases "x < - 1", auto simp add: less_eq_float_def less_float_def)
qed
lemma exp_boundaries': assumes "x \<le> 0"
@@ -1509,7 +1490,7 @@
let "?lb_exp_horner x" = "lb_exp_horner prec (get_even (prec + 2)) 1 1 x"
let "?ub_exp_horner x" = "ub_exp_horner prec (get_odd (prec + 2)) 1 1 x"
- have "real x \<le> 0" and "\<not> x > 0" using `x \<le> 0` unfolding le_float_def less_float_def by auto
+ have "real x \<le> 0" and "\<not> x > 0" using `x \<le> 0` unfolding less_eq_float_def less_float_def by auto
show ?thesis
proof (cases "x < - 1")
case False hence "- 1 \<le> real x" unfolding less_float_def by auto
@@ -1527,62 +1508,62 @@
next
case True
- obtain m e where Float_floor: "floor_fl x = Float m e" by (cases "floor_fl x", auto)
- let ?num = "nat (- m) * 2 ^ nat e"
-
- have "real (floor_fl x) < - 1" using floor_fl `x < - 1` unfolding le_float_def less_float_def real_of_float_minus real_of_float_1 by (rule order_le_less_trans)
- hence "real (floor_fl x) < 0" unfolding Float_floor real_of_float_simp using zero_less_pow2[of xe] by auto
- hence "m < 0"
- unfolding less_float_def real_of_float_0 Float_floor real_of_float_simp
- unfolding pos_prod_lt[OF zero_less_pow2[of e], unfolded mult_commute] by auto
- hence "1 \<le> - m" by auto
- hence "0 < nat (- m)" by auto
- moreover
- have "0 \<le> e" using floor_pos_exp Float_floor[symmetric] by auto
- hence "(0::nat) < 2 ^ nat e" by auto
- ultimately have "0 < ?num" by auto
+ let ?num = "nat (- int_floor_fl x)"
+
+ have "real (int_floor_fl x) < - 1" using int_floor_fl `x < - 1` unfolding less_eq_float_def less_float_def real_of_float_uminus real_of_float_one
+ by (rule order_le_less_trans)
+ hence "real (int_floor_fl x) < 0" by simp
+ hence "int_floor_fl x < 0" by auto
+ hence "1 \<le> - int_floor_fl x" by auto
+ hence "0 < nat (- int_floor_fl x)" by auto
+ hence "0 < ?num" by auto
hence "real ?num \<noteq> 0" by auto
- have e_nat: "(nat e) = e" using `0 \<le> e` by auto
- have num_eq: "real ?num = - floor_fl x" using `0 < nat (- m)`
- unfolding Float_floor real_of_float_minus real_of_float_simp real_of_nat_mult pow2_int[of "nat e", unfolded e_nat] real_of_nat_power by auto
- have "0 < - floor_fl x" using `0 < ?num`[unfolded real_of_nat_less_iff[symmetric]] unfolding less_float_def num_eq[symmetric] real_of_float_0 real_of_nat_zero .
- hence "real (floor_fl x) < 0" unfolding less_float_def by auto
-
+ have num_eq: "real ?num = - int_floor_fl x" using `0 < nat (- int_floor_fl x)` by auto
+ have "0 < - int_floor_fl x" using `0 < ?num`[unfolded real_of_nat_less_iff[symmetric]] by simp
+ hence "real (int_floor_fl x) < 0" unfolding less_float_def by auto
+ have fl_eq: "real (- int_floor_fl x) = real (- floor_fl x)"
+ by (simp add: floor_fl_def int_floor_fl_def)
+ from `0 < - int_floor_fl x` have "0 < real (- floor_fl x)"
+ by (simp add: floor_fl_def int_floor_fl_def)
+ from `real (int_floor_fl x) < 0` have "real (floor_fl x) < 0"
+ by (simp add: floor_fl_def int_floor_fl_def)
have "exp x \<le> ub_exp prec x"
proof -
have div_less_zero: "real (float_divr prec x (- floor_fl x)) \<le> 0"
- using float_divr_nonpos_pos_upper_bound[OF `x \<le> 0` `0 < - floor_fl x`] unfolding le_float_def real_of_float_0 .
+ using float_divr_nonpos_pos_upper_bound[OF `real x \<le> 0` `0 < real (- floor_fl x)`]
+ unfolding less_eq_float_def real_of_float_zero .
have "exp x = exp (?num * (x / ?num))" using `real ?num \<noteq> 0` by auto
also have "\<dots> = exp (x / ?num) ^ ?num" unfolding exp_real_of_nat_mult ..
- also have "\<dots> \<le> exp (float_divr prec x (- floor_fl x)) ^ ?num" unfolding num_eq
+ also have "\<dots> \<le> exp (float_divr prec x (- floor_fl x)) ^ ?num" unfolding num_eq fl_eq
by (rule power_mono, rule exp_le_cancel_iff[THEN iffD2], rule float_divr) auto
- also have "\<dots> \<le> (?ub_exp_horner (float_divr prec x (- floor_fl x))) ^ ?num" unfolding float_power
+ also have "\<dots> \<le> (?ub_exp_horner (float_divr prec x (- floor_fl x))) ^ ?num"
+ unfolding real_of_float_power
by (rule power_mono, rule bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct2], auto)
- finally show ?thesis unfolding ub_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] float.cases Float_floor Let_def .
+ finally show ?thesis unfolding ub_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] floor_fl_def Let_def .
qed
moreover
have "lb_exp prec x \<le> exp x"
proof -
- let ?divl = "float_divl prec x (- Float m e)"
+ let ?divl = "float_divl prec x (- floor_fl x)"
let ?horner = "?lb_exp_horner ?divl"
show ?thesis
proof (cases "?horner \<le> 0")
- case False hence "0 \<le> real ?horner" unfolding le_float_def by auto
+ case False hence "0 \<le> real ?horner" unfolding less_eq_float_def by auto
have div_less_zero: "real (float_divl prec x (- floor_fl x)) \<le> 0"
using `real (floor_fl x) < 0` `real x \<le> 0` by (auto intro!: order_trans[OF float_divl] divide_nonpos_neg)
have "(?lb_exp_horner (float_divl prec x (- floor_fl x))) ^ ?num \<le>
- exp (float_divl prec x (- floor_fl x)) ^ ?num" unfolding float_power
- using `0 \<le> real ?horner`[unfolded Float_floor[symmetric]] bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct1] by (auto intro!: power_mono)
- also have "\<dots> \<le> exp (x / ?num) ^ ?num" unfolding num_eq
- using float_divl by (auto intro!: power_mono simp del: real_of_float_minus)
+ exp (float_divl prec x (- floor_fl x)) ^ ?num"
+ using `0 \<le> real ?horner`[unfolded floor_fl_def[symmetric]] bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct1] by (auto intro!: power_mono)
+ also have "\<dots> \<le> exp (x / ?num) ^ ?num" unfolding num_eq fl_eq
+ using float_divl by (auto intro!: power_mono simp del: real_of_float_uminus)
also have "\<dots> = exp (?num * (x / ?num))" unfolding exp_real_of_nat_mult ..
also have "\<dots> = exp x" using `real ?num \<noteq> 0` by auto
finally show ?thesis
- unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] float.cases Float_floor Let_def if_not_P[OF False] by auto
+ unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] int_floor_fl_def Let_def if_not_P[OF False] by auto
next
case True
have "real (floor_fl x) \<noteq> 0" and "real (floor_fl x) \<le> 0" using `real (floor_fl x) < 0` by auto
@@ -1592,9 +1573,9 @@
have "Float 1 -2 \<le> exp (x / (- floor_fl x))" unfolding Float_num .
hence "real (Float 1 -2) ^ ?num \<le> exp (x / (- floor_fl x)) ^ ?num"
by (auto intro!: power_mono)
- also have "\<dots> = exp x" unfolding num_eq exp_real_of_nat_mult[symmetric] using `real (floor_fl x) \<noteq> 0` by auto
+ also have "\<dots> = exp x" unfolding num_eq fl_eq exp_real_of_nat_mult[symmetric] using `real (floor_fl x) \<noteq> 0` by auto
finally show ?thesis
- unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] float.cases Float_floor Let_def if_P[OF True] float_power .
+ unfolding lb_exp.simps if_not_P[OF `\<not> 0 < x`] if_P[OF `x < - 1`] int_floor_fl_def Let_def if_P[OF True] real_of_float_power .
qed
qed
ultimately show ?thesis by auto
@@ -1605,10 +1586,10 @@
proof -
show ?thesis
proof (cases "0 < x")
- case False hence "x \<le> 0" unfolding less_float_def le_float_def by auto
+ case False hence "x \<le> 0" unfolding less_float_def less_eq_float_def by auto
from exp_boundaries'[OF this] show ?thesis .
next
- case True hence "-x \<le> 0" unfolding less_float_def le_float_def by auto
+ case True hence "-x \<le> 0" unfolding less_float_def less_eq_float_def by auto
have "lb_exp prec x \<le> exp x"
proof -
@@ -1630,9 +1611,9 @@
have lb_exp: "lb_exp prec (-x) \<le> exp (- real x)" unfolding atLeastAtMost_iff real_of_float_minus by auto
have "exp x \<le> (1 :: float) / lb_exp prec (-x)"
- using lb_exp[unfolded inverse_le_iff_le[OF exp_gt_zero lb_exp_pos[OF `\<not> 0 < -x`, unfolded less_float_def real_of_float_0],
+ using lb_exp[unfolded inverse_le_iff_le[OF exp_gt_zero lb_exp_pos[OF `\<not> 0 < -x`, unfolded less_float_def real_of_float_zero],
symmetric]]
- unfolding exp_minus nonzero_inverse_inverse_eq[OF exp_not_eq_zero] inverse_eq_divide real_of_float_1 by auto
+ unfolding exp_minus nonzero_inverse_inverse_eq[OF exp_not_eq_zero] inverse_eq_divide real_of_float_one by auto
also have "\<dots> \<le> float_divr prec 1 (lb_exp prec (-x))" using float_divr .
finally show ?thesis unfolding ub_exp.simps if_P[OF True] .
qed
@@ -1703,7 +1684,7 @@
let "?s n" = "-1^n * (1 / real (1 + n)) * (real x)^(Suc n)"
- have "?lb \<le> setsum ?s {0 ..< 2 * ev}" unfolding power_Suc2 mult_assoc[symmetric] real_of_float_mult setsum_left_distrib[symmetric] unfolding mult_commute[of "real x"] ev
+ have "?lb \<le> setsum ?s {0 ..< 2 * ev}" unfolding power_Suc2 mult_assoc[symmetric] real_of_float_times setsum_left_distrib[symmetric] unfolding mult_commute[of "real x"] ev
using horner_bounds(1)[where G="\<lambda> i k. Suc k" and F="\<lambda>x. x" and f="\<lambda>x. x" and lb="\<lambda>n i k x. lb_ln_horner prec n k x" and ub="\<lambda>n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*ev",
OF `0 \<le> real x` refl lb_ln_horner.simps ub_ln_horner.simps] `0 \<le> real x`
by (rule mult_right_mono)
@@ -1711,7 +1692,7 @@
finally show "?lb \<le> ?ln" .
have "?ln \<le> setsum ?s {0 ..< 2 * od + 1}" using ln_bounds(2)[OF `0 \<le> real x` `real x < 1`] by auto
- also have "\<dots> \<le> ?ub" unfolding power_Suc2 mult_assoc[symmetric] real_of_float_mult setsum_left_distrib[symmetric] unfolding mult_commute[of "real x"] od
+ also have "\<dots> \<le> ?ub" unfolding power_Suc2 mult_assoc[symmetric] real_of_float_times setsum_left_distrib[symmetric] unfolding mult_commute[of "real x"] od
using horner_bounds(2)[where G="\<lambda> i k. Suc k" and F="\<lambda>x. x" and f="\<lambda>x. x" and lb="\<lambda>n i k x. lb_ln_horner prec n k x" and ub="\<lambda>n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*od+1",
OF `0 \<le> real x` refl lb_ln_horner.simps ub_ln_horner.simps] `0 \<le> real x`
by (rule mult_right_mono)
@@ -1747,28 +1728,27 @@
using ln_add[of "3 / 2" "1 / 2"] by auto
have lb3: "?lthird \<le> 1 / 3" using lapprox_rat[of prec 1 3] by auto
hence lb3_ub: "real ?lthird < 1" by auto
- have lb3_lb: "0 \<le> real ?lthird" using lapprox_rat_bottom[of 1 3] by auto
+ have lb3_lb: "0 \<le> real ?lthird" using lapprox_rat_nonneg[of 1 3] by auto
have ub3: "1 / 3 \<le> ?uthird" using rapprox_rat[of 1 3] by auto
hence ub3_lb: "0 \<le> real ?uthird" by auto
have lb2: "0 \<le> real (Float 1 -1)" and ub2: "real (Float 1 -1) < 1" unfolding Float_num by auto
have "0 \<le> (1::int)" and "0 < (3::int)" by auto
- have ub3_ub: "real ?uthird < 1" unfolding rapprox_rat.simps(2)[OF `0 \<le> 1` `0 < 3`]
- by (rule rapprox_posrat_less1, auto)
+ have ub3_ub: "real ?uthird < 1" by (simp add: compute_rapprox_rat rapprox_posrat_less1)
have third_gt0: "(0 :: real) < 1 / 3 + 1" by auto
have uthird_gt0: "0 < real ?uthird + 1" using ub3_lb by auto
have lthird_gt0: "0 < real ?lthird + 1" using lb3_lb by auto
- show ?ub_ln2 unfolding ub_ln2_def Let_def real_of_float_add ln2_sum Float_num(4)[symmetric]
+ show ?ub_ln2 unfolding ub_ln2_def Let_def real_of_float_plus ln2_sum Float_num(4)[symmetric]
proof (rule add_mono, fact ln_float_bounds(2)[OF lb2 ub2])
have "ln (1 / 3 + 1) \<le> ln (real ?uthird + 1)" unfolding ln_le_cancel_iff[OF third_gt0 uthird_gt0] using ub3 by auto
also have "\<dots> \<le> ?uthird * ub_ln_horner prec (get_odd prec) 1 ?uthird"
using ln_float_bounds(2)[OF ub3_lb ub3_ub] .
finally show "ln (1 / 3 + 1) \<le> ?uthird * ub_ln_horner prec (get_odd prec) 1 ?uthird" .
qed
- show ?lb_ln2 unfolding lb_ln2_def Let_def real_of_float_add ln2_sum Float_num(4)[symmetric]
+ show ?lb_ln2 unfolding lb_ln2_def Let_def real_of_float_plus ln2_sum Float_num(4)[symmetric]
proof (rule add_mono, fact ln_float_bounds(1)[OF lb2 ub2])
have "?lthird * lb_ln_horner prec (get_even prec) 1 ?lthird \<le> ln (real ?lthird + 1)"
using ln_float_bounds(1)[OF lb3_lb lb3_ub] .
@@ -1786,7 +1766,7 @@
if x \<le> Float 3 -1 then Some (horner (x - 1))
else if x < Float 1 1 then Some (horner (Float 1 -1) + horner (x * rapprox_rat prec 2 3 - 1))
else let l = bitlen (mantissa x) - 1 in
- Some (ub_ln2 prec * (Float (scale x + l) 0) + horner (Float (mantissa x) (- l) - 1)))" |
+ Some (ub_ln2 prec * (Float (exponent x + l) 0) + horner (Float (mantissa x) (- l) - 1)))" |
"lb_ln prec x = (if x \<le> 0 then None
else if x < 1 then Some (- the (ub_ln prec (float_divr prec 1 x)))
else let horner = \<lambda>x. x * lb_ln_horner prec (get_even prec) 1 x in
@@ -1794,41 +1774,117 @@
else if x < Float 1 1 then Some (horner (Float 1 -1) +
horner (max (x * lapprox_rat prec 2 3 - 1) 0))
else let l = bitlen (mantissa x) - 1 in
- Some (lb_ln2 prec * (Float (scale x + l) 0) + horner (Float (mantissa x) (- l) - 1)))"
+ Some (lb_ln2 prec * (Float (exponent x + l) 0) + horner (Float (mantissa x) (- l) - 1)))"
by pat_completeness auto
termination proof (relation "measure (\<lambda> v. let (prec, x) = sum_case id id v in (if x < 1 then 1 else 0))", auto)
fix prec x assume "\<not> x \<le> 0" and "x < 1" and "float_divl (max prec (Suc 0)) 1 x < 1"
- hence "0 < x" and "0 < max prec (Suc 0)" unfolding less_float_def le_float_def by auto
- from float_divl_pos_less1_bound[OF `0 < x` `x < 1` `0 < max prec (Suc 0)`]
- show False using `float_divl (max prec (Suc 0)) 1 x < 1` unfolding less_float_def le_float_def by auto
+ hence "0 < real x" "1 \<le> max prec (Suc 0)" "real x < 1"
+ unfolding less_float_def less_eq_float_def by auto
+ from float_divl_pos_less1_bound[OF `0 < real x` `real x < 1` `1 \<le> max prec (Suc 0)`]
+ show False using `float_divl (max prec (Suc 0)) 1 x < 1` unfolding less_float_def less_eq_float_def by auto
next
fix prec x assume "\<not> x \<le> 0" and "x < 1" and "float_divr prec 1 x < 1"
- hence "0 < x" unfolding less_float_def le_float_def by auto
+ hence "0 < x" unfolding less_float_def less_eq_float_def by auto
from float_divr_pos_less1_lower_bound[OF `0 < x` `x < 1`, of prec]
- show False using `float_divr prec 1 x < 1` unfolding less_float_def le_float_def by auto
+ show False using `float_divr prec 1 x < 1` unfolding less_float_def less_eq_float_def by auto
+qed
+
+lemma float_pos_eq_mantissa_pos: "x > 0 \<longleftrightarrow> mantissa x > 0"
+ apply (subst Float_mantissa_exponent[of x, symmetric])
+ apply (auto simp add: zero_less_mult_iff zero_float_def powr_gt_zero[of 2 "exponent x"] dest: less_zeroE)
+ using powr_gt_zero[of 2 "exponent x"]
+ apply simp
+ done
+
+lemma Float_pos_eq_mantissa_pos: "Float m e > 0 \<longleftrightarrow> m > 0"
+ apply (auto simp add: zero_less_mult_iff zero_float_def powr_gt_zero[of 2 "exponent x"] dest: less_zeroE)
+ using powr_gt_zero[of 2 "e"]
+ apply simp
+ done
+
+lemma Float_representation_aux:
+ fixes m e
+ defines "x \<equiv> Float m e"
+ assumes "x > 0"
+ shows "Float (exponent x + (bitlen (mantissa x) - 1)) 0 = Float (e + (bitlen m - 1)) 0" (is ?th1)
+ and "Float (mantissa x) (- (bitlen (mantissa x) - 1)) = Float m ( - (bitlen m - 1))" (is ?th2)
+proof -
+ from assms have mantissa_pos: "m > 0" "mantissa x > 0"
+ using Float_pos_eq_mantissa_pos float_pos_eq_mantissa_pos by simp_all
+ thus ?th1 using bitlen_Float[of m e] assms by (simp add: less_float_def zero_less_mult_iff)
+ from assms have "x \<noteq> float_of 0" by (auto simp add: zero_float_def zero_less_mult_iff)
+ from denormalize_shift[OF assms(1) this] guess i . note i = this
+ from `x \<noteq> float_of 0` have "mantissa x \<noteq> 0" by (simp add: mantissa_noteq_0)
+ from assms
+ have "real (mantissa x) * 2 powr (1 - real (bitlen (mantissa x))) \<in> float"
+ using two_powr_int_float[of "1 - bitlen (mantissa x)"] by simp
+ moreover
+ have "2 powr (1 - (real (bitlen (mantissa x)) + real i)) =
+ 2 powr (1 - (real (bitlen (mantissa x)))) * inverse (2 powr (real i))"
+ by (simp add: powr_minus[symmetric] powr_add[symmetric] field_simps)
+ hence "real (mantissa x) * 2 powr (1 - real (bitlen (mantissa x))) =
+ (real (mantissa x) * 2 ^ i) * 2 powr (1 - real (bitlen (mantissa x * 2 ^ i)))"
+ using `mantissa x > 0` by (simp add: powr_realpow)
+ ultimately
+ show ?th2
+ using two_powr_int_float[of "1 - bitlen (mantissa x)"] by (simp add: i)
+qed
+
+lemma compute_ln[code]:
+ fixes m e
+ defines "x \<equiv> Float m e"
+ shows "ub_ln prec x = (if x \<le> 0 then None
+ else if x < 1 then Some (- the (lb_ln prec (float_divl (max prec 1) 1 x)))
+ else let horner = \<lambda>x. x * ub_ln_horner prec (get_odd prec) 1 x in
+ if x \<le> Float 3 -1 then Some (horner (x - 1))
+ else if x < Float 1 1 then Some (horner (Float 1 -1) + horner (x * rapprox_rat prec 2 3 - 1))
+ else let l = bitlen m - 1 in
+ Some (ub_ln2 prec * (Float (e + l) 0) + horner (Float m (- l) - 1)))"
+ (is ?th1)
+ and "lb_ln prec x = (if x \<le> 0 then None
+ else if x < 1 then Some (- the (ub_ln prec (float_divr prec 1 x)))
+ else let horner = \<lambda>x. x * lb_ln_horner prec (get_even prec) 1 x in
+ if x \<le> Float 3 -1 then Some (horner (x - 1))
+ else if x < Float 1 1 then Some (horner (Float 1 -1) +
+ horner (max (x * lapprox_rat prec 2 3 - 1) 0))
+ else let l = bitlen m - 1 in
+ Some (lb_ln2 prec * (Float (e + l) 0) + horner (Float m (- l) - 1)))"
+ (is ?th2)
+proof -
+ from assms Float_pos_eq_mantissa_pos have "x > 0 \<Longrightarrow> m > 0" by simp
+ thus ?th1 ?th2 using Float_representation_aux[of m e] unfolding x_def[symmetric]
+ by (auto simp add: simp del: Float_def dest: not_leE)
qed
lemma ln_shifted_float: assumes "0 < m" shows "ln (Float m e) = ln 2 * (e + (bitlen m - 1)) + ln (Float m (- (bitlen m - 1)))"
proof -
let ?B = "2^nat (bitlen m - 1)"
+ def bl \<equiv> "bitlen m - 1"
have "0 < real m" and "\<And>X. (0 :: real) < 2^X" and "0 < (2 :: real)" and "m \<noteq> 0" using assms by auto
- hence "0 \<le> bitlen m - 1" using bitlen_ge1[OF `m \<noteq> 0`] by auto
+ hence "0 \<le> bl" by (simp add: bitlen_def bl_def)
show ?thesis
proof (cases "0 \<le> e")
- case True
- show ?thesis unfolding normalized_float[OF `m \<noteq> 0`]
- unfolding ln_div[OF `0 < real m` `0 < ?B`] real_of_int_add ln_realpow[OF `0 < 2`]
- unfolding real_of_float_ge0_exp[OF True] ln_mult[OF `0 < real m` `0 < 2^nat e`]
- ln_realpow[OF `0 < 2`] algebra_simps using `0 \<le> bitlen m - 1` True by auto
+ case True
+ thus ?thesis
+ unfolding bl_def[symmetric] using `0 < real m` `0 \<le> bl`
+ apply (simp add: ln_mult)
+ apply (cases "e=0")
+ apply (cases "bl = 0", simp_all add: powr_minus ln_inverse ln_powr)
+ apply (cases "bl = 0", simp_all add: powr_minus ln_inverse ln_powr field_simps)
+ done
next
case False hence "0 < -e" by auto
+ have lne: "ln (2 powr real e) = ln (inverse (2 powr - e))" by (simp add: powr_minus)
hence pow_gt0: "(0::real) < 2^nat (-e)" by auto
hence inv_gt0: "(0::real) < inverse (2^nat (-e))" by auto
- show ?thesis unfolding normalized_float[OF `m \<noteq> 0`]
- unfolding ln_div[OF `0 < real m` `0 < ?B`] real_of_int_add ln_realpow[OF `0 < 2`]
- unfolding real_of_float_nge0_exp[OF False] ln_mult[OF `0 < real m` inv_gt0] ln_inverse[OF pow_gt0]
- ln_realpow[OF `0 < 2`] algebra_simps using `0 \<le> bitlen m - 1` False by auto
+ show ?thesis using False unfolding bl_def[symmetric] using `0 < real m` `0 \<le> bl`
+ apply (simp add: ln_mult lne)
+ apply (cases "e=0")
+ apply (cases "bl = 0", simp_all add: powr_minus ln_inverse ln_powr)
+ apply (simp add: ln_inverse lne)
+ apply (cases "bl = 0", simp_all add: ln_inverse ln_powr field_simps)
+ done
qed
qed
@@ -1838,10 +1894,10 @@
proof (cases "x < Float 1 1")
case True
hence "real (x - 1) < 1" and "real x < 2" unfolding less_float_def Float_num by auto
- have "\<not> x \<le> 0" and "\<not> x < 1" using `1 \<le> x` unfolding less_float_def le_float_def by auto
+ have "\<not> x \<le> 0" and "\<not> x < 1" using `1 \<le> x` unfolding less_float_def less_eq_float_def by auto
hence "0 \<le> real (x - 1)" using `1 \<le> x` unfolding less_float_def Float_num by auto
- have [simp]: "(Float 3 -1) = 3 / 2" by (simp add: real_of_float_def pow2_def)
+ have [simp]: "(Float 3 -1) = 3 / 2" by simp
show ?thesis
proof (cases "x \<le> Float 3 -1")
@@ -1850,7 +1906,7 @@
using ln_float_bounds[OF `0 \<le> real (x - 1)` `real (x - 1) < 1`, of prec] `\<not> x \<le> 0` `\<not> x < 1` True
by auto
next
- case False hence *: "3 / 2 < x" by (auto simp add: le_float_def)
+ case False hence *: "3 / 2 < x" by (auto simp add: less_eq_float_def)
with ln_add[of "3 / 2" "x - 3 / 2"]
have add: "ln x = ln (3 / 2) + ln (real x * 2 / 3)"
@@ -1891,7 +1947,7 @@
by (rule order_trans[OF lapprox_rat], simp)
have low: "0 \<le> real (lapprox_rat prec 2 3)"
- using lapprox_rat_bottom[of 2 3 prec] by simp
+ using lapprox_rat_nonneg[of 2 3 prec] by simp
have "?lb_horner ?max
\<le> ln (real ?max + 1)"
@@ -1907,7 +1963,7 @@
show "0 < real x * 2/3" using * by auto
show "real ?max + 1 \<le> real x * 2/3" using * up
by (cases "0 < real x * real (lapprox_posrat prec 2 3) - 1",
- auto simp add: real_of_float_max min_max.sup_absorb1)
+ auto simp add: max_def)
qed
finally have "?lb_horner (Float 1 -1) + ?lb_horner ?max
\<le> ln x"
@@ -1919,55 +1975,61 @@
next
case False
hence "\<not> x \<le> 0" and "\<not> x < 1" "0 < x" "\<not> x \<le> Float 3 -1"
- using `1 \<le> x` unfolding less_float_def le_float_def real_of_float_simp pow2_def
+ using `1 \<le> x` unfolding less_float_def less_eq_float_def
by auto
show ?thesis
- proof (cases x)
- case (Float m e)
+ proof -
+ def m \<equiv> "mantissa x"
+ def e \<equiv> "exponent x"
+ from Float_mantissa_exponent[of x] have Float: "x = Float m e" by (simp add: m_def e_def)
let ?s = "Float (e + (bitlen m - 1)) 0"
let ?x = "Float m (- (bitlen m - 1))"
- have "0 < m" and "m \<noteq> 0" using float_pos_m_pos `0 < x` Float by auto
+ have "0 < m" and "m \<noteq> 0" using `0 < x` Float powr_gt_zero[of 2 e]
+ by (auto simp: less_float_def less_eq_float_def zero_less_mult_iff)
+ def bl \<equiv> "bitlen m - 1" hence "bl \<ge> 0" using `m > 0` by (simp add: bitlen_def)
+ have "1 \<le> Float m e" using `1 \<le> x` Float unfolding less_eq_float_def by auto
+ from bitlen_div[OF `0 < m`] float_gt1_scale[OF `1 \<le> Float m e`] `bl \<ge> 0`
+ have x_bnds: "0 \<le> real (?x - 1)" "real (?x - 1) < 1"
+ unfolding bl_def[symmetric]
+ by (auto simp: powr_realpow[symmetric] field_simps inverse_eq_divide)
+ (auto simp : powr_minus field_simps inverse_eq_divide)
{
have "lb_ln2 prec * ?s \<le> ln 2 * (e + (bitlen m - 1))" (is "?lb2 \<le> _")
- unfolding real_of_float_mult real_of_float_ge0_exp[OF order_refl] nat_0 power_0 mult_1_right
+ unfolding nat_0 power_0 mult_1_right real_of_float_times
using lb_ln2[of prec]
- proof (rule mult_right_mono)
- have "1 \<le> Float m e" using `1 \<le> x` Float unfolding le_float_def by auto
- from float_gt1_scale[OF this]
- show "0 \<le> real (e + (bitlen m - 1))" by auto
- qed
+ proof (rule mult_mono)
+ from float_gt1_scale[OF `1 \<le> Float m e`]
+ show "0 \<le> real (Float (e + (bitlen m - 1)) 0)" by simp
+ qed auto
moreover
- from bitlen_div[OF `0 < m`, unfolded normalized_float[OF `m \<noteq> 0`, symmetric]]
- have "0 \<le> real (?x - 1)" and "real (?x - 1) < 1" by auto
- from ln_float_bounds(1)[OF this]
+ from ln_float_bounds(1)[OF x_bnds]
have "(?x - 1) * lb_ln_horner prec (get_even prec) 1 (?x - 1) \<le> ln ?x" (is "?lb_horner \<le> _") by auto
ultimately have "?lb2 + ?lb_horner \<le> ln x"
unfolding Float ln_shifted_float[OF `0 < m`, of e] by auto
}
moreover
{
- from bitlen_div[OF `0 < m`, unfolded normalized_float[OF `m \<noteq> 0`, symmetric]]
- have "0 \<le> real (?x - 1)" and "real (?x - 1) < 1" by auto
- from ln_float_bounds(2)[OF this]
+ from ln_float_bounds(2)[OF x_bnds]
have "ln ?x \<le> (?x - 1) * ub_ln_horner prec (get_odd prec) 1 (?x - 1)" (is "_ \<le> ?ub_horner") by auto
moreover
have "ln 2 * (e + (bitlen m - 1)) \<le> ub_ln2 prec * ?s" (is "_ \<le> ?ub2")
- unfolding real_of_float_mult real_of_float_ge0_exp[OF order_refl] nat_0 power_0 mult_1_right
+ unfolding nat_0 power_0 mult_1_right real_of_float_times
using ub_ln2[of prec]
- proof (rule mult_right_mono)
- have "1 \<le> Float m e" using `1 \<le> x` Float unfolding le_float_def by auto
- from float_gt1_scale[OF this]
+ proof (rule mult_mono)
+ from float_gt1_scale[OF `1 \<le> Float m e`]
show "0 \<le> real (e + (bitlen m - 1))" by auto
- qed
+ next
+ have "0 \<le> ln 2" by simp
+ thus "0 \<le> real (ub_ln2 prec)" using ub_ln2[of prec] by arith
+ qed auto
ultimately have "ln x \<le> ?ub2 + ?ub_horner"
unfolding Float ln_shifted_float[OF `0 < m`, of e] by auto
}
ultimately show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps
unfolding if_not_P[OF `\<not> x \<le> 0`] if_not_P[OF `\<not> x < 1`] if_not_P[OF False] if_not_P[OF `\<not> x \<le> Float 3 -1`] Let_def
- unfolding scale.simps[of m e, unfolded Float[symmetric]] mantissa.simps[of m e, unfolded Float[symmetric]] real_of_float_add
- by auto
+ unfolding real_of_float_plus e_def[symmetric] m_def[symmetric] by simp
qed
qed
@@ -1975,33 +2037,33 @@
shows "the (lb_ln prec x) \<le> ln x \<and> ln x \<le> the (ub_ln prec x)"
(is "?lb \<le> ?ln \<and> ?ln \<le> ?ub")
proof (cases "x < 1")
- case False hence "1 \<le> x" unfolding less_float_def le_float_def by auto
+ case False hence "1 \<le> x" unfolding less_float_def less_eq_float_def by auto
show ?thesis using ub_ln_lb_ln_bounds'[OF `1 \<le> x`] .
next
- case True have "\<not> x \<le> 0" using `0 < x` unfolding less_float_def le_float_def by auto
-
+ case True have "\<not> x \<le> 0" using `0 < x` unfolding less_float_def less_eq_float_def by auto
+ from True have "real x < 1" by (simp add: less_float_def)
have "0 < real x" and "real x \<noteq> 0" using `0 < x` unfolding less_float_def by auto
hence A: "0 < 1 / real x" by auto
{
let ?divl = "float_divl (max prec 1) 1 x"
- have A': "1 \<le> ?divl" using float_divl_pos_less1_bound[OF `0 < x` `x < 1`] unfolding le_float_def less_float_def by auto
- hence B: "0 < real ?divl" unfolding le_float_def by auto
+ have A': "1 \<le> ?divl" using float_divl_pos_less1_bound[OF `0 < real x` `real x < 1`] unfolding less_eq_float_def less_float_def by auto
+ hence B: "0 < real ?divl" unfolding less_eq_float_def by auto
have "ln ?divl \<le> ln (1 / x)" unfolding ln_le_cancel_iff[OF B A] using float_divl[of _ 1 x] by auto
hence "ln x \<le> - ln ?divl" unfolding nonzero_inverse_eq_divide[OF `real x \<noteq> 0`, symmetric] ln_inverse[OF `0 < real x`] by auto
from this ub_ln_lb_ln_bounds'[OF A', THEN conjunct1, THEN le_imp_neg_le]
- have "?ln \<le> - the (lb_ln prec ?divl)" unfolding real_of_float_minus by (rule order_trans)
+ have "?ln \<le> - the (lb_ln prec ?divl)" unfolding real_of_float_uminus by (rule order_trans)
} moreover
{
let ?divr = "float_divr prec 1 x"
- have A': "1 \<le> ?divr" using float_divr_pos_less1_lower_bound[OF `0 < x` `x < 1`] unfolding le_float_def less_float_def by auto
- hence B: "0 < real ?divr" unfolding le_float_def by auto
+ have A': "1 \<le> ?divr" using float_divr_pos_less1_lower_bound[OF `0 < x` `x < 1`] unfolding less_eq_float_def less_float_def by auto
+ hence B: "0 < real ?divr" unfolding less_eq_float_def by auto
have "ln (1 / x) \<le> ln ?divr" unfolding ln_le_cancel_iff[OF A B] using float_divr[of 1 x] by auto
hence "- ln ?divr \<le> ln x" unfolding nonzero_inverse_eq_divide[OF `real x \<noteq> 0`, symmetric] ln_inverse[OF `0 < real x`] by auto
from ub_ln_lb_ln_bounds'[OF A', THEN conjunct2, THEN le_imp_neg_le] this
- have "- the (ub_ln prec ?divr) \<le> ?ln" unfolding real_of_float_minus by (rule order_trans)
+ have "- the (ub_ln prec ?divr) \<le> ?ln" unfolding real_of_float_uminus by (rule order_trans)
}
ultimately show ?thesis unfolding lb_ln.simps[where x=x] ub_ln.simps[where x=x]
unfolding if_not_P[OF `\<not> x \<le> 0`] if_P[OF True] by auto
@@ -2012,7 +2074,7 @@
proof -
have "0 < x"
proof (rule ccontr)
- assume "\<not> 0 < x" hence "x \<le> 0" unfolding le_float_def less_float_def by auto
+ assume "\<not> 0 < x" hence "x \<le> 0" unfolding less_eq_float_def less_float_def by auto
thus False using assms by auto
qed
thus "0 < real x" unfolding less_float_def by auto
@@ -2025,7 +2087,7 @@
proof -
have "0 < x"
proof (rule ccontr)
- assume "\<not> 0 < x" hence "x \<le> 0" unfolding le_float_def less_float_def by auto
+ assume "\<not> 0 < x" hence "x \<le> 0" unfolding less_eq_float_def less_float_def by auto
thus False using assms by auto
qed
thus "0 < real x" unfolding less_float_def by auto
@@ -2174,12 +2236,12 @@
unfolding bounded_by_def by auto
fun approx approx' :: "nat \<Rightarrow> floatarith \<Rightarrow> (float * float) option list \<Rightarrow> (float * float) option" where
-"approx' prec a bs = (case (approx prec a bs) of Some (l, u) \<Rightarrow> Some (round_down prec l, round_up prec u) | None \<Rightarrow> None)" |
+"approx' prec a bs = (case (approx prec a bs) of Some (l, u) \<Rightarrow> Some (float_round_down prec l, float_round_up prec u) | None \<Rightarrow> None)" |
"approx prec (Add a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs) (\<lambda> l1 u1 l2 u2. (l1 + l2, u1 + u2))" |
"approx prec (Minus a) bs = lift_un' (approx' prec a bs) (\<lambda> l u. (-u, -l))" |
"approx prec (Mult a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs)
- (\<lambda> a1 a2 b1 b2. (float_nprt a1 * float_pprt b2 + float_nprt a2 * float_nprt b2 + float_pprt a1 * float_pprt b1 + float_pprt a2 * float_nprt b1,
- float_pprt a2 * float_pprt b2 + float_pprt a1 * float_nprt b2 + float_nprt a2 * float_pprt b1 + float_nprt a1 * float_nprt b1))" |
+ (\<lambda> a1 a2 b1 b2. (nprt a1 * pprt b2 + nprt a2 * nprt b2 + pprt a1 * pprt b1 + pprt a2 * nprt b1,
+ pprt a2 * pprt b2 + pprt a1 * nprt b2 + nprt a2 * pprt b1 + nprt a1 * nprt b1))" |
"approx prec (Inverse a) bs = lift_un (approx' prec a bs) (\<lambda> l u. if (0 < l \<or> u < 0) then (Some (float_divl prec 1 u), Some (float_divr prec 1 l)) else (None, None))" |
"approx prec (Cos a) bs = lift_un' (approx' prec a bs) (bnds_cos prec)" |
"approx prec Pi bs = Some (lb_pi prec, ub_pi prec)" |
@@ -2233,11 +2295,11 @@
proof -
obtain l' u' where S: "Some (l', u') = approx prec a vs"
using approx' unfolding approx'.simps by (cases "approx prec a vs", auto)
- have l': "l = round_down prec l'" and u': "u = round_up prec u'"
+ have l': "l = float_round_down prec l'" and u': "u = float_round_up prec u'"
using approx' unfolding approx'.simps S[symmetric] by auto
show ?thesis unfolding l' u'
- using order_trans[OF Pa[OF S, THEN conjunct2] round_up[of u']]
- using order_trans[OF round_down[of _ l'] Pa[OF S, THEN conjunct1]] by auto
+ using order_trans[OF Pa[OF S, THEN conjunct2] float_round_up[of u']]
+ using order_trans[OF float_round_down[of _ l'] Pa[OF S, THEN conjunct1]] by auto
qed
lemma lift_bin':
@@ -2394,11 +2456,11 @@
case (Mult a b)
from lift_bin'[OF Mult.prems[unfolded approx.simps]] Mult.hyps
obtain l1 u1 l2 u2
- where l: "l = float_nprt l1 * float_pprt u2 + float_nprt u1 * float_nprt u2 + float_pprt l1 * float_pprt l2 + float_pprt u1 * float_nprt l2"
- and u: "u = float_pprt u1 * float_pprt u2 + float_pprt l1 * float_nprt u2 + float_nprt u1 * float_pprt l2 + float_nprt l1 * float_nprt l2"
+ where l: "l = nprt l1 * pprt u2 + nprt u1 * nprt u2 + pprt l1 * pprt l2 + pprt u1 * nprt l2"
+ and u: "u = pprt u1 * pprt u2 + pprt l1 * nprt u2 + nprt u1 * pprt l2 + nprt l1 * nprt l2"
and "l1 \<le> interpret_floatarith a xs" and "interpret_floatarith a xs \<le> u1"
and "l2 \<le> interpret_floatarith b xs" and "interpret_floatarith b xs \<le> u2" unfolding fst_conv snd_conv by blast
- thus ?case unfolding interpret_floatarith.simps l u real_of_float_add real_of_float_mult real_of_float_nprt real_of_float_pprt
+ thus ?case unfolding interpret_floatarith.simps l u
using mult_le_prts mult_ge_prts by auto
next
case (Inverse a)
@@ -2443,7 +2505,7 @@
from lift_un'[OF Abs.prems[unfolded approx.simps], unfolded fst_conv snd_conv] Abs.hyps
obtain l1 u1 where l': "l = (if l1 < 0 \<and> 0 < u1 then 0 else min \<bar>l1\<bar> \<bar>u1\<bar>)" and u': "u = max \<bar>l1\<bar> \<bar>u1\<bar>"
and l1: "l1 \<le> interpret_floatarith x xs" and u1: "interpret_floatarith x xs \<le> u1" by blast
- thus ?case unfolding l' u' by (cases "l1 < 0 \<and> 0 < u1", auto simp add: real_of_float_min real_of_float_max real_of_float_abs less_float_def)
+ thus ?case unfolding l' u' by (cases "l1 < 0 \<and> 0 < u1", auto simp add: real_of_float_min real_of_float_max less_float_def)
next
case (Min a b)
from lift_bin'[OF Min.prems[unfolded approx.simps], unfolded fst_conv snd_conv] Min.hyps
@@ -2527,7 +2589,7 @@
let ?m = "(l + u) * Float 1 -1"
have "real l \<le> ?m" and "?m \<le> real u"
- unfolding le_float_def using Suc.prems by auto
+ unfolding less_eq_float_def using Suc.prems by auto
with `x \<in> { l .. u }`
have "x \<in> { l .. ?m} \<or> x \<in> { ?m .. u }" by auto
@@ -2576,7 +2638,7 @@
then obtain n
where x_eq: "x = Var n" by (cases x) auto
- with Assign.prems obtain l u' l' u
+ with Assign.prems obtain l u
where bnd_eq: "Some (l, u) = approx prec a vs"
and x_eq: "x = Var n"
and approx_form': "approx_form' prec f (ss ! n) n l u vs ss"
@@ -2612,7 +2674,7 @@
and inequality: "u \<le> l'"
by (cases "approx prec a vs", auto,
cases "approx prec b vs", auto)
- from inequality[unfolded le_float_def] approx[OF LessEqual.prems(2) l_eq] approx[OF LessEqual.prems(2) u_eq]
+ from inequality[unfolded less_eq_float_def] approx[OF LessEqual.prems(2) l_eq] approx[OF LessEqual.prems(2) u_eq]
show ?case by auto
next
case (AtLeastAtMost x a b)
@@ -2624,7 +2686,7 @@
by (cases "approx prec x vs", auto,
cases "approx prec a vs", auto,
cases "approx prec b vs", auto, blast)
- from inequality[unfolded le_float_def] approx[OF AtLeastAtMost.prems(2) l_eq] approx[OF AtLeastAtMost.prems(2) u_eq] approx[OF AtLeastAtMost.prems(2) x_eq]
+ from inequality[unfolded less_eq_float_def] approx[OF AtLeastAtMost.prems(2) l_eq] approx[OF AtLeastAtMost.prems(2) u_eq] approx[OF AtLeastAtMost.prems(2) x_eq]
show ?case by auto
qed
@@ -2685,12 +2747,14 @@
by (auto intro!: DERIV_intros
simp add: algebra_simps power2_eq_square)
next case (Cos a) thus ?case
- by (auto intro!: DERIV_intros
+ apply (auto intro!: DERIV_intros
simp del: interpret_floatarith.simps(5)
simp add: interpret_floatarith_sin interpret_floatarith.simps(5)[of a])
+ apply (simp add: cos_sin_eq)
+ done
next case (Power a n) thus ?case
by (cases n, auto intro!: DERIV_intros
- simp del: power_Suc simp add: real_eq_of_nat)
+ simp del: power_Suc)
next case (Ln a) thus ?case
by (auto intro!: DERIV_intros simp add: divide_inverse)
next case (Var i) thus ?case using `n < length vs` by auto
@@ -3056,7 +3120,7 @@
by (auto simp add: Let_def lazy_conj)
have m_l: "real l \<le> ?m" and m_u: "?m \<le> real u"
- unfolding le_float_def using Suc.prems by auto
+ unfolding less_eq_float_def using Suc.prems by auto
with `x \<in> { l .. u }`
have "x \<in> { l .. ?m} \<or> x \<in> { ?m .. u }" by auto
@@ -3118,7 +3182,7 @@
from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'
have "ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
by (auto simp add: diff_minus)
- from order_trans[OF `0 \<le> ly`[unfolded le_float_def] this]
+ from order_trans[OF `0 \<le> ly`[unfolded less_eq_float_def] this]
show ?thesis by auto
qed
--- a/src/HOL/Library/Float.thy Wed Apr 18 14:29:20 2012 +0200
+++ b/src/HOL/Library/Float.thy Wed Apr 18 14:29:21 2012 +0200
@@ -1,509 +1,925 @@
-(* Title: HOL/Library/Float.thy
- Author: Steven Obua 2008
- Author: Johannes Hoelzl, TU Muenchen <hoelzl@in.tum.de> 2008 / 2009
-*)
-
header {* Floating-Point Numbers *}
theory Float
-imports Complex_Main Lattice_Algebras
+imports Complex_Main "~~/src/HOL/Library/Lattice_Algebras"
begin
-definition pow2 :: "int \<Rightarrow> real" where
- [simp]: "pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))"
-
-datatype float = Float int int
-
-primrec of_float :: "float \<Rightarrow> real" where
- "of_float (Float a b) = real a * pow2 b"
-
-defs (overloaded)
- real_of_float_def [code_unfold]: "real == of_float"
-
-declare [[coercion "% x . Float x 0"]]
-declare [[coercion "real::float\<Rightarrow>real"]]
-
-primrec mantissa :: "float \<Rightarrow> int" where
- "mantissa (Float a b) = a"
-
-primrec scale :: "float \<Rightarrow> int" where
- "scale (Float a b) = b"
-
-instantiation float :: zero
-begin
-definition zero_float where "0 = Float 0 0"
-instance ..
-end
-
-instantiation float :: one
-begin
-definition one_float where "1 = Float 1 0"
-instance ..
-end
-
-lemma real_of_float_simp[simp]: "real (Float a b) = real a * pow2 b"
- unfolding real_of_float_def using of_float.simps .
-
-lemma real_of_float_neg_exp: "e < 0 \<Longrightarrow> real (Float m e) = real m * inverse (2^nat (-e))" by auto
-lemma real_of_float_nge0_exp: "\<not> 0 \<le> e \<Longrightarrow> real (Float m e) = real m * inverse (2^nat (-e))" by auto
-lemma real_of_float_ge0_exp: "0 \<le> e \<Longrightarrow> real (Float m e) = real m * (2^nat e)" by auto
-
-lemma Float_num[simp]: shows
- "real (Float 1 0) = 1" and "real (Float 1 1) = 2" and "real (Float 1 2) = 4" and
- "real (Float 1 -1) = 1/2" and "real (Float 1 -2) = 1/4" and "real (Float 1 -3) = 1/8" and
- "real (Float -1 0) = -1" and "real (Float (numeral n) 0) = numeral n"
+typedef float = "{m * 2 powr e | (m :: int) (e :: int). True }"
+ morphisms real_of_float float_of
by auto
-lemma float_number_of_int[simp]: "real (Float n 0) = real n"
- by simp
+declare [[coercion "real::float\<Rightarrow>real"]]
+
+lemmas float_of_inject[simp]
+lemmas float_of_cases2 = float_of_cases[case_product float_of_cases]
+lemmas float_of_cases3 = float_of_cases2[case_product float_of_cases]
+
+defs (overloaded)
+ real_of_float_def[code_unfold]: "real == real_of_float"
-lemma pow2_0[simp]: "pow2 0 = 1" by simp
-lemma pow2_1[simp]: "pow2 1 = 2" by simp
-lemma pow2_neg: "pow2 x = inverse (pow2 (-x))" by simp
+lemma real_of_float_eq[simp]:
+ fixes f1 f2 :: float shows "real f1 = real f2 \<longleftrightarrow> f1 = f2"
+ unfolding real_of_float_def real_of_float_inject ..
+
+lemma float_of_real[simp]: "float_of (real x) = x"
+ unfolding real_of_float_def by (rule real_of_float_inverse)
-lemma pow2_powr: "pow2 a = 2 powr a"
- by (simp add: powr_realpow[symmetric] powr_minus)
+lemma real_float[simp]: "x \<in> float \<Longrightarrow> real (float_of x) = x"
+ unfolding real_of_float_def by (rule float_of_inverse)
-declare pow2_def[simp del]
+subsection {* Real operations preserving the representation as floating point number *}
+
+lemma floatI: fixes m e :: int shows "m * 2 powr e = x \<Longrightarrow> x \<in> float"
+ by (auto simp: float_def)
-lemma pow2_add: "pow2 (a+b) = (pow2 a) * (pow2 b)"
- by (simp add: pow2_powr powr_add)
+lemma zero_float[simp]: "0 \<in> float" by (auto simp: float_def)
+lemma one_float[simp]: "1 \<in> float" by (intro floatI[of 1 0]) simp
+lemma numeral_float[simp]: "numeral i \<in> float" by (intro floatI[of "numeral i" 0]) simp
+lemma neg_numeral_float[simp]: "neg_numeral i \<in> float" by (intro floatI[of "neg_numeral i" 0]) simp
+lemma real_of_int_float[simp]: "real (x :: int) \<in> float" by (intro floatI[of x 0]) simp
+lemma real_of_nat_float[simp]: "real (x ::nat) \<in> float" by (intro floatI[of x 0]) simp
+lemma two_powr_int_float[simp]: "2 powr (real (i::int)) \<in> float" by (intro floatI[of 1 i]) simp
+lemma two_powr_nat_float[simp]: "2 powr (real (i::nat)) \<in> float" by (intro floatI[of 1 i]) simp
+lemma two_powr_minus_int_float[simp]: "2 powr - (real (i::int)) \<in> float" by (intro floatI[of 1 "-i"]) simp
+lemma two_powr_minus_nat_float[simp]: "2 powr - (real (i::nat)) \<in> float" by (intro floatI[of 1 "-i"]) simp
+lemma two_powr_numeral_float[simp]: "2 powr numeral i \<in> float" by (intro floatI[of 1 "numeral i"]) simp
+lemma two_powr_neg_numeral_float[simp]: "2 powr neg_numeral i \<in> float" by (intro floatI[of 1 "neg_numeral i"]) simp
+lemma two_pow_float[simp]: "2 ^ n \<in> float" by (intro floatI[of 1 "n"]) (simp add: powr_realpow)
+lemma real_of_float_float[simp]: "real (f::float) \<in> float" by (cases f) simp
-lemma pow2_diff: "pow2 (a - b) = pow2 a / pow2 b"
- by (simp add: pow2_powr powr_divide2)
-
-lemma pow2_add1: "pow2 (1 + a) = 2 * (pow2 a)"
- by (simp add: pow2_add)
-
-lemma float_components[simp]: "Float (mantissa f) (scale f) = f" by (cases f) auto
-
-lemma float_split: "\<exists> a b. x = Float a b" by (cases x) auto
+lemma plus_float[simp]: "r \<in> float \<Longrightarrow> p \<in> float \<Longrightarrow> r + p \<in> float"
+ unfolding float_def
+proof (safe, simp)
+ fix e1 m1 e2 m2 :: int
+ { fix e1 m1 e2 m2 :: int assume "e1 \<le> e2"
+ then have "m1 * 2 powr e1 + m2 * 2 powr e2 = (m1 + m2 * 2 ^ nat (e2 - e1)) * 2 powr e1"
+ by (simp add: powr_realpow[symmetric] powr_divide2[symmetric] field_simps)
+ then have "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e"
+ by blast }
+ note * = this
+ show "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e"
+ proof (cases e1 e2 rule: linorder_le_cases)
+ assume "e2 \<le> e1" from *[OF this, of m2 m1] show ?thesis by (simp add: ac_simps)
+ qed (rule *)
+qed
-lemma float_split2: "(\<forall> a b. x \<noteq> Float a b) = False" by (auto simp add: float_split)
+lemma uminus_float[simp]: "x \<in> float \<Longrightarrow> -x \<in> float"
+ apply (auto simp: float_def)
+ apply (rule_tac x="-x" in exI)
+ apply (rule_tac x="xa" in exI)
+ apply (simp add: field_simps)
+ done
-lemma float_zero[simp]: "real (Float 0 e) = 0" by simp
+lemma times_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> x * y \<in> float"
+ apply (auto simp: float_def)
+ apply (rule_tac x="x * xa" in exI)
+ apply (rule_tac x="xb + xc" in exI)
+ apply (simp add: powr_add)
+ done
-lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> abs((a::int) div 2) < abs a"
-by arith
+lemma minus_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> x - y \<in> float"
+ unfolding ab_diff_minus by (intro uminus_float plus_float)
+
+lemma abs_float[simp]: "x \<in> float \<Longrightarrow> abs x \<in> float"
+ by (cases x rule: linorder_cases[of 0]) auto
+
+lemma sgn_of_float[simp]: "x \<in> float \<Longrightarrow> sgn x \<in> float"
+ by (cases x rule: linorder_cases[of 0]) (auto intro!: uminus_float)
-function normfloat :: "float \<Rightarrow> float" where
- "normfloat (Float a b) =
- (if a \<noteq> 0 \<and> even a then normfloat (Float (a div 2) (b+1))
- else if a=0 then Float 0 0 else Float a b)"
-by pat_completeness auto
-termination by (relation "measure (nat o abs o mantissa)") (auto intro: abs_div_2_less)
-declare normfloat.simps[simp del]
+lemma div_power_2_float[simp]: "x \<in> float \<Longrightarrow> x / 2^d \<in> float"
+ apply (auto simp add: float_def)
+ apply (rule_tac x="x" in exI)
+ apply (rule_tac x="xa - d" in exI)
+ apply (simp add: powr_realpow[symmetric] field_simps powr_add[symmetric])
+ done
+
+lemma div_power_2_int_float[simp]: "x \<in> float \<Longrightarrow> x / (2::int)^d \<in> float"
+ apply (auto simp add: float_def)
+ apply (rule_tac x="x" in exI)
+ apply (rule_tac x="xa - d" in exI)
+ apply (simp add: powr_realpow[symmetric] field_simps powr_add[symmetric])
+ done
-theorem normfloat[symmetric, simp]: "real f = real (normfloat f)"
-proof (induct f rule: normfloat.induct)
- case (1 a b)
- have real2: "2 = real (2::int)"
- by auto
- show ?case
- apply (subst normfloat.simps)
- apply auto
- apply (subst 1[symmetric])
- apply (auto simp add: pow2_add even_def)
- done
+lemma div_numeral_Bit0_float[simp]:
+ assumes x: "x / numeral n \<in> float" shows "x / (numeral (Num.Bit0 n)) \<in> float"
+proof -
+ have "(x / numeral n) / 2^1 \<in> float"
+ by (intro x div_power_2_float)
+ also have "(x / numeral n) / 2^1 = x / (numeral (Num.Bit0 n))"
+ by (induct n) auto
+ finally show ?thesis .
+qed
+
+lemma div_neg_numeral_Bit0_float[simp]:
+ assumes x: "x / numeral n \<in> float" shows "x / (neg_numeral (Num.Bit0 n)) \<in> float"
+proof -
+ have "- (x / numeral (Num.Bit0 n)) \<in> float" using x by simp
+ also have "- (x / numeral (Num.Bit0 n)) = x / neg_numeral (Num.Bit0 n)"
+ unfolding neg_numeral_def by (simp del: minus_numeral)
+ finally show ?thesis .
qed
-lemma pow2_neq_zero[simp]: "pow2 x \<noteq> 0"
- by (auto simp add: pow2_def)
+subsection {* Arithmetic operations on floating point numbers *}
+
+instantiation float :: ring_1
+begin
+
+definition [simp]: "(0::float) = float_of 0"
+
+definition [simp]: "(1::float) = float_of 1"
+
+definition "(x + y::float) = float_of (real x + real y)"
+
+lemma float_plus_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> float_of x + float_of y = float_of (x + y)"
+ by (simp add: plus_float_def)
-lemma pow2_int: "pow2 (int c) = 2^c"
- by (simp add: pow2_def)
+definition "(-x::float) = float_of (- real x)"
+
+lemma uminus_of_float[simp]: "x \<in> float \<Longrightarrow> - float_of x = float_of (- x)"
+ by (simp add: uminus_float_def)
+
+definition "(x - y::float) = float_of (real x - real y)"
-lemma zero_less_pow2[simp]: "0 < pow2 x"
- by (simp add: pow2_powr)
+lemma float_minus_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> float_of x - float_of y = float_of (x - y)"
+ by (simp add: minus_float_def)
+
+definition "(x * y::float) = float_of (real x * real y)"
+
+lemma float_times_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> float_of x * float_of y = float_of (x * y)"
+ by (simp add: times_float_def)
-lemma normfloat_imp_odd_or_zero:
- "normfloat f = Float a b \<Longrightarrow> odd a \<or> (a = 0 \<and> b = 0)"
-proof (induct f rule: normfloat.induct)
- case (1 u v)
- from 1 have ab: "normfloat (Float u v) = Float a b" by auto
- {
- assume eu: "even u"
- assume z: "u \<noteq> 0"
- have "normfloat (Float u v) = normfloat (Float (u div 2) (v + 1))"
- apply (subst normfloat.simps)
- by (simp add: eu z)
- with ab have "normfloat (Float (u div 2) (v + 1)) = Float a b" by simp
- with 1 eu z have ?case by auto
- }
- note case1 = this
- {
- assume "odd u \<or> u = 0"
- then have ou: "\<not> (u \<noteq> 0 \<and> even u)" by auto
- have "normfloat (Float u v) = (if u = 0 then Float 0 0 else Float u v)"
- apply (subst normfloat.simps)
- apply (simp add: ou)
- done
- with ab have "Float a b = (if u = 0 then Float 0 0 else Float u v)" by auto
- then have ?case
- apply (case_tac "u=0")
- apply (auto)
- by (insert ou, auto)
- }
- note case2 = this
- show ?case
- apply (case_tac "odd u \<or> u = 0")
- apply (rule case2)
- apply simp
- apply (rule case1)
- apply auto
- done
+instance
+proof
+ fix a b c :: float
+ show "0 + a = a"
+ by (cases a rule: float_of_cases) simp
+ show "1 * a = a"
+ by (cases a rule: float_of_cases) simp
+ show "a * 1 = a"
+ by (cases a rule: float_of_cases) simp
+ show "-a + a = 0"
+ by (cases a rule: float_of_cases) simp
+ show "a + b = b + a"
+ by (cases a b rule: float_of_cases2) (simp add: ac_simps)
+ show "a - b = a + -b"
+ by (cases a b rule: float_of_cases2) (simp add: field_simps)
+ show "a + b + c = a + (b + c)"
+ by (cases a b c rule: float_of_cases3) (simp add: ac_simps)
+ show "a * b * c = a * (b * c)"
+ by (cases a b c rule: float_of_cases3) (simp add: ac_simps)
+ show "(a + b) * c = a * c + b * c"
+ by (cases a b c rule: float_of_cases3) (simp add: field_simps)
+ show "a * (b + c) = a * b + a * c"
+ by (cases a b c rule: float_of_cases3) (simp add: field_simps)
+ show "0 \<noteq> (1::float)" by simp
qed
+end
-lemma float_eq_odd_helper:
- assumes odd: "odd a'"
- and floateq: "real (Float a b) = real (Float a' b')"
- shows "b \<le> b'"
-proof -
- from odd have "a' \<noteq> 0" by auto
- with floateq have a': "real a' = real a * pow2 (b - b')"
- by (simp add: pow2_diff field_simps)
+lemma real_of_float_uminus[simp]:
+ fixes f g::float shows "real (- g) = - real g"
+ by (simp add: uminus_float_def)
+
+lemma real_of_float_plus[simp]:
+ fixes f g::float shows "real (f + g) = real f + real g"
+ by (simp add: plus_float_def)
+
+lemma real_of_float_minus[simp]:
+ fixes f g::float shows "real (f - g) = real f - real g"
+ by (simp add: minus_float_def)
+
+lemma real_of_float_times[simp]:
+ fixes f g::float shows "real (f * g) = real f * real g"
+ by (simp add: times_float_def)
+
+lemma real_of_float_zero[simp]: "real (0::float) = 0" by simp
+lemma real_of_float_one[simp]: "real (1::float) = 1" by simp
+
+lemma real_of_float_power[simp]: fixes f::float shows "real (f^n) = real f^n"
+ by (induct n) simp_all
- {
- assume bcmp: "b > b'"
- then obtain c :: nat where "b - b' = int c + 1"
- by atomize_elim arith
- with a' have "real a' = real (a * 2^c * 2)"
- by (simp add: pow2_def nat_add_distrib)
- with odd have False
- unfolding real_of_int_inject by simp
- }
- then show ?thesis by arith
-qed
+instantiation float :: linorder
+begin
+
+definition "x \<le> (y::float) \<longleftrightarrow> real x \<le> real y"
+
+lemma float_le_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> float_of x \<le> float_of y \<longleftrightarrow> x \<le> y"
+ by (simp add: less_eq_float_def)
+
+definition "x < (y::float) \<longleftrightarrow> real x < real y"
+
+lemma float_less_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> float_of x < float_of y \<longleftrightarrow> x < y"
+ by (simp add: less_float_def)
-lemma float_eq_odd:
- assumes odd1: "odd a"
- and odd2: "odd a'"
- and floateq: "real (Float a b) = real (Float a' b')"
- shows "a = a' \<and> b = b'"
-proof -
- from
- float_eq_odd_helper[OF odd2 floateq]
- float_eq_odd_helper[OF odd1 floateq[symmetric]]
- have beq: "b = b'" by arith
- with floateq show ?thesis by auto
+instance
+proof
+ fix a b c :: float
+ show "a \<le> a"
+ by (cases a rule: float_of_cases) simp
+ show "a < b \<longleftrightarrow> a \<le> b \<and> \<not> b \<le> a"
+ by (cases a b rule: float_of_cases2) auto
+ show "a \<le> b \<or> b \<le> a"
+ by (cases a b rule: float_of_cases2) auto
+ { assume "a \<le> b" "b \<le> a" then show "a = b"
+ by (cases a b rule: float_of_cases2) auto }
+ { assume "a \<le> b" "b \<le> c" then show "a \<le> c"
+ by (cases a b c rule: float_of_cases3) auto }
qed
+end
+
+lemma real_of_float_min: fixes a b::float shows "real (min a b) = min (real a) (real b)"
+ by (simp add: min_def less_eq_float_def)
+
+lemma real_of_float_max: fixes a b::float shows "real (max a b) = max (real a) (real b)"
+ by (simp add: max_def less_eq_float_def)
+
+instantiation float :: linordered_ring
+begin
+
+definition "(abs x :: float) = float_of (abs (real x))"
-theorem normfloat_unique:
- assumes real_of_float_eq: "real f = real g"
- shows "normfloat f = normfloat g"
-proof -
- from float_split[of "normfloat f"] obtain a b where normf:"normfloat f = Float a b" by auto
- from float_split[of "normfloat g"] obtain a' b' where normg:"normfloat g = Float a' b'" by auto
- have "real (normfloat f) = real (normfloat g)"
- by (simp add: real_of_float_eq)
- then have float_eq: "real (Float a b) = real (Float a' b')"
- by (simp add: normf normg)
- have ab: "odd a \<or> (a = 0 \<and> b = 0)" by (rule normfloat_imp_odd_or_zero[OF normf])
- have ab': "odd a' \<or> (a' = 0 \<and> b' = 0)" by (rule normfloat_imp_odd_or_zero[OF normg])
- {
- assume odd: "odd a"
- then have "a \<noteq> 0" by (simp add: even_def) arith
- with float_eq have "a' \<noteq> 0" by auto
- with ab' have "odd a'" by simp
- from odd this float_eq have "a = a' \<and> b = b'" by (rule float_eq_odd)
- }
- note odd_case = this
- {
- assume even: "even a"
- with ab have a0: "a = 0" by simp
- with float_eq have a0': "a' = 0" by auto
- from a0 a0' ab ab' have "a = a' \<and> b = b'" by auto
- }
- note even_case = this
- from odd_case even_case show ?thesis
- apply (simp add: normf normg)
- apply (case_tac "even a")
- apply auto
+lemma float_abs[simp]: "x \<in> float \<Longrightarrow> abs (float_of x) = float_of (abs x)"
+ by (simp add: abs_float_def)
+
+instance
+proof
+ fix a b c :: float
+ show "\<bar>a\<bar> = (if a < 0 then - a else a)"
+ by (cases a rule: float_of_cases) simp
+ assume "a \<le> b"
+ then show "c + a \<le> c + b"
+ by (cases a b c rule: float_of_cases3) simp
+ assume "0 \<le> c"
+ with `a \<le> b` show "c * a \<le> c * b"
+ by (cases a b c rule: float_of_cases3) (auto intro: mult_left_mono)
+ from `0 \<le> c` `a \<le> b` show "a * c \<le> b * c"
+ by (cases a b c rule: float_of_cases3) (auto intro: mult_right_mono)
+qed
+end
+
+lemma real_of_abs_float[simp]: fixes f::float shows "real (abs f) = abs (real f)"
+ unfolding abs_float_def by simp
+
+instance float :: dense_linorder
+proof
+ fix a b :: float
+ show "\<exists>c. a < c"
+ apply (intro exI[of _ "a + 1"])
+ apply (cases a rule: float_of_cases)
+ apply simp
+ done
+ show "\<exists>c. c < a"
+ apply (intro exI[of _ "a - 1"])
+ apply (cases a rule: float_of_cases)
+ apply simp
+ done
+ assume "a < b"
+ then show "\<exists>c. a < c \<and> c < b"
+ apply (intro exI[of _ "(b + a) * float_of (1/2)"])
+ apply (cases a b rule: float_of_cases2)
+ apply simp
done
qed
-instantiation float :: plus
-begin
-fun plus_float where
-[simp del]: "(Float a_m a_e) + (Float b_m b_e) =
- (if a_e \<le> b_e then Float (a_m + b_m * 2^(nat(b_e - a_e))) a_e
- else Float (a_m * 2^(nat (a_e - b_e)) + b_m) b_e)"
-instance ..
-end
-
-instantiation float :: uminus
-begin
-primrec uminus_float where [simp del]: "uminus_float (Float m e) = Float (-m) e"
-instance ..
-end
-
-instantiation float :: minus
+instantiation float :: linordered_idom
begin
-definition minus_float where "(z::float) - w = z + (- w)"
-instance ..
-end
+
+definition "sgn x = float_of (sgn (real x))"
-instantiation float :: times
-begin
-fun times_float where [simp del]: "(Float a_m a_e) * (Float b_m b_e) = Float (a_m * b_m) (a_e + b_e)"
-instance ..
-end
+lemma sgn_float[simp]: "x \<in> float \<Longrightarrow> sgn (float_of x) = float_of (sgn x)"
+ by (simp add: sgn_float_def)
-primrec float_pprt :: "float \<Rightarrow> float" where
- "float_pprt (Float a e) = (if 0 <= a then (Float a e) else 0)"
-
-primrec float_nprt :: "float \<Rightarrow> float" where
- "float_nprt (Float a e) = (if 0 <= a then 0 else (Float a e))"
-
-instantiation float :: ord
-begin
-definition le_float_def: "z \<le> (w :: float) \<equiv> real z \<le> real w"
-definition less_float_def: "z < (w :: float) \<equiv> real z < real w"
-instance ..
+instance
+proof
+ fix a b c :: float
+ show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
+ by (cases a rule: float_of_cases) simp
+ show "a * b = b * a"
+ by (cases a b rule: float_of_cases2) (simp add: field_simps)
+ show "1 * a = a" "(a + b) * c = a * c + b * c"
+ by (simp_all add: field_simps del: one_float_def)
+ assume "a < b" "0 < c" then show "c * a < c * b"
+ by (cases a b c rule: float_of_cases3) (simp add: field_simps)
+qed
end
-lemma real_of_float_add[simp]: "real (a + b) = real a + real (b :: float)"
- by (cases a, cases b) (simp add: algebra_simps plus_float.simps,
- auto simp add: pow2_int[symmetric] pow2_add[symmetric])
-
-lemma real_of_float_minus[simp]: "real (- a) = - real (a :: float)"
- by (cases a) simp
-
-lemma real_of_float_sub[simp]: "real (a - b) = real a - real (b :: float)"
- by (cases a, cases b) (simp add: minus_float_def)
-
-lemma real_of_float_mult[simp]: "real (a*b) = real a * real (b :: float)"
- by (cases a, cases b) (simp add: times_float.simps pow2_add)
-
-lemma real_of_float_0[simp]: "real (0 :: float) = 0"
- by (auto simp add: zero_float_def)
+definition Float :: "int \<Rightarrow> int \<Rightarrow> float" where
+ [simp, code del]: "Float m e = float_of (m * 2 powr e)"
-lemma real_of_float_1[simp]: "real (1 :: float) = 1"
- by (auto simp add: one_float_def)
+lemma real_of_float_Float[code]: "real_of_float (Float m e) =
+ (if e \<ge> 0 then m * 2 ^ nat e else m * inverse (2 ^ nat (- e)))"
+by (auto simp add: powr_realpow[symmetric] powr_minus real_of_float_def[symmetric])
-lemma zero_le_float:
- "(0 <= real (Float a b)) = (0 <= a)"
- apply auto
- apply (auto simp add: zero_le_mult_iff)
- apply (insert zero_less_pow2[of b])
- apply (simp_all)
- done
+code_datatype Float
-lemma float_le_zero:
- "(real (Float a b) <= 0) = (a <= 0)"
- apply auto
- apply (auto simp add: mult_le_0_iff)
- apply (insert zero_less_pow2[of b])
- apply auto
- done
+lemma real_Float: "real (Float m e) = m * 2 powr e" by simp
-lemma zero_less_float:
- "(0 < real (Float a b)) = (0 < a)"
- apply auto
- apply (auto simp add: zero_less_mult_iff)
- apply (insert zero_less_pow2[of b])
- apply (simp_all)
- done
-
-lemma float_less_zero:
- "(real (Float a b) < 0) = (a < 0)"
- apply auto
- apply (auto simp add: mult_less_0_iff)
- apply (insert zero_less_pow2[of b])
- apply (simp_all)
- done
-
-declare real_of_float_simp[simp del]
+definition normfloat :: "float \<Rightarrow> float" where
+ [simp]: "normfloat x = x"
-lemma real_of_float_pprt[simp]: "real (float_pprt a) = pprt (real a)"
- by (cases a) (auto simp add: zero_le_float float_le_zero)
-
-lemma real_of_float_nprt[simp]: "real (float_nprt a) = nprt (real a)"
- by (cases a) (auto simp add: zero_le_float float_le_zero)
+lemma compute_normfloat[code]: "normfloat (Float m e) =
+ (if m mod 2 = 0 \<and> m \<noteq> 0 then normfloat (Float (m div 2) (e + 1))
+ else if m = 0 then 0 else Float m e)"
+ by (simp del: real_of_int_add split: prod.split)
+ (auto simp add: powr_add zmod_eq_0_iff)
-instance float :: ab_semigroup_add
-proof (intro_classes)
- fix a b c :: float
- show "a + b + c = a + (b + c)"
- by (cases a, cases b, cases c)
- (auto simp add: algebra_simps power_add[symmetric] plus_float.simps)
-next
- fix a b :: float
- show "a + b = b + a"
- by (cases a, cases b) (simp add: plus_float.simps)
-qed
+lemma compute_zero[code_unfold, code]: "0 = Float 0 0"
+ by simp
+
+lemma compute_one[code_unfold, code]: "1 = Float 1 0"
+ by simp
instance float :: numeral ..
-lemma Float_add_same_scale: "Float x e + Float y e = Float (x + y) e"
- by (simp add: plus_float.simps)
+lemma float_of_numeral[simp]: "numeral k = float_of (numeral k)"
+ by (induct k)
+ (simp_all only: numeral.simps one_float_def float_plus_float numeral_float one_float plus_float)
+
+lemma float_of_neg_numeral[simp]: "neg_numeral k = float_of (neg_numeral k)"
+ by (simp add: minus_numeral[symmetric] del: minus_numeral)
-(* FIXME: define other constant for code_unfold_post *)
-lemma numeral_float_Float (*[code_unfold_post]*):
- "numeral k = Float (numeral k) 0"
- by (induct k, simp_all only: numeral.simps one_float_def
- Float_add_same_scale)
+lemma
+ shows float_numeral[simp]: "real (numeral x :: float) = numeral x"
+ and float_neg_numeral[simp]: "real (neg_numeral x :: float) = neg_numeral x"
+ by simp_all
-lemma float_number_of[simp]: "real (numeral x :: float) = numeral x"
- by (simp only: numeral_float_Float Float_num)
+subsection {* Represent floats as unique mantissa and exponent *}
-instance float :: comm_monoid_mult
-proof (intro_classes)
- fix a b c :: float
- show "a * b * c = a * (b * c)"
- by (cases a, cases b, cases c) (simp add: times_float.simps)
-next
- fix a b :: float
- show "a * b = b * a"
- by (cases a, cases b) (simp add: times_float.simps)
-next
- fix a :: float
- show "1 * a = a"
- by (cases a) (simp add: times_float.simps one_float_def)
+lemma int_induct_abs[case_names less]:
+ fixes j :: int
+ assumes H: "\<And>n. (\<And>i. \<bar>i\<bar> < \<bar>n\<bar> \<Longrightarrow> P i) \<Longrightarrow> P n"
+ shows "P j"
+proof (induct "nat \<bar>j\<bar>" arbitrary: j rule: less_induct)
+ case less show ?case by (rule H[OF less]) simp
+qed
+
+lemma int_cancel_factors:
+ fixes n :: int assumes "1 < r" shows "n = 0 \<or> (\<exists>k i. n = k * r ^ i \<and> \<not> r dvd k)"
+proof (induct n rule: int_induct_abs)
+ case (less n)
+ { fix m assume n: "n \<noteq> 0" "n = m * r"
+ then have "\<bar>m \<bar> < \<bar>n\<bar>"
+ by (metis abs_dvd_iff abs_ge_self assms comm_semiring_1_class.normalizing_semiring_rules(7)
+ dvd_imp_le_int dvd_refl dvd_triv_right linorder_neq_iff linorder_not_le
+ mult_eq_0_iff zdvd_mult_cancel1)
+ from less[OF this] n have "\<exists>k i. n = k * r ^ Suc i \<and> \<not> r dvd k" by auto }
+ then show ?case
+ by (metis comm_semiring_1_class.normalizing_semiring_rules(12,7) dvdE power_0)
+qed
+
+lemma mult_powr_eq_mult_powr_iff_asym:
+ fixes m1 m2 e1 e2 :: int
+ assumes m1: "\<not> 2 dvd m1" and "e1 \<le> e2"
+ shows "m1 * 2 powr e1 = m2 * 2 powr e2 \<longleftrightarrow> m1 = m2 \<and> e1 = e2"
+proof
+ have "m1 \<noteq> 0" using m1 unfolding dvd_def by auto
+ assume eq: "m1 * 2 powr e1 = m2 * 2 powr e2"
+ with `e1 \<le> e2` have "m1 = m2 * 2 powr nat (e2 - e1)"
+ by (simp add: powr_divide2[symmetric] field_simps)
+ also have "\<dots> = m2 * 2^nat (e2 - e1)"
+ by (simp add: powr_realpow)
+ finally have m1_eq: "m1 = m2 * 2^nat (e2 - e1)"
+ unfolding real_of_int_inject .
+ with m1 have "m1 = m2"
+ by (cases "nat (e2 - e1)") (auto simp add: dvd_def)
+ then show "m1 = m2 \<and> e1 = e2"
+ using eq `m1 \<noteq> 0` by (simp add: powr_inj)
+qed simp
+
+lemma mult_powr_eq_mult_powr_iff:
+ fixes m1 m2 e1 e2 :: int
+ 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"
+ using mult_powr_eq_mult_powr_iff_asym[of m1 e1 e2 m2]
+ using mult_powr_eq_mult_powr_iff_asym[of m2 e2 e1 m1]
+ by (cases e1 e2 rule: linorder_le_cases) auto
+
+lemma floatE_normed:
+ assumes x: "x \<in> float"
+ obtains (zero) "x = 0"
+ | (powr) m e :: int where "x = m * 2 powr e" "\<not> 2 dvd m" "x \<noteq> 0"
+proof atomize_elim
+ { assume "x \<noteq> 0"
+ from x obtain m e :: int where x: "x = m * 2 powr e" by (auto simp: float_def)
+ with `x \<noteq> 0` int_cancel_factors[of 2 m] obtain k i where "m = k * 2 ^ i" "\<not> 2 dvd k"
+ by auto
+ with `\<not> 2 dvd k` x have "\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m"
+ by (rule_tac exI[of _ "k"], rule_tac exI[of _ "e + int i"])
+ (simp add: powr_add powr_realpow) }
+ then show "x = 0 \<or> (\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m \<and> x \<noteq> 0)"
+ by blast
+qed
+
+lemma float_normed_cases:
+ fixes f :: float
+ obtains (zero) "f = 0"
+ | (powr) m e :: int where "real f = m * 2 powr e" "\<not> 2 dvd m" "f \<noteq> 0"
+proof (atomize_elim, induct f)
+ case (float_of y) then show ?case
+ by (cases rule: floatE_normed) auto
+qed
+
+definition mantissa :: "float \<Rightarrow> int" where
+ "mantissa f = fst (SOME p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0)
+ \<or> (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p))"
+
+definition exponent :: "float \<Rightarrow> int" where
+ "exponent f = snd (SOME p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0)
+ \<or> (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p))"
+
+lemma
+ shows exponent_0[simp]: "exponent (float_of 0) = 0" (is ?E)
+ and mantissa_0[simp]: "mantissa (float_of 0) = 0" (is ?M)
+proof -
+ have "\<And>p::int \<times> int. fst p = 0 \<and> snd p = 0 \<longleftrightarrow> p = (0, 0)" by auto
+ then show ?E ?M
+ by (auto simp add: mantissa_def exponent_def)
qed
-(* Floats do NOT form a cancel_semigroup_add: *)
-lemma "0 + Float 0 1 = 0 + Float 0 2"
- by (simp add: plus_float.simps zero_float_def)
+lemma
+ shows mantissa_exponent: "real f = mantissa f * 2 powr exponent f" (is ?E)
+ and mantissa_not_dvd: "f \<noteq> (float_of 0) \<Longrightarrow> \<not> 2 dvd mantissa f" (is "_ \<Longrightarrow> ?D")
+proof cases
+ assume [simp]: "f \<noteq> (float_of 0)"
+ have "f = mantissa f * 2 powr exponent f \<and> \<not> 2 dvd mantissa f"
+ proof (cases f rule: float_normed_cases)
+ case (powr m e)
+ then have "\<exists>p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0)
+ \<or> (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p)"
+ by auto
+ then show ?thesis
+ unfolding exponent_def mantissa_def
+ by (rule someI2_ex) simp
+ qed simp
+ then show ?E ?D by auto
+qed simp
+
+lemma mantissa_noteq_0: "f \<noteq> float_of 0 \<Longrightarrow> mantissa f \<noteq> 0"
+ using mantissa_not_dvd[of f] by auto
+
+lemma Float_mantissa_exponent: "Float (mantissa f) (exponent f) = f"
+ unfolding real_of_float_eq[symmetric] mantissa_exponent[of f] by simp
+
+lemma Float_cases[case_names Float, cases type: float]:
+ fixes f :: float
+ obtains (Float) m e :: int where "f = Float m e"
+ using Float_mantissa_exponent[symmetric]
+ by (atomize_elim) auto
-instance float :: comm_semiring
-proof (intro_classes)
- fix a b c :: float
- show "(a + b) * c = a * c + b * c"
- by (cases a, cases b, cases c) (simp add: plus_float.simps times_float.simps algebra_simps)
+lemma
+ fixes m e :: int
+ defines "f \<equiv> float_of (m * 2 powr e)"
+ assumes dvd: "\<not> 2 dvd m"
+ shows mantissa_float: "mantissa f = m" (is "?M")
+ and exponent_float: "m \<noteq> 0 \<Longrightarrow> exponent f = e" (is "_ \<Longrightarrow> ?E")
+proof cases
+ assume "m = 0" with dvd show "mantissa f = m" by auto
+next
+ assume "m \<noteq> 0"
+ then have f_not_0: "f \<noteq> float_of 0" by (simp add: f_def)
+ from mantissa_exponent[of f]
+ have "m * 2 powr e = mantissa f * 2 powr exponent f"
+ by (auto simp add: f_def)
+ then show "?M" "?E"
+ using mantissa_not_dvd[OF f_not_0] dvd
+ by (auto simp: mult_powr_eq_mult_powr_iff)
+qed
+
+lemma denormalize_shift:
+ assumes f_def: "f \<equiv> Float m e" and not_0: "f \<noteq> float_of 0"
+ obtains i where "m = mantissa f * 2 ^ i" "e = exponent f - i"
+proof
+ from mantissa_exponent[of f] f_def
+ have "m * 2 powr e = mantissa f * 2 powr exponent f"
+ by simp
+ then have eq: "m = mantissa f * 2 powr (exponent f - e)"
+ by (simp add: powr_divide2[symmetric] field_simps)
+ moreover
+ have "e \<le> exponent f"
+ proof (rule ccontr)
+ assume "\<not> e \<le> exponent f"
+ then have pos: "exponent f < e" by simp
+ then have "2 powr (exponent f - e) = 2 powr - real (e - exponent f)"
+ by simp
+ also have "\<dots> = 1 / 2^nat (e - exponent f)"
+ using pos by (simp add: powr_realpow[symmetric] powr_divide2[symmetric])
+ finally have "m * 2^nat (e - exponent f) = real (mantissa f)"
+ using eq by simp
+ then have "mantissa f = m * 2^nat (e - exponent f)"
+ unfolding real_of_int_inject by simp
+ with `exponent f < e` have "2 dvd mantissa f"
+ apply (intro dvdI[where k="m * 2^(nat (e-exponent f)) div 2"])
+ apply (cases "nat (e - exponent f)")
+ apply auto
+ done
+ then show False using mantissa_not_dvd[OF not_0] by simp
+ qed
+ ultimately have "real m = mantissa f * 2^nat (exponent f - e)"
+ by (simp add: powr_realpow[symmetric])
+ with `e \<le> exponent f`
+ show "m = mantissa f * 2 ^ nat (exponent f - e)" "e = exponent f - nat (exponent f - e)"
+ unfolding real_of_int_inject by auto
qed
-(* Floats do NOT form an order, because "(x < y) = (x <= y & x <> y)" does NOT hold *)
+subsection {* Compute arithmetic operations *}
+
+lemma compute_float_numeral[code_abbrev]: "Float (numeral k) 0 = numeral k"
+ by simp
+
+lemma compute_float_neg_numeral[code_abbrev]: "Float (neg_numeral k) 0 = neg_numeral k"
+ by simp
+
+lemma compute_float_uminus[code]: "- Float m1 e1 = Float (- m1) e1"
+ by simp
+
+lemma compute_float_times[code]: "Float m1 e1 * Float m2 e2 = Float (m1 * m2) (e1 + e2)"
+ by (simp add: field_simps powr_add)
+
+lemma compute_float_plus[code]: "Float m1 e1 + Float m2 e2 =
+ (if e1 \<le> e2 then Float (m1 + m2 * 2^nat (e2 - e1)) e1
+ else Float (m2 + m1 * 2^nat (e1 - e2)) e2)"
+ by (simp add: field_simps)
+ (simp add: powr_realpow[symmetric] powr_divide2[symmetric])
+
+lemma compute_float_minus[code]: fixes f g::float shows "f - g = f + (-g)" by simp
+
+lemma compute_float_sgn[code]: "sgn (Float m1 e1) = (if 0 < m1 then 1 else if m1 < 0 then -1 else 0)"
+ by (simp add: sgn_times)
+
+definition is_float_pos :: "float \<Rightarrow> bool" where
+ "is_float_pos f \<longleftrightarrow> 0 < f"
+
+lemma compute_is_float_pos[code]: "is_float_pos (Float m e) \<longleftrightarrow> 0 < m"
+ by (auto simp add: is_float_pos_def zero_less_mult_iff) (simp add: not_le[symmetric])
+
+lemma compute_float_less[code]: "a < b \<longleftrightarrow> is_float_pos (b - a)"
+ by (simp add: is_float_pos_def field_simps del: zero_float_def)
+
+definition is_float_nonneg :: "float \<Rightarrow> bool" where
+ "is_float_nonneg f \<longleftrightarrow> 0 \<le> f"
+
+lemma compute_is_float_nonneg[code]: "is_float_nonneg (Float m e) \<longleftrightarrow> 0 \<le> m"
+ by (auto simp add: is_float_nonneg_def zero_le_mult_iff) (simp add: not_less[symmetric])
+
+lemma compute_float_le[code]: "a \<le> b \<longleftrightarrow> is_float_nonneg (b - a)"
+ by (simp add: is_float_nonneg_def field_simps del: zero_float_def)
+
+definition is_float_zero :: "float \<Rightarrow> bool" where
+ "is_float_zero f \<longleftrightarrow> 0 = f"
+
+lemma compute_is_float_zero[code]: "is_float_zero (Float m e) \<longleftrightarrow> 0 = m"
+ by (auto simp add: is_float_zero_def)
-instance float :: zero_neq_one
-proof (intro_classes)
- show "(0::float) \<noteq> 1"
- by (simp add: zero_float_def one_float_def)
+lemma compute_float_abs[code]: "abs (Float m e) = Float (abs m) e" by (simp add: abs_mult)
+
+instantiation float :: equal
+begin
+
+definition "equal_float (f1 :: float) f2 \<longleftrightarrow> is_float_zero (f1 - f2)"
+
+instance proof qed (auto simp: equal_float_def is_float_zero_def simp del: zero_float_def)
+end
+
+subsection {* Rounding Real numbers *}
+
+definition round_down :: "int \<Rightarrow> real \<Rightarrow> real" where
+ "round_down prec x = floor (x * 2 powr prec) * 2 powr -prec"
+
+definition round_up :: "int \<Rightarrow> real \<Rightarrow> real" where
+ "round_up prec x = ceiling (x * 2 powr prec) * 2 powr -prec"
+
+lemma round_down_float[simp]: "round_down prec x \<in> float"
+ unfolding round_down_def
+ by (auto intro!: times_float simp: real_of_int_minus[symmetric] simp del: real_of_int_minus)
+
+lemma round_up_float[simp]: "round_up prec x \<in> float"
+ unfolding round_up_def
+ by (auto intro!: times_float simp: real_of_int_minus[symmetric] simp del: real_of_int_minus)
+
+lemma round_up: "x \<le> round_up prec x"
+ by (simp add: powr_minus_divide le_divide_eq round_up_def)
+
+lemma round_down: "round_down prec x \<le> x"
+ by (simp add: powr_minus_divide divide_le_eq round_down_def)
+
+lemma round_up_0[simp]: "round_up p 0 = 0"
+ unfolding round_up_def by simp
+
+lemma round_down_0[simp]: "round_down p 0 = 0"
+ unfolding round_down_def by simp
+
+lemma round_up_diff_round_down:
+ "round_up prec x - round_down prec x \<le> 2 powr -prec"
+proof -
+ have "round_up prec x - round_down prec x =
+ (ceiling (x * 2 powr prec) - floor (x * 2 powr prec)) * 2 powr -prec"
+ by (simp add: round_up_def round_down_def field_simps)
+ also have "\<dots> \<le> 1 * 2 powr -prec"
+ by (rule mult_mono)
+ (auto simp del: real_of_int_diff
+ simp: real_of_int_diff[symmetric] real_of_int_le_one_cancel_iff ceiling_diff_floor_le_1)
+ finally show ?thesis by simp
qed
-lemma float_le_simp: "((x::float) \<le> y) = (0 \<le> y - x)"
- by (auto simp add: le_float_def)
+lemma round_down_shift: "round_down p (x * 2 powr k) = 2 powr k * round_down (p + k) x"
+ unfolding round_down_def
+ by (simp add: powr_add powr_mult field_simps powr_divide2[symmetric])
+ (simp add: powr_add[symmetric])
-lemma float_less_simp: "((x::float) < y) = (0 < y - x)"
- by (auto simp add: less_float_def)
+lemma round_up_shift: "round_up p (x * 2 powr k) = 2 powr k * round_up (p + k) x"
+ unfolding round_up_def
+ by (simp add: powr_add powr_mult field_simps powr_divide2[symmetric])
+ (simp add: powr_add[symmetric])
+
+subsection {* Rounding Floats *}
-lemma real_of_float_min: "real (min x y :: float) = min (real x) (real y)" unfolding min_def le_float_def by auto
-lemma real_of_float_max: "real (max a b :: float) = max (real a) (real b)" unfolding max_def le_float_def by auto
+definition float_up :: "int \<Rightarrow> float \<Rightarrow> float" where
+ "float_up prec x = float_of (round_up prec (real x))"
+
+lemma float_up_float:
+ "x \<in> float \<Longrightarrow> float_up prec (float_of x) = float_of (round_up prec x)"
+ unfolding float_up_def by simp
-lemma float_power: "real (x ^ n :: float) = real x ^ n"
- by (induct n) simp_all
+lemma float_up_correct:
+ shows "real (float_up e f) - real f \<in> {0..2 powr -e}"
+unfolding atLeastAtMost_iff
+proof
+ have "round_up e f - f \<le> round_up e f - round_down e f" using round_down by simp
+ also have "\<dots> \<le> 2 powr -e" using round_up_diff_round_down by simp
+ finally show "real (float_up e f) - real f \<le> 2 powr real (- e)"
+ by (simp add: float_up_def)
+qed (simp add: algebra_simps float_up_def round_up)
-lemma zero_le_pow2[simp]: "0 \<le> pow2 s"
- apply (subgoal_tac "0 < pow2 s")
- apply (auto simp only:)
- apply auto
- done
+definition float_down :: "int \<Rightarrow> float \<Rightarrow> float" where
+ "float_down prec x = float_of (round_down prec (real x))"
+
+lemma float_down_float:
+ "x \<in> float \<Longrightarrow> float_down prec (float_of x) = float_of (round_down prec x)"
+ unfolding float_down_def by simp
-lemma pow2_less_0_eq_False[simp]: "(pow2 s < 0) = False"
- apply auto
- apply (subgoal_tac "0 \<le> pow2 s")
- apply simp
- apply simp
- done
+lemma float_down_correct:
+ shows "real f - real (float_down e f) \<in> {0..2 powr -e}"
+unfolding atLeastAtMost_iff
+proof
+ have "f - round_down e f \<le> round_up e f - round_down e f" using round_up by simp
+ also have "\<dots> \<le> 2 powr -e" using round_up_diff_round_down by simp
+ finally show "real f - real (float_down e f) \<le> 2 powr real (- e)"
+ by (simp add: float_down_def)
+qed (simp add: algebra_simps float_down_def round_down)
+
+lemma round_down_Float_id:
+ assumes "p + e \<ge> 0"
+ shows "round_down p (Float m e) = Float m e"
+proof -
+ from assms have r: "real e + real p = real (nat (e + p))" by simp
+ have r: "\<lfloor>real (Float m e) * 2 powr real p\<rfloor> = real (Float m e) * 2 powr real p"
+ by (auto intro: exI[where x="m*2^nat (e+p)"]
+ simp add: ac_simps powr_add[symmetric] r powr_realpow)
+ show ?thesis using assms
+ unfolding round_down_def floor_divide_eq_div r
+ by (simp add: ac_simps powr_add[symmetric])
+qed
-lemma pow2_le_0_eq_False[simp]: "(pow2 s \<le> 0) = False"
- apply auto
- apply (subgoal_tac "0 < pow2 s")
- apply simp
- apply simp
- done
+lemma compute_float_down[code]:
+ "float_down p (Float m e) =
+ (if p + e < 0 then Float (m div 2^nat (-(p + e))) (-p) else Float m e)"
+proof cases
+ assume "p + e < 0"
+ hence "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))"
+ using powr_realpow[of 2 "nat (-(p + e))"] by simp
+ also have "... = 1 / 2 powr p / 2 powr e"
+ unfolding powr_minus_divide real_of_int_minus by (simp add: powr_add)
+ finally show ?thesis
+ unfolding float_down_def round_down_def floor_divide_eq_div[symmetric]
+ using `p + e < 0` by (simp add: ac_simps)
+next
+ assume "\<not> p + e < 0" with round_down_Float_id show ?thesis by (simp add: float_down_def)
+qed
-lemma float_pos_m_pos: "0 < Float m e \<Longrightarrow> 0 < m"
- unfolding less_float_def real_of_float_simp real_of_float_0 zero_less_mult_iff
- by auto
+lemma ceil_divide_floor_conv:
+assumes "b \<noteq> 0"
+shows "\<lceil>real a / real b\<rceil> = (if b dvd a then a div b else \<lfloor>real a / real b\<rfloor> + 1)"
+proof cases
+ assume "\<not> b dvd a"
+ hence "a mod b \<noteq> 0" by auto
+ hence ne: "real (a mod b) / real b \<noteq> 0" using `b \<noteq> 0` by auto
+ have "\<lceil>real a / real b\<rceil> = \<lfloor>real a / real b\<rfloor> + 1"
+ apply (rule ceiling_eq) apply (auto simp: floor_divide_eq_div[symmetric])
+ proof -
+ have "real \<lfloor>real a / real b\<rfloor> \<le> real a / real b" by simp
+ moreover have "real \<lfloor>real a / real b\<rfloor> \<noteq> real a / real b"
+ apply (subst (2) real_of_int_div_aux) unfolding floor_divide_eq_div using ne `b \<noteq> 0` by auto
+ ultimately show "real \<lfloor>real a / real b\<rfloor> < real a / real b" by arith
+ qed
+ thus ?thesis using `\<not> b dvd a` by simp
+qed (simp add: ceiling_def real_of_int_minus[symmetric] divide_minus_left[symmetric]
+ floor_divide_eq_div dvd_neg_div del: divide_minus_left real_of_int_minus)
-lemma float_pos_less1_e_neg: assumes "0 < Float m e" and "Float m e < 1" shows "e < 0"
+lemma round_up_Float_id:
+ assumes "p + e \<ge> 0"
+ shows "round_up p (Float m e) = Float m e"
proof -
- have "0 < m" using float_pos_m_pos `0 < Float m e` by auto
- hence "0 \<le> real m" and "1 \<le> real m" by auto
-
- show "e < 0"
- proof (rule ccontr)
- assume "\<not> e < 0" hence "0 \<le> e" by auto
- hence "1 \<le> pow2 e" unfolding pow2_def by auto
- from mult_mono[OF `1 \<le> real m` this `0 \<le> real m`]
- have "1 \<le> Float m e" by (simp add: le_float_def real_of_float_simp)
- thus False using `Float m e < 1` unfolding less_float_def le_float_def by auto
+ from assms have r1: "real e + real p = real (nat (e + p))" by simp
+ have r: "\<lceil>real (Float m e) * 2 powr real p\<rceil> = real (Float m e) * 2 powr real p"
+ by (auto simp add: ac_simps powr_add[symmetric] r1 powr_realpow
+ intro: exI[where x="m*2^nat (e+p)"])
+ show ?thesis using assms
+ unfolding float_up_def round_up_def floor_divide_eq_div Let_def r
+ by (simp add: ac_simps powr_add[symmetric])
+qed
+
+lemma compute_float_up[code]:
+ "float_up p (Float m e) =
+ (let P = 2^nat (-(p + e)); r = m mod P in
+ if p + e < 0 then Float (m div P + (if r = 0 then 0 else 1)) (-p) else Float m e)"
+proof cases
+ assume "p + e < 0"
+ hence "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))"
+ using powr_realpow[of 2 "nat (-(p + e))"] by simp
+ also have "... = 1 / 2 powr p / 2 powr e"
+ unfolding powr_minus_divide real_of_int_minus by (simp add: powr_add)
+ finally have twopow_rewrite:
+ "real ((2::int) ^ nat (- (p + e))) = 1 / 2 powr real p / 2 powr real e" .
+ with `p + e < 0` have powr_rewrite:
+ "2 powr real e * 2 powr real p = 1 / real ((2::int) ^ nat (- (p + e)))"
+ unfolding powr_divide2 by simp
+ show ?thesis
+ proof cases
+ assume "2^nat (-(p + e)) dvd m"
+ with `p + e < 0` twopow_rewrite show ?thesis
+ by (auto simp: ac_simps float_up_def round_up_def floor_divide_eq_div dvd_eq_mod_eq_0)
+ next
+ assume ndvd: "\<not> 2 ^ nat (- (p + e)) dvd m"
+ have one_div: "real m * (1 / real ((2::int) ^ nat (- (p + e)))) =
+ real m / real ((2::int) ^ nat (- (p + e)))"
+ by (simp add: field_simps)
+ have "real \<lceil>real m * (2 powr real e * 2 powr real p)\<rceil> =
+ real \<lfloor>real m * (2 powr real e * 2 powr real p)\<rfloor> + 1"
+ using ndvd unfolding powr_rewrite one_div
+ by (subst ceil_divide_floor_conv) (auto simp: field_simps)
+ thus ?thesis using `p + e < 0` twopow_rewrite
+ by (auto simp: ac_simps Let_def float_up_def round_up_def floor_divide_eq_div[symmetric])
qed
+next
+ assume "\<not> p + e < 0" with round_up_Float_id show ?thesis by (simp add: float_up_def)
qed
-lemma float_less1_mantissa_bound: assumes "0 < Float m e" "Float m e < 1" shows "m < 2^(nat (-e))"
+lemmas real_of_ints =
+ real_of_int_zero
+ real_of_one
+ real_of_int_add
+ real_of_int_minus
+ real_of_int_diff
+ real_of_int_mult
+ real_of_int_power
+ real_numeral
+lemmas real_of_nats =
+ real_of_nat_zero
+ real_of_nat_one
+ real_of_nat_1
+ real_of_nat_add
+ real_of_nat_mult
+ real_of_nat_power
+
+lemmas int_of_reals = real_of_ints[symmetric]
+lemmas nat_of_reals = real_of_nats[symmetric]
+
+lemma two_real_int: "(2::real) = real (2::int)" by simp
+lemma two_real_nat: "(2::real) = real (2::nat)" by simp
+
+lemma mult_cong: "a = c ==> b = d ==> a*b = c*d" by simp
+
+subsection {* Compute bitlen of integers *}
+
+definition bitlen::"int => int"
+where "bitlen a = (if a > 0 then \<lfloor>log 2 a\<rfloor> + 1 else 0)"
+
+lemma bitlen_nonneg: "0 \<le> bitlen x"
proof -
- have "e < 0" using float_pos_less1_e_neg assms by auto
- have "\<And>x. (0::real) < 2^x" by auto
- have "real m < 2^(nat (-e))" using `Float m e < 1`
- unfolding less_float_def real_of_float_neg_exp[OF `e < 0`] real_of_float_1
- real_mult_less_iff1[of _ _ 1, OF `0 < 2^(nat (-e))`, symmetric]
- mult_assoc by auto
- thus ?thesis unfolding real_of_int_less_iff[symmetric] by auto
+ {
+ assume "0 > x"
+ have "-1 = log 2 (inverse 2)" by (subst log_inverse) simp_all
+ also have "... < log 2 (-x)" using `0 > x` by auto
+ finally have "-1 < log 2 (-x)" .
+ } thus "0 \<le> bitlen x" unfolding bitlen_def by (auto intro!: add_nonneg_nonneg)
+qed
+
+lemma bitlen_bounds:
+ assumes "x > 0"
+ shows "2 ^ nat (bitlen x - 1) \<le> x \<and> x < 2 ^ nat (bitlen x)"
+proof
+ have "(2::real) ^ nat \<lfloor>log 2 (real x)\<rfloor> = 2 powr real (floor (log 2 (real x)))"
+ using powr_realpow[symmetric, of 2 "nat \<lfloor>log 2 (real x)\<rfloor>"] `x > 0`
+ using real_nat_eq_real[of "floor (log 2 (real x))"]
+ by simp
+ also have "... \<le> 2 powr log 2 (real x)"
+ by simp
+ also have "... = real x"
+ using `0 < x` by simp
+ finally have "2 ^ nat \<lfloor>log 2 (real x)\<rfloor> \<le> real x" by simp
+ thus "2 ^ nat (bitlen x - 1) \<le> x" using `x > 0`
+ by (simp add: bitlen_def)
+next
+ have "x \<le> 2 powr (log 2 x)" using `x > 0` by simp
+ also have "... < 2 ^ nat (\<lfloor>log 2 (real x)\<rfloor> + 1)"
+ apply (simp add: powr_realpow[symmetric])
+ using `x > 0` by simp
+ finally show "x < 2 ^ nat (bitlen x)" using `x > 0`
+ by (simp add: bitlen_def ac_simps int_of_reals del: real_of_ints)
+qed
+
+lemma bitlen_pow2[simp]:
+ assumes "b > 0"
+ shows "bitlen (b * 2 ^ c) = bitlen b + c"
+proof -
+ from assms have "b * 2 ^ c > 0" by (auto intro: mult_pos_pos)
+ thus ?thesis
+ using floor_add[of "log 2 b" c] assms
+ by (auto simp add: log_mult log_nat_power bitlen_def)
qed
-function bitlen :: "int \<Rightarrow> int" where
-"bitlen 0 = 0" |
-"bitlen -1 = 1" |
-"0 < x \<Longrightarrow> bitlen x = 1 + (bitlen (x div 2))" |
-"x < -1 \<Longrightarrow> bitlen x = 1 + (bitlen (x div 2))"
- apply (case_tac "x = 0 \<or> x = -1 \<or> x < -1 \<or> x > 0")
- apply auto
- done
-termination by (relation "measure (nat o abs)", auto)
-
-lemma bitlen_ge0: "0 \<le> bitlen x" by (induct x rule: bitlen.induct, auto)
-lemma bitlen_ge1: "x \<noteq> 0 \<Longrightarrow> 1 \<le> bitlen x" by (induct x rule: bitlen.induct, auto simp add: bitlen_ge0)
+lemma bitlen_Float:
+fixes m e
+defines "f \<equiv> Float m e"
+shows "bitlen (\<bar>mantissa f\<bar>) + exponent f = (if m = 0 then 0 else bitlen \<bar>m\<bar> + e)"
+proof cases
+ assume "m \<noteq> 0" hence "f \<noteq> float_of 0" by (simp add: f_def) hence "mantissa f \<noteq> 0"
+ by (simp add: mantissa_noteq_0)
+ moreover
+ from f_def[THEN denormalize_shift, OF `f \<noteq> float_of 0`] guess i .
+ ultimately show ?thesis by (simp add: abs_mult)
+qed (simp add: f_def bitlen_def)
-lemma bitlen_bounds': assumes "0 < x" shows "2^nat (bitlen x - 1) \<le> x \<and> x + 1 \<le> 2^nat (bitlen x)" (is "?P x")
- using `0 < x`
-proof (induct x rule: bitlen.induct)
- fix x
- assume "0 < x" and hyp: "0 < x div 2 \<Longrightarrow> ?P (x div 2)" hence "0 \<le> x" and "x \<noteq> 0" by auto
- { fix x have "0 \<le> 1 + bitlen x" using bitlen_ge0[of x] by auto } note gt0_pls1 = this
+lemma compute_bitlen[code]:
+ shows "bitlen x = (if x > 0 then bitlen (x div 2) + 1 else 0)"
+proof -
+ { assume "2 \<le> x"
+ then have "\<lfloor>log 2 (x div 2)\<rfloor> + 1 = \<lfloor>log 2 (x - x mod 2)\<rfloor>"
+ by (simp add: log_mult zmod_zdiv_equality')
+ also have "\<dots> = \<lfloor>log 2 (real x)\<rfloor>"
+ proof cases
+ assume "x mod 2 = 0" then show ?thesis by simp
+ next
+ def n \<equiv> "\<lfloor>log 2 (real x)\<rfloor>"
+ then have "0 \<le> n"
+ using `2 \<le> x` by simp
+ assume "x mod 2 \<noteq> 0"
+ with `2 \<le> x` have "x mod 2 = 1" "\<not> 2 dvd x" by (auto simp add: dvd_eq_mod_eq_0)
+ with `2 \<le> x` have "x \<noteq> 2^nat n" by (cases "nat n") auto
+ moreover
+ { have "real (2^nat n :: int) = 2 powr (nat n)"
+ by (simp add: powr_realpow)
+ also have "\<dots> \<le> 2 powr (log 2 x)"
+ using `2 \<le> x` by (simp add: n_def del: powr_log_cancel)
+ finally have "2^nat n \<le> x" using `2 \<le> x` by simp }
+ ultimately have "2^nat n \<le> x - 1" by simp
+ then have "2^nat n \<le> real (x - 1)"
+ unfolding real_of_int_le_iff[symmetric] by simp
+ { have "n = \<lfloor>log 2 (2^nat n)\<rfloor>"
+ using `0 \<le> n` by (simp add: log_nat_power)
+ also have "\<dots> \<le> \<lfloor>log 2 (x - 1)\<rfloor>"
+ using `2^nat n \<le> real (x - 1)` `0 \<le> n` `2 \<le> x` by (auto intro: floor_mono)
+ finally have "n \<le> \<lfloor>log 2 (x - 1)\<rfloor>" . }
+ moreover have "\<lfloor>log 2 (x - 1)\<rfloor> \<le> n"
+ using `2 \<le> x` by (auto simp add: n_def intro!: floor_mono)
+ ultimately show "\<lfloor>log 2 (x - x mod 2)\<rfloor> = \<lfloor>log 2 x\<rfloor>"
+ unfolding n_def `x mod 2 = 1` by auto
+ qed
+ finally have "\<lfloor>log 2 (x div 2)\<rfloor> + 1 = \<lfloor>log 2 x\<rfloor>" . }
+ moreover
+ { assume "x < 2" "0 < x"
+ then have "x = 1" by simp
+ then have "\<lfloor>log 2 (real x)\<rfloor> = 0" by simp }
+ ultimately show ?thesis
+ unfolding bitlen_def
+ by (auto simp: pos_imp_zdiv_pos_iff not_le)
+qed
- have "0 < (2::int)" by auto
-
- show "?P x"
- proof (cases "x = 1")
- case True show "?P x" unfolding True by auto
+lemma float_gt1_scale: assumes "1 \<le> Float m e"
+ shows "0 \<le> e + (bitlen m - 1)"
+proof -
+ have "0 < Float m e" using assms by auto
+ hence "0 < m" using powr_gt_zero[of 2 e]
+ by (auto simp: less_float_def less_eq_float_def zero_less_mult_iff)
+ hence "m \<noteq> 0" by auto
+ show ?thesis
+ proof (cases "0 \<le> e")
+ case True thus ?thesis using `0 < m` by (simp add: bitlen_def)
next
- case False hence "2 \<le> x" using `0 < x` `x \<noteq> 1` by auto
- hence "2 div 2 \<le> x div 2" by (rule zdiv_mono1, auto)
- hence "0 < x div 2" and "x div 2 \<noteq> 0" by auto
- hence bitlen_s1_ge0: "0 \<le> bitlen (x div 2) - 1" using bitlen_ge1[OF `x div 2 \<noteq> 0`] by auto
-
- { from hyp[OF `0 < x div 2`]
- have "2 ^ nat (bitlen (x div 2) - 1) \<le> x div 2" by auto
- hence "2 ^ nat (bitlen (x div 2) - 1) * 2 \<le> x div 2 * 2" by (rule mult_right_mono, auto)
- also have "\<dots> \<le> x" using `0 < x` by auto
- finally have "2^nat (1 + bitlen (x div 2) - 1) \<le> x" unfolding power_Suc2[symmetric] Suc_nat_eq_nat_zadd1[OF bitlen_s1_ge0] by auto
- } moreover
- { have "x + 1 \<le> x - x mod 2 + 2"
- proof -
- have "x mod 2 < 2" using `0 < x` by auto
- hence "x < x - x mod 2 + 2" unfolding algebra_simps by auto
- thus ?thesis by auto
- qed
- also have "x - x mod 2 + 2 = (x div 2 + 1) * 2" unfolding algebra_simps using `0 < x` div_mod_equality[of x 2 0] by auto
- also have "\<dots> \<le> 2^nat (bitlen (x div 2)) * 2" using hyp[OF `0 < x div 2`, THEN conjunct2] by (rule mult_right_mono, auto)
- also have "\<dots> = 2^(1 + nat (bitlen (x div 2)))" unfolding power_Suc2[symmetric] by auto
- finally have "x + 1 \<le> 2^(1 + nat (bitlen (x div 2)))" .
- }
- ultimately show ?thesis
- unfolding bitlen.simps(3)[OF `0 < x`] nat_add_distrib[OF zero_le_one bitlen_ge0]
- unfolding add_commute nat_add_distrib[OF zero_le_one gt0_pls1]
- by auto
+ have "(1::int) < 2" by simp
+ case False let ?S = "2^(nat (-e))"
+ have "inverse (2 ^ nat (- e)) = 2 powr e" using assms False powr_realpow[of 2 "nat (-e)"]
+ by (auto simp: powr_minus field_simps inverse_eq_divide)
+ hence "1 \<le> real m * inverse ?S" using assms False powr_realpow[of 2 "nat (-e)"]
+ by (auto simp: powr_minus)
+ hence "1 * ?S \<le> real m * inverse ?S * ?S" by (rule mult_right_mono, auto)
+ hence "?S \<le> real m" unfolding mult_assoc by auto
+ hence "?S \<le> m" unfolding real_of_int_le_iff[symmetric] by auto
+ from this bitlen_bounds[OF `0 < m`, THEN conjunct2]
+ have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF `1 < 2`, symmetric] by (rule order_le_less_trans)
+ hence "-e < bitlen m" using False by auto
+ thus ?thesis by auto
qed
-next
- fix x :: int assume "x < -1" and "0 < x" hence False by auto
- thus "?P x" by auto
-qed auto
-
-lemma bitlen_bounds: assumes "0 < x" shows "2^nat (bitlen x - 1) \<le> x \<and> x < 2^nat (bitlen x)"
- using bitlen_bounds'[OF `0<x`] by auto
+qed
lemma bitlen_div: assumes "0 < m" shows "1 \<le> real m / 2^nat (bitlen m - 1)" and "real m / 2^nat (bitlen m - 1) < 2"
proof -
@@ -514,840 +930,571 @@
thus "1 \<le> real m / ?B" by auto
have "m \<noteq> 0" using assms by auto
- have "0 \<le> bitlen m - 1" using bitlen_ge1[OF `m \<noteq> 0`] by auto
+ have "0 \<le> bitlen m - 1" using `0 < m` by (auto simp: bitlen_def)
have "m < 2^nat(bitlen m)" using bitlen_bounds[OF `0 <m`] ..
- also have "\<dots> = 2^nat(bitlen m - 1 + 1)" using bitlen_ge1[OF `m \<noteq> 0`] by auto
+ also have "\<dots> = 2^nat(bitlen m - 1 + 1)" using `0 < m` by (auto simp: bitlen_def)
also have "\<dots> = ?B * 2" unfolding nat_add_distrib[OF `0 \<le> bitlen m - 1` zero_le_one] by auto
finally have "real m < 2 * ?B" unfolding real_of_int_less_iff[symmetric] by auto
hence "real m / ?B < 2 * ?B / ?B" by (rule divide_strict_right_mono, auto)
thus "real m / ?B < 2" by auto
qed
-lemma float_gt1_scale: assumes "1 \<le> Float m e"
- shows "0 \<le> e + (bitlen m - 1)"
-proof (cases "0 \<le> e")
- have "0 < Float m e" using assms unfolding less_float_def le_float_def by auto
- hence "0 < m" using float_pos_m_pos by auto
- hence "m \<noteq> 0" by auto
- case True with bitlen_ge1[OF `m \<noteq> 0`] show ?thesis by auto
-next
- have "0 < Float m e" using assms unfolding less_float_def le_float_def by auto
- hence "0 < m" using float_pos_m_pos by auto
- hence "m \<noteq> 0" and "1 < (2::int)" by auto
- case False let ?S = "2^(nat (-e))"
- have "1 \<le> real m * inverse ?S" using assms unfolding le_float_def real_of_float_nge0_exp[OF False] by auto
- hence "1 * ?S \<le> real m * inverse ?S * ?S" by (rule mult_right_mono, auto)
- hence "?S \<le> real m" unfolding mult_assoc by auto
- hence "?S \<le> m" unfolding real_of_int_le_iff[symmetric] by auto
- from this bitlen_bounds[OF `0 < m`, THEN conjunct2]
- have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF `1 < 2`, symmetric] by (rule order_le_less_trans)
- hence "-e < bitlen m" using False bitlen_ge0 by auto
- thus ?thesis by auto
-qed
+subsection {* Approximation of positive rationals *}
+
+lemma zdiv_zmult_twopow_eq: fixes a b::int shows "a div b div (2 ^ n) = a div (b * 2 ^ n)"
+by (simp add: zdiv_zmult2_eq)
-lemma normalized_float: assumes "m \<noteq> 0" shows "real (Float m (- (bitlen m - 1))) = real m / 2^nat (bitlen m - 1)"
-proof (cases "- (bitlen m - 1) = 0")
- case True show ?thesis unfolding real_of_float_simp pow2_def using True by auto
-next
- case False hence P: "\<not> 0 \<le> - (bitlen m - 1)" using bitlen_ge1[OF `m \<noteq> 0`] by auto
- show ?thesis unfolding real_of_float_nge0_exp[OF P] divide_inverse by auto
-qed
+lemma div_mult_twopow_eq: fixes a b::nat shows "a div ((2::nat) ^ n) div b = a div (b * 2 ^ n)"
+ by (cases "b=0") (simp_all add: div_mult2_eq[symmetric] ac_simps)
-(* BROKEN
-lemma bitlen_Pls: "bitlen (Int.Pls) = Int.Pls" by (subst Pls_def, subst Pls_def, simp)
+lemma real_div_nat_eq_floor_of_divide:
+ fixes a b::nat
+ shows "a div b = real (floor (a/b))"
+by (metis floor_divide_eq_div real_of_int_of_nat_eq zdiv_int)
-lemma bitlen_Min: "bitlen (Int.Min) = Int.Bit1 Int.Pls" by (subst Min_def, simp add: Bit1_def)
+definition "rat_precision prec x y = int prec - (bitlen x - bitlen y)"
-lemma bitlen_B0: "bitlen (Int.Bit0 b) = (if iszero b then Int.Pls else Int.succ (bitlen b))"
- apply (auto simp add: iszero_def succ_def)
- apply (simp add: Bit0_def Pls_def)
- apply (subst Bit0_def)
- apply simp
- apply (subgoal_tac "0 < 2 * b \<or> 2 * b < -1")
- apply auto
- done
+definition lapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float" where
+ "lapprox_posrat prec x y = float_of (round_down (rat_precision prec x y) (x / y))"
-lemma bitlen_B1: "bitlen (Int.Bit1 b) = (if iszero (Int.succ b) then Int.Bit1 Int.Pls else Int.succ (bitlen b))"
-proof -
- have h: "! x. (2*x + 1) div 2 = (x::int)"
- by arith
- show ?thesis
- apply (auto simp add: iszero_def succ_def)
- apply (subst Bit1_def)+
- apply simp
- apply (subgoal_tac "2 * b + 1 = -1")
- apply (simp only:)
- apply simp_all
- apply (subst Bit1_def)
- apply simp
- apply (subgoal_tac "0 < 2 * b + 1 \<or> 2 * b + 1 < -1")
- apply (auto simp add: h)
- done
-qed
+lemma compute_lapprox_posrat[code]:
+ fixes prec x y
+ shows "lapprox_posrat prec x y =
+ (let
+ l = rat_precision prec x y;
+ d = if 0 \<le> l then x * 2^nat l div y else x div 2^nat (- l) div y
+ in normfloat (Float d (- l)))"
+ unfolding lapprox_posrat_def div_mult_twopow_eq
+ by (simp add: round_down_def powr_int real_div_nat_eq_floor_of_divide
+ field_simps Let_def
+ del: two_powr_minus_int_float)
-lemma bitlen_number_of: "bitlen (number_of w) = number_of (bitlen w)"
- by (simp add: number_of_is_id)
-BH *)
-
-lemma [code]: "bitlen x =
- (if x = 0 then 0
- else if x = -1 then 1
- else (1 + (bitlen (x div 2))))"
- by (cases "x = 0 \<or> x = -1 \<or> 0 < x") auto
-
-definition lapprox_posrat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float"
-where
- "lapprox_posrat prec x y =
- (let
- l = nat (int prec + bitlen y - bitlen x) ;
- d = (x * 2^l) div y
- in normfloat (Float d (- (int l))))"
-
-lemma pow2_minus: "pow2 (-x) = inverse (pow2 x)"
- unfolding pow2_neg[of "-x"] by auto
+definition rapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float" where
+ "rapprox_posrat prec x y = float_of (round_up (rat_precision prec x y) (x / y))"
-lemma lapprox_posrat:
- assumes x: "0 \<le> x"
- and y: "0 < y"
- shows "real (lapprox_posrat prec x y) \<le> real x / real y"
-proof -
- let ?l = "nat (int prec + bitlen y - bitlen x)"
-
- have "real (x * 2^?l div y) * inverse (2^?l) \<le> (real (x * 2^?l) / real y) * inverse (2^?l)"
- by (rule mult_right_mono, fact real_of_int_div4, simp)
- also have "\<dots> \<le> (real x / real y) * 2^?l * inverse (2^?l)" by auto
- finally have "real (x * 2^?l div y) * inverse (2^?l) \<le> real x / real y" unfolding mult_assoc by auto
- thus ?thesis unfolding lapprox_posrat_def Let_def normfloat real_of_float_simp
- unfolding pow2_minus pow2_int minus_minus .
-qed
-
-lemma real_of_int_div_mult:
- fixes x y c :: int assumes "0 < y" and "0 < c"
- shows "real (x div y) \<le> real (x * c div y) * inverse (real c)"
-proof -
- have "c * (x div y) + 0 \<le> c * x div y" unfolding zdiv_zmult1_eq[of c x y]
- by (rule add_left_mono,
- auto intro!: mult_nonneg_nonneg
- simp add: pos_imp_zdiv_nonneg_iff[OF `0 < y`] `0 < c`[THEN less_imp_le] pos_mod_sign[OF `0 < y`])
- hence "real (x div y) * real c \<le> real (x * c div y)"
- unfolding real_of_int_mult[symmetric] real_of_int_le_iff mult_commute by auto
- hence "real (x div y) * real c * inverse (real c) \<le> real (x * c div y) * inverse (real c)"
- using `0 < c` by auto
- thus ?thesis unfolding mult_assoc using `0 < c` by auto
-qed
-
-lemma lapprox_posrat_bottom: assumes "0 < y"
- shows "real (x div y) \<le> real (lapprox_posrat n x y)"
-proof -
- have pow: "\<And>x. (0::int) < 2^x" by auto
- show ?thesis
- unfolding lapprox_posrat_def Let_def real_of_float_add normfloat real_of_float_simp pow2_minus pow2_int
- using real_of_int_div_mult[OF `0 < y` pow] by auto
-qed
-
-lemma lapprox_posrat_nonneg: assumes "0 \<le> x" and "0 < y"
- shows "0 \<le> real (lapprox_posrat n x y)"
-proof -
+(* TODO: optimize using zmod_zmult2_eq, pdivmod ? *)
+lemma compute_rapprox_posrat[code]:
+ fixes prec x y
+ defines "l \<equiv> rat_precision prec x y"
+ shows "rapprox_posrat prec x y = (let
+ l = l ;
+ X = if 0 \<le> l then (x * 2^nat l, y) else (x, y * 2^nat(-l)) ;
+ d = fst X div snd X ;
+ m = fst X mod snd X
+ in normfloat (Float (d + (if m = 0 \<or> y = 0 then 0 else 1)) (- l)))"
+proof (cases "y = 0")
+ assume "y = 0" thus ?thesis by (simp add: rapprox_posrat_def Let_def)
+next
+ assume "y \<noteq> 0"
show ?thesis
- unfolding lapprox_posrat_def Let_def real_of_float_add normfloat real_of_float_simp pow2_minus pow2_int
- using pos_imp_zdiv_nonneg_iff[OF `0 < y`] assms by (auto intro!: mult_nonneg_nonneg)
-qed
-
-definition rapprox_posrat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float"
-where
- "rapprox_posrat prec x y = (let
- l = nat (int prec + bitlen y - bitlen x) ;
- X = x * 2^l ;
- d = X div y ;
- m = X mod y
- in normfloat (Float (d + (if m = 0 then 0 else 1)) (- (int l))))"
-
-lemma rapprox_posrat:
- assumes x: "0 \<le> x"
- and y: "0 < y"
- shows "real x / real y \<le> real (rapprox_posrat prec x y)"
-proof -
- let ?l = "nat (int prec + bitlen y - bitlen x)" let ?X = "x * 2^?l"
- show ?thesis
- proof (cases "?X mod y = 0")
- case True hence "y dvd ?X" using `0 < y` by auto
- from real_of_int_div[OF this]
- have "real (?X div y) * inverse (2 ^ ?l) = real ?X / real y * inverse (2 ^ ?l)" by auto
- also have "\<dots> = real x / real y * (2^?l * inverse (2^?l))" by auto
- finally have "real (?X div y) * inverse (2^?l) = real x / real y" by auto
- thus ?thesis unfolding rapprox_posrat_def Let_def normfloat if_P[OF True]
- unfolding real_of_float_simp pow2_minus pow2_int minus_minus by auto
- next
- case False
- have "0 \<le> real y" and "real y \<noteq> 0" using `0 < y` by auto
- have "0 \<le> real y * 2^?l" by (rule mult_nonneg_nonneg, rule `0 \<le> real y`, auto)
-
- have "?X = y * (?X div y) + ?X mod y" by auto
- also have "\<dots> \<le> y * (?X div y) + y" by (rule add_mono, auto simp add: pos_mod_bound[OF `0 < y`, THEN less_imp_le])
- also have "\<dots> = y * (?X div y + 1)" unfolding right_distrib by auto
- finally have "real ?X \<le> real y * real (?X div y + 1)" unfolding real_of_int_le_iff real_of_int_mult[symmetric] .
- hence "real ?X / (real y * 2^?l) \<le> real y * real (?X div y + 1) / (real y * 2^?l)"
- by (rule divide_right_mono, simp only: `0 \<le> real y * 2^?l`)
- also have "\<dots> = real y * real (?X div y + 1) / real y / 2^?l" by auto
- also have "\<dots> = real (?X div y + 1) * inverse (2^?l)" unfolding nonzero_mult_divide_cancel_left[OF `real y \<noteq> 0`]
- unfolding divide_inverse ..
- finally show ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_not_P[OF False]
- unfolding pow2_minus pow2_int minus_minus by auto
+ proof (cases "0 \<le> l")
+ assume "0 \<le> l"
+ def x' == "x * 2 ^ nat l"
+ have "int x * 2 ^ nat l = x'" by (simp add: x'_def int_mult int_power)
+ moreover have "real x * 2 powr real l = real x'"
+ by (simp add: powr_realpow[symmetric] `0 \<le> l` x'_def)
+ ultimately show ?thesis
+ unfolding rapprox_posrat_def round_up_def l_def[symmetric]
+ using ceil_divide_floor_conv[of y x'] powr_realpow[of 2 "nat l"] `0 \<le> l` `y \<noteq> 0`
+ by (simp add: Let_def floor_divide_eq_div[symmetric] dvd_eq_mod_eq_0 int_of_reals
+ del: real_of_ints)
+ next
+ assume "\<not> 0 \<le> l"
+ def y' == "y * 2 ^ nat (- l)"
+ from `y \<noteq> 0` have "y' \<noteq> 0" by (simp add: y'_def)
+ have "int y * 2 ^ nat (- l) = y'" by (simp add: y'_def int_mult int_power)
+ moreover have "real x * real (2::int) powr real l / real y = x / real y'"
+ using `\<not> 0 \<le> l`
+ by (simp add: powr_realpow[symmetric] powr_minus y'_def field_simps inverse_eq_divide)
+ ultimately show ?thesis
+ using ceil_divide_floor_conv[of y' x] `\<not> 0 \<le> l` `y' \<noteq> 0` `y \<noteq> 0`
+ by (simp add: rapprox_posrat_def l_def round_up_def ceil_divide_floor_conv
+ floor_divide_eq_div[symmetric] dvd_eq_mod_eq_0 int_of_reals
+ del: real_of_ints)
qed
qed
-lemma rapprox_posrat_le1: assumes "0 \<le> x" and "0 < y" and "x \<le> y"
- shows "real (rapprox_posrat n x y) \<le> 1"
+
+lemma rat_precision_pos:
+ assumes "0 \<le> x" and "0 < y" and "2 * x < y" and "0 < n"
+ shows "rat_precision n (int x) (int y) > 0"
proof -
- let ?l = "nat (int n + bitlen y - bitlen x)" let ?X = "x * 2^?l"
- show ?thesis
- proof (cases "?X mod y = 0")
- case True hence "y dvd ?X" using `0 < y` by auto
- from real_of_int_div[OF this]
- have "real (?X div y) * inverse (2 ^ ?l) = real ?X / real y * inverse (2 ^ ?l)" by auto
- also have "\<dots> = real x / real y * (2^?l * inverse (2^?l))" by auto
- finally have "real (?X div y) * inverse (2^?l) = real x / real y" by auto
- also have "real x / real y \<le> 1" using `0 \<le> x` and `0 < y` and `x \<le> y` by auto
- finally show ?thesis unfolding rapprox_posrat_def Let_def normfloat if_P[OF True]
- unfolding real_of_float_simp pow2_minus pow2_int minus_minus by auto
- next
- case False
- have "x \<noteq> y"
- proof (rule ccontr)
- assume "\<not> x \<noteq> y" hence "x = y" by auto
- have "?X mod y = 0" unfolding `x = y` using mod_mult_self1_is_0 by auto
- thus False using False by auto
- qed
- hence "x < y" using `x \<le> y` by auto
- hence "real x / real y < 1" using `0 < y` and `0 \<le> x` by auto
-
- from real_of_int_div4[of "?X" y]
- have "real (?X div y) \<le> (real x / real y) * 2^?l" unfolding real_of_int_mult times_divide_eq_left real_of_int_power real_numeral .
- also have "\<dots> < 1 * 2^?l" using `real x / real y < 1` by (rule mult_strict_right_mono, auto)
- finally have "?X div y < 2^?l" unfolding real_of_int_less_iff[of _ "2^?l", symmetric] by auto
- hence "?X div y + 1 \<le> 2^?l" by auto
- hence "real (?X div y + 1) * inverse (2^?l) \<le> 2^?l * inverse (2^?l)"
- unfolding real_of_int_le_iff[of _ "2^?l", symmetric] real_of_int_power real_numeral
- by (rule mult_right_mono, auto)
- hence "real (?X div y + 1) * inverse (2^?l) \<le> 1" by auto
- thus ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_not_P[OF False]
- unfolding pow2_minus pow2_int minus_minus by auto
- qed
+ { assume "0 < x" hence "log 2 x + 1 = log 2 (2 * x)" by (simp add: log_mult) }
+ hence "bitlen (int x) < bitlen (int y)" using assms
+ by (simp add: bitlen_def del: floor_add_one)
+ (auto intro!: floor_mono simp add: floor_add_one[symmetric] simp del: floor_add floor_add_one)
+ thus ?thesis
+ using assms by (auto intro!: pos_add_strict simp add: field_simps rat_precision_def)
qed
-lemma zdiv_greater_zero: fixes a b :: int assumes "0 < a" and "a \<le> b"
- shows "0 < b div a"
-proof (rule ccontr)
- have "0 \<le> b" using assms by auto
- assume "\<not> 0 < b div a" hence "b div a = 0" using `0 \<le> b`[unfolded pos_imp_zdiv_nonneg_iff[OF `0<a`, of b, symmetric]] by auto
- have "b = a * (b div a) + b mod a" by auto
- hence "b = b mod a" unfolding `b div a = 0` by auto
- hence "b < a" using `0 < a`[THEN pos_mod_bound, of b] by auto
- thus False using `a \<le> b` by auto
+lemma power_aux: assumes "x > 0" shows "(2::int) ^ nat (x - 1) \<le> 2 ^ nat x - 1"
+proof -
+ def y \<equiv> "nat (x - 1)" moreover
+ have "(2::int) ^ y \<le> (2 ^ (y + 1)) - 1" by simp
+ ultimately show ?thesis using assms by simp
qed
lemma rapprox_posrat_less1: assumes "0 \<le> x" and "0 < y" and "2 * x < y" and "0 < n"
shows "real (rapprox_posrat n x y) < 1"
-proof (cases "x = 0")
- case True thus ?thesis unfolding rapprox_posrat_def True Let_def normfloat real_of_float_simp by auto
-next
- case False hence "0 < x" using `0 \<le> x` by auto
- hence "x < y" using assms by auto
-
- let ?l = "nat (int n + bitlen y - bitlen x)" let ?X = "x * 2^?l"
- show ?thesis
- proof (cases "?X mod y = 0")
- case True hence "y dvd ?X" using `0 < y` by auto
- from real_of_int_div[OF this]
- have "real (?X div y) * inverse (2 ^ ?l) = real ?X / real y * inverse (2 ^ ?l)" by auto
- also have "\<dots> = real x / real y * (2^?l * inverse (2^?l))" by auto
- finally have "real (?X div y) * inverse (2^?l) = real x / real y" by auto
- also have "real x / real y < 1" using `0 \<le> x` and `0 < y` and `x < y` by auto
- finally show ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_P[OF True]
- unfolding pow2_minus pow2_int minus_minus by auto
- next
- case False
- hence "(real x / real y) < 1 / 2" using `0 < y` and `0 \<le> x` `2 * x < y` by auto
-
- have "0 < ?X div y"
- proof -
- have "2^nat (bitlen x - 1) \<le> y" and "y < 2^nat (bitlen y)"
- using bitlen_bounds[OF `0 < x`, THEN conjunct1] bitlen_bounds[OF `0 < y`, THEN conjunct2] `x < y` by auto
- hence "(2::int)^nat (bitlen x - 1) < 2^nat (bitlen y)" by (rule order_le_less_trans)
- hence "bitlen x \<le> bitlen y" by auto
- hence len_less: "nat (bitlen x - 1) \<le> nat (int (n - 1) + bitlen y)" by auto
-
- have "x \<noteq> 0" and "y \<noteq> 0" using `0 < x` `0 < y` by auto
-
- have exp_eq: "nat (int (n - 1) + bitlen y) - nat (bitlen x - 1) = ?l"
- using `bitlen x \<le> bitlen y` bitlen_ge1[OF `x \<noteq> 0`] bitlen_ge1[OF `y \<noteq> 0`] `0 < n` by auto
-
- have "y * 2^nat (bitlen x - 1) \<le> y * x"
- using bitlen_bounds[OF `0 < x`, THEN conjunct1] `0 < y`[THEN less_imp_le] by (rule mult_left_mono)
- also have "\<dots> \<le> 2^nat (bitlen y) * x" using bitlen_bounds[OF `0 < y`, THEN conjunct2, THEN less_imp_le] `0 \<le> x` by (rule mult_right_mono)
- also have "\<dots> \<le> x * 2^nat (int (n - 1) + bitlen y)" unfolding mult_commute[of x] by (rule mult_right_mono, auto simp add: `0 \<le> x`)
- finally have "real y * 2^nat (bitlen x - 1) * inverse (2^nat (bitlen x - 1)) \<le> real x * 2^nat (int (n - 1) + bitlen y) * inverse (2^nat (bitlen x - 1))"
- unfolding real_of_int_le_iff[symmetric] by auto
- hence "real y \<le> real x * (2^nat (int (n - 1) + bitlen y) / (2^nat (bitlen x - 1)))"
- unfolding mult_assoc divide_inverse by auto
- also have "\<dots> = real x * (2^(nat (int (n - 1) + bitlen y) - nat (bitlen x - 1)))" using power_diff[of "2::real", OF _ len_less] by auto
- finally have "y \<le> x * 2^?l" unfolding exp_eq unfolding real_of_int_le_iff[symmetric] by auto
- thus ?thesis using zdiv_greater_zero[OF `0 < y`] by auto
- qed
-
- from real_of_int_div4[of "?X" y]
- have "real (?X div y) \<le> (real x / real y) * 2^?l" unfolding real_of_int_mult times_divide_eq_left real_of_int_power real_numeral .
- also have "\<dots> < 1/2 * 2^?l" using `real x / real y < 1/2` by (rule mult_strict_right_mono, auto)
- finally have "?X div y * 2 < 2^?l" unfolding real_of_int_less_iff[of _ "2^?l", symmetric] by auto
- hence "?X div y + 1 < 2^?l" using `0 < ?X div y` by auto
- hence "real (?X div y + 1) * inverse (2^?l) < 2^?l * inverse (2^?l)"
- unfolding real_of_int_less_iff[of _ "2^?l", symmetric] real_of_int_power real_numeral
- by (rule mult_strict_right_mono, auto)
- hence "real (?X div y + 1) * inverse (2^?l) < 1" by auto
- thus ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_not_P[OF False]
- unfolding pow2_minus pow2_int minus_minus by auto
- qed
+proof -
+ have powr1: "2 powr real (rat_precision n (int x) (int y)) =
+ 2 ^ nat (rat_precision n (int x) (int y))" using rat_precision_pos[of x y n] assms
+ by (simp add: powr_realpow[symmetric])
+ have "x * 2 powr real (rat_precision n (int x) (int y)) / y = (x / y) *
+ 2 powr real (rat_precision n (int x) (int y))" by simp
+ also have "... < (1 / 2) * 2 powr real (rat_precision n (int x) (int y))"
+ apply (rule mult_strict_right_mono) by (insert assms) auto
+ also have "\<dots> = 2 powr real (rat_precision n (int x) (int y) - 1)"
+ by (simp add: powr_add diff_def powr_neg_numeral)
+ also have "\<dots> = 2 ^ nat (rat_precision n (int x) (int y) - 1)"
+ using rat_precision_pos[of x y n] assms by (simp add: powr_realpow[symmetric])
+ also have "\<dots> \<le> 2 ^ nat (rat_precision n (int x) (int y)) - 1"
+ unfolding int_of_reals real_of_int_le_iff
+ using rat_precision_pos[OF assms] by (rule power_aux)
+ finally show ?thesis unfolding rapprox_posrat_def
+ apply (simp add: round_up_def)
+ apply (simp add: round_up_def field_simps powr_minus inverse_eq_divide)
+ unfolding powr1
+ unfolding int_of_reals real_of_int_less_iff
+ unfolding ceiling_less_eq using rat_precision_pos[of x y n] assms apply simp done
qed
-lemma approx_rat_pattern: fixes P and ps :: "nat * int * int"
- assumes Y: "\<And>y prec x. \<lbrakk>y = 0; ps = (prec, x, 0)\<rbrakk> \<Longrightarrow> P"
- and A: "\<And>x y prec. \<lbrakk>0 \<le> x; 0 < y; ps = (prec, x, y)\<rbrakk> \<Longrightarrow> P"
- and B: "\<And>x y prec. \<lbrakk>x < 0; 0 < y; ps = (prec, x, y)\<rbrakk> \<Longrightarrow> P"
- and C: "\<And>x y prec. \<lbrakk>x < 0; y < 0; ps = (prec, x, y)\<rbrakk> \<Longrightarrow> P"
- and D: "\<And>x y prec. \<lbrakk>0 \<le> x; y < 0; ps = (prec, x, y)\<rbrakk> \<Longrightarrow> P"
- shows P
-proof -
- obtain prec x y where [simp]: "ps = (prec, x, y)" by (cases ps) auto
- from Y have "y = 0 \<Longrightarrow> P" by auto
- moreover {
- assume "0 < y"
- have P
- proof (cases "0 \<le> x")
- case True
- with A and `0 < y` show P by auto
- next
- case False
- with B and `0 < y` show P by auto
- qed
- }
- moreover {
- assume "y < 0"
- have P
- proof (cases "0 \<le> x")
- case True
- with D and `y < 0` show P by auto
- next
- case False
- with C and `y < 0` show P by auto
- qed
- }
- ultimately show P by (cases "y = 0 \<or> 0 < y \<or> y < 0") auto
-qed
-
-function lapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float"
-where
- "y = 0 \<Longrightarrow> lapprox_rat prec x y = 0"
-| "0 \<le> x \<Longrightarrow> 0 < y \<Longrightarrow> lapprox_rat prec x y = lapprox_posrat prec x y"
-| "x < 0 \<Longrightarrow> 0 < y \<Longrightarrow> lapprox_rat prec x y = - (rapprox_posrat prec (-x) y)"
-| "x < 0 \<Longrightarrow> y < 0 \<Longrightarrow> lapprox_rat prec x y = lapprox_posrat prec (-x) (-y)"
-| "0 \<le> x \<Longrightarrow> y < 0 \<Longrightarrow> lapprox_rat prec x y = - (rapprox_posrat prec x (-y))"
-apply simp_all by (rule approx_rat_pattern)
-termination by lexicographic_order
+definition lapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" where
+ "lapprox_rat prec x y = float_of (round_down (rat_precision prec \<bar>x\<bar> \<bar>y\<bar>) (x / y))"
lemma compute_lapprox_rat[code]:
- "lapprox_rat prec x y = (if y = 0 then 0 else if 0 \<le> x then (if 0 < y then lapprox_posrat prec x y else - (rapprox_posrat prec x (-y)))
- else (if 0 < y then - (rapprox_posrat prec (-x) y) else lapprox_posrat prec (-x) (-y)))"
- by auto
-
-lemma lapprox_rat: "real (lapprox_rat prec x y) \<le> real x / real y"
-proof -
- have h[rule_format]: "! a b b'. b' \<le> b \<longrightarrow> a \<le> b' \<longrightarrow> a \<le> (b::real)" by auto
- show ?thesis
- apply (case_tac "y = 0")
- apply simp
- apply (case_tac "0 \<le> x \<and> 0 < y")
- apply (simp add: lapprox_posrat)
- apply (case_tac "x < 0 \<and> 0 < y")
- apply simp
- apply (subst minus_le_iff)
- apply (rule h[OF rapprox_posrat])
- apply (simp_all)
- apply (case_tac "x < 0 \<and> y < 0")
- apply simp
- apply (rule h[OF _ lapprox_posrat])
- apply (simp_all)
- apply (case_tac "0 \<le> x \<and> y < 0")
- apply (simp)
- apply (subst minus_le_iff)
- apply (rule h[OF rapprox_posrat])
- apply simp_all
- apply arith
- done
+ "lapprox_rat prec x y =
+ (if y = 0 then 0
+ else if 0 \<le> x then
+ (if 0 < y then lapprox_posrat prec (nat x) (nat y)
+ else - (rapprox_posrat prec (nat x) (nat (-y))))
+ else (if 0 < y
+ then - (rapprox_posrat prec (nat (-x)) (nat y))
+ else lapprox_posrat prec (nat (-x)) (nat (-y))))"
+ apply (cases "y = 0")
+ apply (simp add: lapprox_posrat_def rapprox_posrat_def round_down_def lapprox_rat_def)
+ apply (auto simp: lapprox_rat_def lapprox_posrat_def rapprox_posrat_def round_up_def round_down_def
+ ceiling_def real_of_float_uminus[symmetric] ac_simps int_of_reals simp del: real_of_ints)
+ apply (auto simp: ac_simps)
+ done
+
+definition rapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" where
+ "rapprox_rat prec x y = float_of (round_up (rat_precision prec \<bar>x\<bar> \<bar>y\<bar>) (x / y))"
+
+lemma compute_rapprox_rat[code]:
+ "rapprox_rat prec x y =
+ (if y = 0 then 0
+ else if 0 \<le> x then
+ (if 0 < y then rapprox_posrat prec (nat x) (nat y)
+ else - (lapprox_posrat prec (nat x) (nat (-y))))
+ else (if 0 < y
+ then - (lapprox_posrat prec (nat (-x)) (nat y))
+ else rapprox_posrat prec (nat (-x)) (nat (-y))))"
+ apply (cases "y = 0", simp add: lapprox_posrat_def rapprox_posrat_def round_up_def rapprox_rat_def)
+ apply (auto simp: rapprox_rat_def lapprox_posrat_def rapprox_posrat_def round_up_def round_down_def
+ ceiling_def real_of_float_uminus[symmetric] ac_simps int_of_reals simp del: real_of_ints)
+ apply (auto simp: ac_simps)
+ done
+
+subsection {* Division *}
+
+definition div_precision
+where "div_precision prec x y =
+ rat_precision prec \<bar>mantissa x\<bar> \<bar>mantissa y\<bar> - exponent x + exponent y"
+
+definition float_divl :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float"
+where "float_divl prec a b =
+ float_of (round_down (div_precision prec a b) (a / b))"
+
+lemma compute_float_divl[code]:
+ fixes m1 s1 m2 s2
+ defines "f1 \<equiv> Float m1 s1" and "f2 \<equiv> Float m2 s2"
+ shows "float_divl prec f1 f2 = lapprox_rat prec m1 m2 * Float 1 (s1 - s2)"
+proof cases
+ assume "f1 \<noteq> 0 \<and> f2 \<noteq> 0"
+ then have "f1 \<noteq> float_of 0" "f2 \<noteq> float_of 0" by auto
+ with mantissa_not_dvd[of f1] mantissa_not_dvd[of f2]
+ have "mantissa f1 \<noteq> 0" "mantissa f2 \<noteq> 0"
+ by (auto simp add: dvd_def)
+ then have pos: "0 < \<bar>mantissa f1\<bar>" "0 < \<bar>mantissa f2\<bar>"
+ by simp_all
+ moreover from f1_def[THEN denormalize_shift, OF `f1 \<noteq> float_of 0`] guess i . note i = this
+ moreover from f2_def[THEN denormalize_shift, OF `f2 \<noteq> float_of 0`] guess j . note j = this
+ moreover have "(real (mantissa f1) * 2 ^ i / (real (mantissa f2) * 2 ^ j))
+ = (real (mantissa f1) / real (mantissa f2)) * 2 powr (int i - int j)"
+ by (simp add: powr_divide2[symmetric] powr_realpow)
+ moreover have "real f1 / real f2 = real (mantissa f1) / real (mantissa f2) * 2 powr real (exponent f1 - exponent f2)"
+ unfolding mantissa_exponent by (simp add: powr_divide2[symmetric])
+ moreover have "rat_precision prec (\<bar>mantissa f1\<bar> * 2 ^ i) (\<bar>mantissa f2\<bar> * 2 ^ j) =
+ rat_precision prec \<bar>mantissa f1\<bar> \<bar>mantissa f2\<bar> + j - i"
+ using pos by (simp add: rat_precision_def)
+ ultimately show ?thesis
+ unfolding float_divl_def lapprox_rat_def div_precision_def
+ by (simp add: abs_mult round_down_shift powr_divide2[symmetric]
+ del: int_nat_eq real_of_int_diff times_divide_eq_left )
+ (simp add: field_simps powr_divide2[symmetric] powr_add)
+next
+ assume "\<not> (f1 \<noteq> 0 \<and> f2 \<noteq> 0)" then show ?thesis
+ by (auto simp add: float_divl_def f1_def f2_def lapprox_rat_def)
+qed
+
+definition float_divr :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float"
+where "float_divr prec a b =
+ float_of (round_up (div_precision prec a b) (a / b))"
+
+lemma compute_float_divr[code]:
+ fixes m1 s1 m2 s2
+ defines "f1 \<equiv> Float m1 s1" and "f2 \<equiv> Float m2 s2"
+ shows "float_divr prec f1 f2 = rapprox_rat prec m1 m2 * Float 1 (s1 - s2)"
+proof cases
+ assume "f1 \<noteq> 0 \<and> f2 \<noteq> 0"
+ then have "f1 \<noteq> float_of 0" "f2 \<noteq> float_of 0" by auto
+ with mantissa_not_dvd[of f1] mantissa_not_dvd[of f2]
+ have "mantissa f1 \<noteq> 0" "mantissa f2 \<noteq> 0"
+ by (auto simp add: dvd_def)
+ then have pos: "0 < \<bar>mantissa f1\<bar>" "0 < \<bar>mantissa f2\<bar>"
+ by simp_all
+ moreover from f1_def[THEN denormalize_shift, OF `f1 \<noteq> float_of 0`] guess i . note i = this
+ moreover from f2_def[THEN denormalize_shift, OF `f2 \<noteq> float_of 0`] guess j . note j = this
+ moreover have "(real (mantissa f1) * 2 ^ i / (real (mantissa f2) * 2 ^ j))
+ = (real (mantissa f1) / real (mantissa f2)) * 2 powr (int i - int j)"
+ by (simp add: powr_divide2[symmetric] powr_realpow)
+ moreover have "real f1 / real f2 = real (mantissa f1) / real (mantissa f2) * 2 powr real (exponent f1 - exponent f2)"
+ unfolding mantissa_exponent by (simp add: powr_divide2[symmetric])
+ moreover have "rat_precision prec (\<bar>mantissa f1\<bar> * 2 ^ i) (\<bar>mantissa f2\<bar> * 2 ^ j) =
+ rat_precision prec \<bar>mantissa f1\<bar> \<bar>mantissa f2\<bar> + j - i"
+ using pos by (simp add: rat_precision_def)
+ ultimately show ?thesis
+ unfolding float_divr_def rapprox_rat_def div_precision_def
+ by (simp add: abs_mult round_up_shift powr_divide2[symmetric]
+ del: int_nat_eq real_of_int_diff times_divide_eq_left)
+ (simp add: field_simps powr_divide2[symmetric] powr_add)
+next
+ assume "\<not> (f1 \<noteq> 0 \<and> f2 \<noteq> 0)" then show ?thesis
+ by (auto simp add: float_divr_def f1_def f2_def rapprox_rat_def)
qed
-lemma lapprox_rat_bottom: assumes "0 \<le> x" and "0 < y"
- shows "real (x div y) \<le> real (lapprox_rat n x y)"
- unfolding lapprox_rat.simps(2)[OF assms] using lapprox_posrat_bottom[OF `0<y`] .
+subsection {* Lemmas needed by Approximate *}
+
+declare one_float_def[simp del] zero_float_def[simp del]
+
+lemma Float_num[simp]: shows
+ "real (Float 1 0) = 1" and "real (Float 1 1) = 2" and "real (Float 1 2) = 4" and
+ "real (Float 1 -1) = 1/2" and "real (Float 1 -2) = 1/4" and "real (Float 1 -3) = 1/8" and
+ "real (Float -1 0) = -1" and "real (Float (number_of n) 0) = number_of n"
+using two_powr_int_float[of 2] two_powr_int_float[of "-1"] two_powr_int_float[of "-2"] two_powr_int_float[of "-3"]
+using powr_realpow[of 2 2] powr_realpow[of 2 3]
+using powr_minus[of 2 1] powr_minus[of 2 2] powr_minus[of 2 3]
+by auto
+
+lemma real_of_Float_int[simp]: "real (Float n 0) = real n" by simp
+
+lemma float_zero[simp]: "real (Float 0 e) = 0" by simp
+
+lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> abs((a::int) div 2) < abs a"
+by arith
-function rapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float"
-where
- "y = 0 \<Longrightarrow> rapprox_rat prec x y = 0"
-| "0 \<le> x \<Longrightarrow> 0 < y \<Longrightarrow> rapprox_rat prec x y = rapprox_posrat prec x y"
-| "x < 0 \<Longrightarrow> 0 < y \<Longrightarrow> rapprox_rat prec x y = - (lapprox_posrat prec (-x) y)"
-| "x < 0 \<Longrightarrow> y < 0 \<Longrightarrow> rapprox_rat prec x y = rapprox_posrat prec (-x) (-y)"
-| "0 \<le> x \<Longrightarrow> y < 0 \<Longrightarrow> rapprox_rat prec x y = - (lapprox_posrat prec x (-y))"
-apply simp_all by (rule approx_rat_pattern)
-termination by lexicographic_order
+lemma lapprox_rat:
+ shows "real (lapprox_rat prec x y) \<le> real x / real y"
+ using round_down by (simp add: lapprox_rat_def)
-lemma compute_rapprox_rat[code]:
- "rapprox_rat prec x y = (if y = 0 then 0 else if 0 \<le> x then (if 0 < y then rapprox_posrat prec x y else - (lapprox_posrat prec x (-y))) else
- (if 0 < y then - (lapprox_posrat prec (-x) y) else rapprox_posrat prec (-x) (-y)))"
- by auto
+lemma mult_div_le: fixes a b:: int assumes "b > 0" shows "a \<ge> b * (a div b)"
+proof -
+ from zmod_zdiv_equality'[of a b]
+ have "a = b * (a div b) + a mod b" by simp
+ also have "... \<ge> b * (a div b) + 0" apply (rule add_left_mono) apply (rule pos_mod_sign)
+ using assms by simp
+ finally show ?thesis by simp
+qed
+
+lemma lapprox_rat_nonneg:
+ fixes n x y
+ defines "p == int n - ((bitlen \<bar>x\<bar>) - (bitlen \<bar>y\<bar>))"
+ assumes "0 \<le> x" "0 < y"
+ shows "0 \<le> real (lapprox_rat n x y)"
+using assms unfolding lapprox_rat_def p_def[symmetric] round_down_def real_of_int_minus[symmetric]
+ powr_int[of 2, simplified]
+ by (auto simp add: inverse_eq_divide intro!: mult_nonneg_nonneg divide_nonneg_pos mult_pos_pos)
lemma rapprox_rat: "real x / real y \<le> real (rapprox_rat prec x y)"
-proof -
- have h[rule_format]: "! a b b'. b' \<le> b \<longrightarrow> a \<le> b' \<longrightarrow> a \<le> (b::real)" by auto
- show ?thesis
- apply (case_tac "y = 0")
- apply simp
- apply (case_tac "0 \<le> x \<and> 0 < y")
- apply (simp add: rapprox_posrat)
- apply (case_tac "x < 0 \<and> 0 < y")
- apply simp
- apply (subst le_minus_iff)
- apply (rule h[OF _ lapprox_posrat])
- apply (simp_all)
- apply (case_tac "x < 0 \<and> y < 0")
- apply simp
- apply (rule h[OF rapprox_posrat])
- apply (simp_all)
- apply (case_tac "0 \<le> x \<and> y < 0")
- apply (simp)
- apply (subst le_minus_iff)
- apply (rule h[OF _ lapprox_posrat])
- apply simp_all
- apply arith
- done
+ using round_up by (simp add: rapprox_rat_def)
+
+lemma rapprox_rat_le1:
+ fixes n x y
+ assumes xy: "0 \<le> x" "0 < y" "x \<le> y"
+ shows "real (rapprox_rat n x y) \<le> 1"
+proof -
+ have "bitlen \<bar>x\<bar> \<le> bitlen \<bar>y\<bar>"
+ using xy unfolding bitlen_def by (auto intro!: floor_mono)
+ then have "0 \<le> rat_precision n \<bar>x\<bar> \<bar>y\<bar>" by (simp add: rat_precision_def)
+ have "real \<lceil>real x / real y * 2 powr real (rat_precision n \<bar>x\<bar> \<bar>y\<bar>)\<rceil>
+ \<le> real \<lceil>2 powr real (rat_precision n \<bar>x\<bar> \<bar>y\<bar>)\<rceil>"
+ using xy by (auto intro!: ceiling_mono simp: field_simps)
+ also have "\<dots> = 2 powr real (rat_precision n \<bar>x\<bar> \<bar>y\<bar>)"
+ using `0 \<le> rat_precision n \<bar>x\<bar> \<bar>y\<bar>`
+ by (auto intro!: exI[of _ "2^nat (rat_precision n \<bar>x\<bar> \<bar>y\<bar>)"] simp: powr_int)
+ finally show ?thesis
+ by (simp add: rapprox_rat_def round_up_def)
+ (simp add: powr_minus inverse_eq_divide)
qed
-lemma rapprox_rat_le1: assumes "0 \<le> x" and "0 < y" and "x \<le> y"
- shows "real (rapprox_rat n x y) \<le> 1"
- unfolding rapprox_rat.simps(2)[OF `0 \<le> x` `0 < y`] using rapprox_posrat_le1[OF assms] .
-
-lemma rapprox_rat_neg: assumes "x < 0" and "0 < y"
- shows "real (rapprox_rat n x y) \<le> 0"
- unfolding rapprox_rat.simps(3)[OF assms] using lapprox_posrat_nonneg[of "-x" y n] assms by auto
-
-lemma rapprox_rat_nonneg_neg: assumes "0 \<le> x" and "y < 0"
- shows "real (rapprox_rat n x y) \<le> 0"
- unfolding rapprox_rat.simps(5)[OF assms] using lapprox_posrat_nonneg[of x "-y" n] assms by auto
+lemma rapprox_rat_nonneg_neg:
+ "0 \<le> x \<Longrightarrow> y < 0 \<Longrightarrow> real (rapprox_rat n x y) \<le> 0"
+ unfolding rapprox_rat_def round_up_def
+ by (auto simp: field_simps mult_le_0_iff zero_le_mult_iff)
-lemma rapprox_rat_nonpos_pos: assumes "x \<le> 0" and "0 < y"
- shows "real (rapprox_rat n x y) \<le> 0"
-proof (cases "x = 0")
- case True
- hence "0 \<le> x" by auto show ?thesis
- unfolding rapprox_rat.simps(2)[OF `0 \<le> x` `0 < y`]
- unfolding True rapprox_posrat_def Let_def
- by auto
-next
- case False
- hence "x < 0" using assms by auto
- show ?thesis using rapprox_rat_neg[OF `x < 0` `0 < y`] .
-qed
+lemma rapprox_rat_neg:
+ "x < 0 \<Longrightarrow> 0 < y \<Longrightarrow> real (rapprox_rat n x y) \<le> 0"
+ unfolding rapprox_rat_def round_up_def
+ by (auto simp: field_simps mult_le_0_iff)
-fun float_divl :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float"
-where
- "float_divl prec (Float m1 s1) (Float m2 s2) =
- (let
- l = lapprox_rat prec m1 m2;
- f = Float 1 (s1 - s2)
- in
- f * l)"
+lemma rapprox_rat_nonpos_pos:
+ "x \<le> 0 \<Longrightarrow> 0 < y \<Longrightarrow> real (rapprox_rat n x y) \<le> 0"
+ unfolding rapprox_rat_def round_up_def
+ by (auto simp: field_simps mult_le_0_iff)
lemma float_divl: "real (float_divl prec x y) \<le> real x / real y"
- using lapprox_rat[of prec "mantissa x" "mantissa y"]
- by (cases x y rule: float.exhaust[case_product float.exhaust])
- (simp split: split_if_asm
- add: real_of_float_simp pow2_diff field_simps le_divide_eq mult_less_0_iff zero_less_mult_iff)
+ using round_down by (simp add: float_divl_def)
+
+lemma float_divl_lower_bound:
+ fixes x y prec
+ defines "p == rat_precision prec \<bar>mantissa x\<bar> \<bar>mantissa y\<bar> - exponent x + exponent y"
+ assumes xy: "0 \<le> x" "0 < y" shows "0 \<le> real (float_divl prec x y)"
+ using xy unfolding float_divl_def p_def[symmetric] round_down_def
+ by (simp add: zero_le_mult_iff zero_le_divide_iff less_eq_float_def less_float_def)
+
+lemma exponent_1: "exponent 1 = 0"
+ using exponent_float[of 1 0] by (simp add: one_float_def)
+
+lemma mantissa_1: "mantissa 1 = 1"
+ using mantissa_float[of 1 0] by (simp add: one_float_def)
-lemma float_divl_lower_bound: assumes "0 \<le> x" and "0 < y" shows "0 \<le> float_divl prec x y"
-proof (cases x, cases y)
- fix xm xe ym ye :: int
- assume x_eq: "x = Float xm xe" and y_eq: "y = Float ym ye"
- have "0 \<le> xm"
- using `0 \<le> x`[unfolded x_eq le_float_def real_of_float_simp real_of_float_0 zero_le_mult_iff]
- by auto
- have "0 < ym"
- using `0 < y`[unfolded y_eq less_float_def real_of_float_simp real_of_float_0 zero_less_mult_iff]
- by auto
+lemma bitlen_1: "bitlen 1 = 1"
+ by (simp add: bitlen_def)
+
+lemma mantissa_eq_zero_iff: "mantissa x = 0 \<longleftrightarrow> x = 0"
+proof
+ assume "mantissa x = 0" hence z: "0 = real x" using mantissa_exponent by simp
+ show "x = 0" by (simp add: zero_float_def z)
+qed (simp add: zero_float_def)
- have "\<And>n. 0 \<le> real (Float 1 n)"
- unfolding real_of_float_simp using zero_le_pow2 by auto
- moreover have "0 \<le> real (lapprox_rat prec xm ym)"
- apply (rule order_trans[OF _ lapprox_rat_bottom[OF `0 \<le> xm` `0 < ym`]])
- apply (auto simp add: `0 \<le> xm` pos_imp_zdiv_nonneg_iff[OF `0 < ym`])
- done
- ultimately show "0 \<le> float_divl prec x y"
- unfolding x_eq y_eq float_divl.simps Let_def le_float_def real_of_float_0
- by (auto intro!: mult_nonneg_nonneg)
+lemma float_upper_bound: "x \<le> 2 powr (bitlen \<bar>mantissa x\<bar> + exponent x)"
+proof (cases "x = 0", simp)
+ assume "x \<noteq> 0" hence "mantissa x \<noteq> 0" using mantissa_eq_zero_iff by auto
+ have "x = mantissa x * 2 powr (exponent x)" by (rule mantissa_exponent)
+ also have "mantissa x \<le> \<bar>mantissa x\<bar>" by simp
+ also have "... \<le> 2 powr (bitlen \<bar>mantissa x\<bar>)"
+ using bitlen_bounds[of "\<bar>mantissa x\<bar>"] bitlen_nonneg `mantissa x \<noteq> 0`
+ by (simp add: powr_int) (simp only: two_real_int int_of_reals real_of_int_abs[symmetric]
+ real_of_int_le_iff less_imp_le)
+ finally show ?thesis by (simp add: powr_add)
qed
lemma float_divl_pos_less1_bound:
- assumes "0 < x" and "x < 1" and "0 < prec"
- shows "1 \<le> float_divl prec 1 x"
-proof (cases x)
- case (Float m e)
- from `0 < x` `x < 1` have "0 < m" "e < 0"
- using float_pos_m_pos float_pos_less1_e_neg unfolding Float by auto
- let ?b = "nat (bitlen m)" and ?e = "nat (-e)"
- have "1 \<le> m" and "m \<noteq> 0" using `0 < m` by auto
- with bitlen_bounds[OF `0 < m`] have "m < 2^?b" and "(2::int) \<le> 2^?b" by auto
- hence "1 \<le> bitlen m" using power_le_imp_le_exp[of "2::int" 1 ?b] by auto
- hence pow_split: "nat (int prec + bitlen m - 1) = (prec - 1) + ?b" using `0 < prec` by auto
-
- have pow_not0: "\<And>x. (2::real)^x \<noteq> 0" by auto
-
- from float_less1_mantissa_bound `0 < x` `x < 1` Float
- have "m < 2^?e" by auto
- with bitlen_bounds[OF `0 < m`, THEN conjunct1] have "(2::int)^nat (bitlen m - 1) < 2^?e"
- by (rule order_le_less_trans)
- from power_less_imp_less_exp[OF _ this]
- have "bitlen m \<le> - e" by auto
- hence "(2::real)^?b \<le> 2^?e" by auto
- hence "(2::real)^?b * inverse (2^?b) \<le> 2^?e * inverse (2^?b)"
- by (rule mult_right_mono) auto
- hence "(1::real) \<le> 2^?e * inverse (2^?b)" by auto
- also
- let ?d = "real (2 ^ nat (int prec + bitlen m - 1) div m) * inverse (2 ^ nat (int prec + bitlen m - 1))"
- {
- have "2^(prec - 1) * m \<le> 2^(prec - 1) * 2^?b"
- using `m < 2^?b`[THEN less_imp_le] by (rule mult_left_mono) auto
- also have "\<dots> = 2 ^ nat (int prec + bitlen m - 1)"
- unfolding pow_split power_add by auto
- finally have "2^(prec - 1) * m div m \<le> 2 ^ nat (int prec + bitlen m - 1) div m"
- using `0 < m` by (rule zdiv_mono1)
- hence "2^(prec - 1) \<le> 2 ^ nat (int prec + bitlen m - 1) div m"
- unfolding div_mult_self2_is_id[OF `m \<noteq> 0`] .
- hence "2^(prec - 1) * inverse (2 ^ nat (int prec + bitlen m - 1)) \<le> ?d"
- unfolding real_of_int_le_iff[of "2^(prec - 1)", symmetric] by auto
- }
- from mult_left_mono[OF this [unfolded pow_split power_add inverse_mult_distrib mult_assoc[symmetric] right_inverse[OF pow_not0] mult_1_left], of "2^?e"]
- have "2^?e * inverse (2^?b) \<le> 2^?e * ?d" unfolding pow_split power_add by auto
- finally have "1 \<le> 2^?e * ?d" .
-
- have e_nat: "0 - e = int (nat (-e))" using `e < 0` by auto
- have "bitlen 1 = 1" using bitlen.simps by auto
-
- show ?thesis
- unfolding one_float_def Float float_divl.simps Let_def
- lapprox_rat.simps(2)[OF zero_le_one `0 < m`]
- lapprox_posrat_def `bitlen 1 = 1`
- unfolding le_float_def real_of_float_mult normfloat real_of_float_simp
- pow2_minus pow2_int e_nat
- using `1 \<le> 2^?e * ?d` by (auto simp add: pow2_def)
+ assumes "0 < real x" and "real x < 1" and "prec \<ge> 1"
+ shows "1 \<le> real (float_divl prec 1 x)"
+proof cases
+ assume nonneg: "div_precision prec 1 x \<ge> 0"
+ hence "2 powr real (div_precision prec 1 x) =
+ floor (real ((2::int) ^ nat (div_precision prec 1 x))) * floor (1::real)"
+ by (simp add: powr_int del: real_of_int_power) simp
+ also have "floor (1::real) \<le> floor (1 / x)" using assms by simp
+ also have "floor (real ((2::int) ^ nat (div_precision prec 1 x))) * floor (1 / x) \<le>
+ floor (real ((2::int) ^ nat (div_precision prec 1 x)) * (1 / x))"
+ by (rule le_mult_floor) (auto simp: assms less_imp_le)
+ finally have "2 powr real (div_precision prec 1 x) <=
+ floor (2 powr nat (div_precision prec 1 x) / x)" by (simp add: powr_realpow)
+ thus ?thesis
+ using assms nonneg
+ unfolding float_divl_def round_down_def
+ by simp (simp add: powr_minus inverse_eq_divide)
+next
+ assume neg: "\<not> 0 \<le> div_precision prec 1 x"
+ have "1 \<le> 1 * 2 powr (prec - 1)" using assms by (simp add: powr_realpow)
+ also have "... \<le> 2 powr (bitlen \<bar>mantissa x\<bar> + exponent x) / x * 2 powr (prec - 1)"
+ apply (rule mult_mono) using assms float_upper_bound
+ by (auto intro!: divide_nonneg_pos)
+ also have "2 powr (bitlen \<bar>mantissa x\<bar> + exponent x) / x * 2 powr (prec - 1) =
+ 2 powr real (div_precision prec 1 x) / real x"
+ using assms
+ apply (simp add: div_precision_def rat_precision_def diff_diff_eq2
+ mantissa_1 exponent_1 bitlen_1 powr_add powr_minus real_of_nat_diff)
+ apply (simp only: diff_def powr_add)
+ apply simp
+ done
+ finally have "1 \<le> \<lfloor>2 powr real (div_precision prec 1 x) / real x\<rfloor>"
+ using floor_mono[of "1::real"] by simp thm mult_mono
+ hence "1 \<le> real \<lfloor>2 powr real (div_precision prec 1 x) / real x\<rfloor>"
+ by (metis floor_real_of_int one_le_floor)
+ hence "1 * 1 \<le>
+ real \<lfloor>2 powr real (div_precision prec 1 x) / real x\<rfloor> * 2 powr - real (div_precision prec 1 x)"
+ apply (rule mult_mono)
+ using assms neg
+ by (auto intro: divide_nonneg_pos mult_nonneg_nonneg simp: real_of_int_minus[symmetric] powr_int simp del: real_of_int_minus) find_theorems "real (- _)"
+ thus ?thesis
+ using assms neg
+ unfolding float_divl_def round_down_def
+ by simp
qed
-fun float_divr :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float"
-where
- "float_divr prec (Float m1 s1) (Float m2 s2) =
- (let
- r = rapprox_rat prec m1 m2;
- f = Float 1 (s1 - s2)
- in
- f * r)"
-
lemma float_divr: "real x / real y \<le> real (float_divr prec x y)"
- using rapprox_rat[of "mantissa x" "mantissa y" prec]
- by (cases x y rule: float.exhaust[case_product float.exhaust])
- (simp split: split_if_asm
- add: real_of_float_simp pow2_diff field_simps divide_le_eq mult_less_0_iff zero_less_mult_iff)
+ using round_up by (simp add: float_divr_def)
lemma float_divr_pos_less1_lower_bound: assumes "0 < x" and "x < 1" shows "1 \<le> float_divr prec 1 x"
proof -
have "1 \<le> 1 / real x" using `0 < x` and `x < 1` unfolding less_float_def by auto
also have "\<dots> \<le> real (float_divr prec 1 x)" using float_divr[where x=1 and y=x] by auto
- finally show ?thesis unfolding le_float_def by auto
-qed
-
-lemma float_divr_nonpos_pos_upper_bound: assumes "x \<le> 0" and "0 < y" shows "float_divr prec x y \<le> 0"
-proof (cases x, cases y)
- fix xm xe ym ye :: int
- assume x_eq: "x = Float xm xe" and y_eq: "y = Float ym ye"
- have "xm \<le> 0" using `x \<le> 0`[unfolded x_eq le_float_def real_of_float_simp real_of_float_0 mult_le_0_iff] by auto
- have "0 < ym" using `0 < y`[unfolded y_eq less_float_def real_of_float_simp real_of_float_0 zero_less_mult_iff] by auto
-
- have "\<And>n. 0 \<le> real (Float 1 n)" unfolding real_of_float_simp using zero_le_pow2 by auto
- moreover have "real (rapprox_rat prec xm ym) \<le> 0" using rapprox_rat_nonpos_pos[OF `xm \<le> 0` `0 < ym`] .
- ultimately show "float_divr prec x y \<le> 0"
- unfolding x_eq y_eq float_divr.simps Let_def le_float_def real_of_float_0 real_of_float_mult by (auto intro!: mult_nonneg_nonpos)
-qed
-
-lemma float_divr_nonneg_neg_upper_bound: assumes "0 \<le> x" and "y < 0" shows "float_divr prec x y \<le> 0"
-proof (cases x, cases y)
- fix xm xe ym ye :: int
- assume x_eq: "x = Float xm xe" and y_eq: "y = Float ym ye"
- have "0 \<le> xm" using `0 \<le> x`[unfolded x_eq le_float_def real_of_float_simp real_of_float_0 zero_le_mult_iff] by auto
- have "ym < 0" using `y < 0`[unfolded y_eq less_float_def real_of_float_simp real_of_float_0 mult_less_0_iff] by auto
- hence "0 < - ym" by auto
-
- have "\<And>n. 0 \<le> real (Float 1 n)" unfolding real_of_float_simp using zero_le_pow2 by auto
- moreover have "real (rapprox_rat prec xm ym) \<le> 0" using rapprox_rat_nonneg_neg[OF `0 \<le> xm` `ym < 0`] .
- ultimately show "float_divr prec x y \<le> 0"
- unfolding x_eq y_eq float_divr.simps Let_def le_float_def real_of_float_0 real_of_float_mult by (auto intro!: mult_nonneg_nonpos)
-qed
-
-primrec round_down :: "nat \<Rightarrow> float \<Rightarrow> float" where
-"round_down prec (Float m e) = (let d = bitlen m - int prec in
- if 0 < d then let P = 2^nat d ; n = m div P in Float n (e + d)
- else Float m e)"
-
-primrec round_up :: "nat \<Rightarrow> float \<Rightarrow> float" where
-"round_up prec (Float m e) = (let d = bitlen m - int prec in
- if 0 < d then let P = 2^nat d ; n = m div P ; r = m mod P in Float (n + (if r = 0 then 0 else 1)) (e + d)
- else Float m e)"
-
-lemma round_up: "real x \<le> real (round_up prec x)"
-proof (cases x)
- case (Float m e)
- let ?d = "bitlen m - int prec"
- let ?p = "(2::int)^nat ?d"
- have "0 < ?p" by auto
- show "?thesis"
- proof (cases "0 < ?d")
- case True
- hence pow_d: "pow2 ?d = real ?p" using pow2_int[symmetric] by simp
- show ?thesis
- proof (cases "m mod ?p = 0")
- case True
- have m: "m = m div ?p * ?p + 0" unfolding True[symmetric] using mod_div_equality [symmetric] .
- have "real (Float m e) = real (Float (m div ?p) (e + ?d))" unfolding real_of_float_simp arg_cong[OF m, of real]
- by (auto simp add: pow2_add `0 < ?d` pow_d)
- thus ?thesis
- unfolding Float round_up.simps Let_def if_P[OF `m mod ?p = 0`] if_P[OF `0 < ?d`]
- by auto
- next
- case False
- have "m = m div ?p * ?p + m mod ?p" unfolding mod_div_equality ..
- also have "\<dots> \<le> (m div ?p + 1) * ?p" unfolding left_distrib mult_1 by (rule add_left_mono, rule pos_mod_bound[OF `0 < ?p`, THEN less_imp_le])
- finally have "real (Float m e) \<le> real (Float (m div ?p + 1) (e + ?d))" unfolding real_of_float_simp add_commute[of e]
- unfolding pow2_add mult_assoc[symmetric] real_of_int_le_iff[of m, symmetric]
- by (auto intro!: mult_mono simp add: pow2_add `0 < ?d` pow_d)
- thus ?thesis
- unfolding Float round_up.simps Let_def if_not_P[OF `\<not> m mod ?p = 0`] if_P[OF `0 < ?d`] .
- qed
- next
- case False
- show ?thesis
- unfolding Float round_up.simps Let_def if_not_P[OF False] ..
- qed
-qed
-
-lemma round_down: "real (round_down prec x) \<le> real x"
-proof (cases x)
- case (Float m e)
- let ?d = "bitlen m - int prec"
- let ?p = "(2::int)^nat ?d"
- have "0 < ?p" by auto
- show "?thesis"
- proof (cases "0 < ?d")
- case True
- hence pow_d: "pow2 ?d = real ?p" using pow2_int[symmetric] by simp
- have "m div ?p * ?p \<le> m div ?p * ?p + m mod ?p" by (auto simp add: pos_mod_bound[OF `0 < ?p`, THEN less_imp_le])
- also have "\<dots> \<le> m" unfolding mod_div_equality ..
- finally have "real (Float (m div ?p) (e + ?d)) \<le> real (Float m e)" unfolding real_of_float_simp add_commute[of e]
- unfolding pow2_add mult_assoc[symmetric] real_of_int_le_iff[of _ m, symmetric]
- by (auto intro!: mult_mono simp add: pow2_add `0 < ?d` pow_d)
- thus ?thesis
- unfolding Float round_down.simps Let_def if_P[OF `0 < ?d`] .
- next
- case False
- show ?thesis
- unfolding Float round_down.simps Let_def if_not_P[OF False] ..
- qed
+ finally show ?thesis unfolding less_eq_float_def by auto
qed
-definition lb_mult :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where
- "lb_mult prec x y =
- (case normfloat (x * y) of Float m e \<Rightarrow>
- let
- l = bitlen m - int prec
- in if l > 0 then Float (m div (2^nat l)) (e + l)
- else Float m e)"
+lemma float_divr_nonpos_pos_upper_bound:
+ assumes "real x \<le> 0" and "0 < real y"
+ shows "real (float_divr prec x y) \<le> 0"
+using assms
+unfolding float_divr_def round_up_def
+by (auto simp: field_simps mult_le_0_iff divide_le_0_iff)
-definition ub_mult :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where
- "ub_mult prec x y =
- (case normfloat (x * y) of Float m e \<Rightarrow>
- let
- l = bitlen m - int prec
- in if l > 0 then Float (m div (2^nat l) + 1) (e + l)
- else Float m e)"
+lemma float_divr_nonneg_neg_upper_bound:
+ assumes "0 \<le> real x" and "real y < 0"
+ shows "real (float_divr prec x y) \<le> 0"
+using assms
+unfolding float_divr_def round_up_def
+by (auto simp: field_simps mult_le_0_iff zero_le_mult_iff divide_le_0_iff)
+
+definition "round_prec p f = int p - (bitlen \<bar>mantissa f\<bar> + exponent f)"
+
+definition float_round_down :: "nat \<Rightarrow> float \<Rightarrow> float" where
+"float_round_down prec f = float_of (round_down (round_prec prec f) f)"
+
+definition float_round_up :: "nat \<Rightarrow> float \<Rightarrow> float" where
+"float_round_up prec f = float_of (round_up (round_prec prec f) f)"
-lemma lb_mult: "real (lb_mult prec x y) \<le> real (x * y)"
-proof (cases "normfloat (x * y)")
- case (Float m e)
- hence "odd m \<or> (m = 0 \<and> e = 0)" by (rule normfloat_imp_odd_or_zero)
- let ?l = "bitlen m - int prec"
- have "real (lb_mult prec x y) \<le> real (normfloat (x * y))"
- proof (cases "?l > 0")
- case False thus ?thesis unfolding lb_mult_def Float Let_def float.cases by auto
- next
- case True
- have "real (m div 2^(nat ?l)) * pow2 ?l \<le> real m"
- proof -
- have "real (m div 2^(nat ?l)) * pow2 ?l = real (2^(nat ?l) * (m div 2^(nat ?l)))" unfolding real_of_int_mult real_of_int_power real_numeral unfolding pow2_int[symmetric]
- using `?l > 0` by auto
- also have "\<dots> \<le> real (2^(nat ?l) * (m div 2^(nat ?l)) + m mod 2^(nat ?l))" unfolding real_of_int_add by auto
- also have "\<dots> = real m" unfolding zmod_zdiv_equality[symmetric] ..
- finally show ?thesis by auto
- qed
- thus ?thesis unfolding lb_mult_def Float Let_def float.cases if_P[OF True] real_of_float_simp pow2_add mult_commute mult_assoc by auto
- qed
- also have "\<dots> = real (x * y)" unfolding normfloat ..
- finally show ?thesis .
+lemma compute_float_round_down[code]:
+fixes prec m e
+defines "d == bitlen (abs m) - int prec"
+defines "P == 2^nat d"
+defines "f == Float m e"
+shows "float_round_down prec f = (let d = d in
+ if 0 < d then let P = P ; n = m div P in Float n (e + d)
+ else f)"
+ unfolding float_round_down_def float_down_def[symmetric]
+ compute_float_down f_def Let_def P_def round_prec_def d_def bitlen_Float
+ by (simp add: field_simps)
+
+lemma compute_float_round_up[code]:
+fixes prec m e
+defines "d == bitlen (abs m) - int prec"
+defines "P == 2^nat d"
+defines "f == Float m e"
+shows "float_round_up prec f = (let d = d in
+ if 0 < d then let P = P ; n = m div P ; r = m mod P in Float (n + (if r = 0 then 0 else 1)) (e + d)
+ else f)"
+ unfolding float_round_up_def float_up_def[symmetric]
+ compute_float_up f_def Let_def P_def round_prec_def d_def bitlen_Float
+ by (simp add: field_simps)
+
+lemma float_round_up: "real x \<le> real (float_round_up prec x)"
+ using round_up
+ by (simp add: float_round_up_def)
+
+lemma float_round_down: "real (float_round_down prec x) \<le> real x"
+ using round_down
+ by (simp add: float_round_down_def)
+
+instantiation float :: lattice_ab_group_add
+begin
+
+definition inf_float::"float\<Rightarrow>float\<Rightarrow>float"
+where "inf_float a b = min a b"
+
+definition sup_float::"float\<Rightarrow>float\<Rightarrow>float"
+where "sup_float a b = max a b"
+
+instance
+proof
+ fix x y :: float show "inf x y \<le> x" unfolding inf_float_def by simp
+ show "inf x y \<le> y" unfolding inf_float_def by simp
+ show "x \<le> sup x y" unfolding sup_float_def by simp
+ show "y \<le> sup x y" unfolding sup_float_def by simp
+ fix z::float
+ assume "x \<le> y" "x \<le> z" thus "x \<le> inf y z" unfolding inf_float_def by simp
+ next fix x y z :: float
+ assume "y \<le> x" "z \<le> x" thus "sup y z \<le> x" unfolding sup_float_def by simp
qed
-lemma ub_mult: "real (x * y) \<le> real (ub_mult prec x y)"
-proof (cases "normfloat (x * y)")
- case (Float m e)
- hence "odd m \<or> (m = 0 \<and> e = 0)" by (rule normfloat_imp_odd_or_zero)
- let ?l = "bitlen m - int prec"
- have "real (x * y) = real (normfloat (x * y))" unfolding normfloat ..
- also have "\<dots> \<le> real (ub_mult prec x y)"
- proof (cases "?l > 0")
- case False thus ?thesis unfolding ub_mult_def Float Let_def float.cases by auto
- next
- case True
- have "real m \<le> real (m div 2^(nat ?l) + 1) * pow2 ?l"
- proof -
- have "m mod 2^(nat ?l) < 2^(nat ?l)" by (rule pos_mod_bound) auto
- hence mod_uneq: "real (m mod 2^(nat ?l)) \<le> 1 * 2^(nat ?l)" unfolding mult_1 real_of_int_less_iff[symmetric] by auto
-
- have "real m = real (2^(nat ?l) * (m div 2^(nat ?l)) + m mod 2^(nat ?l))" unfolding zmod_zdiv_equality[symmetric] ..
- also have "\<dots> = real (m div 2^(nat ?l)) * 2^(nat ?l) + real (m mod 2^(nat ?l))" unfolding real_of_int_add by auto
- also have "\<dots> \<le> (real (m div 2^(nat ?l)) + 1) * 2^(nat ?l)" unfolding left_distrib using mod_uneq by auto
- finally show ?thesis unfolding pow2_int[symmetric] using True by auto
- qed
- thus ?thesis unfolding ub_mult_def Float Let_def float.cases if_P[OF True] real_of_float_simp pow2_add mult_commute mult_assoc by auto
- qed
- finally show ?thesis .
-qed
-
-primrec float_abs :: "float \<Rightarrow> float" where
- "float_abs (Float m e) = Float \<bar>m\<bar> e"
-
-instantiation float :: abs
-begin
-definition abs_float_def: "\<bar>x\<bar> = float_abs x"
-instance ..
end
-lemma real_of_float_abs: "real \<bar>x :: float\<bar> = \<bar>real x\<bar>"
-proof (cases x)
- case (Float m e)
- have "\<bar>real m\<bar> * pow2 e = \<bar>real m * pow2 e\<bar>" unfolding abs_mult by auto
- thus ?thesis unfolding Float abs_float_def float_abs.simps real_of_float_simp by auto
-qed
+lemma Float_le_zero_iff: "Float a b \<le> 0 \<longleftrightarrow> a \<le> 0"
+ apply (auto simp: zero_float_def mult_le_0_iff)
+ using powr_gt_zero[of 2 b] by simp
+
+(* TODO: how to use as code equation? -> pprt_float?! *)
+lemma compute_pprt[code]: "pprt (Float a e) = (if a <= 0 then 0 else (Float a e))"
+unfolding pprt_def sup_float_def max_def Float_le_zero_iff ..
-primrec floor_fl :: "float \<Rightarrow> float" where
- "floor_fl (Float m e) = (if 0 \<le> e then Float m e
- else Float (m div (2 ^ (nat (-e)))) 0)"
+(* TODO: how to use as code equation? *)
+lemma compute_nprt[code]: "nprt (Float a e) = (if a <= 0 then (Float a e) else 0)"
+unfolding nprt_def inf_float_def min_def Float_le_zero_iff ..
+
+lemma of_float_pprt[simp]: fixes a::float shows "real (pprt a) = pprt (real a)"
+ unfolding pprt_def sup_float_def max_def sup_real_def by (auto simp: less_eq_float_def)
+
+lemma of_float_nprt[simp]: fixes a::float shows "real (nprt a) = nprt (real a)"
+ unfolding nprt_def inf_float_def min_def inf_real_def by (auto simp: less_eq_float_def)
+
+definition int_floor_fl :: "float \<Rightarrow> int" where
+ "int_floor_fl f = floor f"
-lemma floor_fl: "real (floor_fl x) \<le> real x"
-proof (cases x)
- case (Float m e)
- show ?thesis
- proof (cases "0 \<le> e")
- case False
- hence me_eq: "pow2 (-e) = pow2 (int (nat (-e)))" by auto
- have "real (Float (m div (2 ^ (nat (-e)))) 0) = real (m div 2 ^ (nat (-e)))" unfolding real_of_float_simp by auto
- also have "\<dots> \<le> real m / real ((2::int) ^ (nat (-e)))" using real_of_int_div4 .
- also have "\<dots> = real m * inverse (2 ^ (nat (-e)))" unfolding real_of_int_power real_numeral divide_inverse ..
- also have "\<dots> = real (Float m e)" unfolding real_of_float_simp me_eq pow2_int pow2_neg[of e] ..
- finally show ?thesis unfolding Float floor_fl.simps if_not_P[OF `\<not> 0 \<le> e`] .
- next
- case True thus ?thesis unfolding Float by auto
- qed
-qed
+lemma compute_int_floor_fl[code]:
+ shows "int_floor_fl (Float m e) = (if 0 \<le> e then m * 2 ^ nat e
+ else m div (2 ^ (nat (-e))))"
+ by (simp add: int_floor_fl_def powr_int int_of_reals floor_divide_eq_div del: real_of_ints)
+
+definition floor_fl :: "float \<Rightarrow> float" where
+ "floor_fl f = float_of (floor f)"
+
+lemma compute_floor_fl[code]:
+ shows "floor_fl (Float m e) = (if 0 \<le> e then Float m e
+ else Float (m div (2 ^ (nat (-e)))) 0)"
+ by (simp add: floor_fl_def powr_int int_of_reals floor_divide_eq_div del: real_of_ints)
-lemma floor_pos_exp: assumes floor: "Float m e = floor_fl x" shows "0 \<le> e"
-proof (cases x)
- case (Float mx me)
- from floor[unfolded Float floor_fl.simps] show ?thesis by (cases "0 \<le> me", auto)
-qed
+lemma floor_fl: "real (floor_fl x) \<le> real x" by (simp add: floor_fl_def)
+lemma int_floor_fl: "real (int_floor_fl x) \<le> real x" by (simp add: int_floor_fl_def)
-declare floor_fl.simps[simp del]
+lemma floor_pos_exp: "exponent (floor_fl x) \<ge> 0"
+proof cases
+ assume nzero: "floor_fl x \<noteq> float_of 0"
+ have "floor_fl x \<equiv> Float \<lfloor>real x\<rfloor> 0" by (simp add: floor_fl_def)
+ from denormalize_shift[OF this nzero] guess i . note i = this
+ thus ?thesis by simp
+qed (simp add: floor_fl_def)
-primrec ceiling_fl :: "float \<Rightarrow> float" where
+(* TODO: not used in approximation
+definition ceiling_fl :: "float_of \<Rightarrow> float" where
+ "ceiling_fl f = float_of (ceiling f)"
+
+lemma compute_ceiling_fl:
"ceiling_fl (Float m e) = (if 0 \<le> e then Float m e
else Float (m div (2 ^ (nat (-e))) + 1) 0)"
lemma ceiling_fl: "real x \<le> real (ceiling_fl x)"
-proof (cases x)
- case (Float m e)
- show ?thesis
- proof (cases "0 \<le> e")
- case False
- hence me_eq: "pow2 (-e) = pow2 (int (nat (-e)))" by auto
- have "real (Float m e) = real m * inverse (2 ^ (nat (-e)))" unfolding real_of_float_simp me_eq pow2_int pow2_neg[of e] ..
- also have "\<dots> = real m / real ((2::int) ^ (nat (-e)))" unfolding real_of_int_power real_numeral divide_inverse ..
- also have "\<dots> \<le> 1 + real (m div 2 ^ (nat (-e)))" using real_of_int_div3[unfolded diff_le_eq] .
- also have "\<dots> = real (Float (m div (2 ^ (nat (-e))) + 1) 0)" unfolding real_of_float_simp by auto
- finally show ?thesis unfolding Float ceiling_fl.simps if_not_P[OF `\<not> 0 \<le> e`] .
- next
- case True thus ?thesis unfolding Float by auto
- qed
-qed
-declare ceiling_fl.simps[simp del]
+definition lb_mod :: "nat \<Rightarrow> float_of \<Rightarrow> float_of \<Rightarrow> float_of \<Rightarrow> float" where
+"lb_mod prec x ub lb = x - ceiling_fl (float_divr prec x lb) * ub"
-definition lb_mod :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where
- "lb_mod prec x ub lb = x - ceiling_fl (float_divr prec x lb) * ub"
-
-definition ub_mod :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" where
- "ub_mod prec x ub lb = x - floor_fl (float_divl prec x ub) * lb"
+definition ub_mod :: "nat \<Rightarrow> float_of \<Rightarrow> float_of \<Rightarrow> float_of \<Rightarrow> float" where
+"ub_mod prec x ub lb = x - floor_fl (float_divl prec x ub) * lb"
lemma lb_mod: fixes k :: int assumes "0 \<le> real x" and "real k * y \<le> real x" (is "?k * y \<le> ?x")
assumes "0 < real lb" "real lb \<le> y" (is "?lb \<le> y") "y \<le> real ub" (is "y \<le> ?ub")
shows "real (lb_mod prec x ub lb) \<le> ?x - ?k * y"
-proof -
- have "?lb \<le> ?ub" using assms by auto
- have "0 \<le> ?lb" and "?lb \<noteq> 0" using assms by auto
- have "?k * y \<le> ?x" using assms by auto
- also have "\<dots> \<le> ?x / ?lb * ?ub" by (metis mult_left_mono[OF `?lb \<le> ?ub` `0 \<le> ?x`] divide_right_mono[OF _ `0 \<le> ?lb` ] times_divide_eq_left nonzero_mult_divide_cancel_right[OF `?lb \<noteq> 0`])
- also have "\<dots> \<le> real (ceiling_fl (float_divr prec x lb)) * ?ub" by (metis mult_right_mono order_trans `0 \<le> ?lb` `?lb \<le> ?ub` float_divr ceiling_fl)
- finally show ?thesis unfolding lb_mod_def real_of_float_sub real_of_float_mult by auto
-qed
-lemma ub_mod:
- fixes k :: int and x :: float
- assumes "0 \<le> real x" and "real x \<le> real k * y" (is "?x \<le> ?k * y")
+lemma ub_mod: fixes k :: int and x :: float_of assumes "0 \<le> real x" and "real x \<le> real k * y" (is "?x \<le> ?k * y")
assumes "0 < real lb" "real lb \<le> y" (is "?lb \<le> y") "y \<le> real ub" (is "y \<le> ?ub")
shows "?x - ?k * y \<le> real (ub_mod prec x ub lb)"
-proof -
- have "?lb \<le> ?ub" using assms by auto
- hence "0 \<le> ?lb" and "0 \<le> ?ub" and "?ub \<noteq> 0" using assms by auto
- have "real (floor_fl (float_divl prec x ub)) * ?lb \<le> ?x / ?ub * ?lb" by (metis mult_right_mono order_trans `0 \<le> ?lb` `?lb \<le> ?ub` float_divl floor_fl)
- also have "\<dots> \<le> ?x" by (metis mult_left_mono[OF `?lb \<le> ?ub` `0 \<le> ?x`] divide_right_mono[OF _ `0 \<le> ?ub` ] times_divide_eq_left nonzero_mult_divide_cancel_right[OF `?ub \<noteq> 0`])
- also have "\<dots> \<le> ?k * y" using assms by auto
- finally show ?thesis unfolding ub_mod_def real_of_float_sub real_of_float_mult by auto
-qed
-lemma le_float_def'[code]: "f \<le> g = (case f - g of Float a b \<Rightarrow> a \<le> 0)"
-proof -
- have le_transfer: "(f \<le> g) = (real (f - g) \<le> 0)" by (auto simp add: le_float_def)
- from float_split[of "f - g"] obtain a b where f_diff_g: "f - g = Float a b" by auto
- with le_transfer have le_transfer': "f \<le> g = (real (Float a b) \<le> 0)" by simp
- show ?thesis by (simp add: le_transfer' f_diff_g float_le_zero)
-qed
-
-lemma less_float_def'[code]: "f < g = (case f - g of Float a b \<Rightarrow> a < 0)"
-proof -
- have less_transfer: "(f < g) = (real (f - g) < 0)" by (auto simp add: less_float_def)
- from float_split[of "f - g"] obtain a b where f_diff_g: "f - g = Float a b" by auto
- with less_transfer have less_transfer': "f < g = (real (Float a b) < 0)" by simp
- show ?thesis by (simp add: less_transfer' f_diff_g float_less_zero)
-qed
+*)
end
+