Conditions for innocuous double rounding.
Require Import Fcore_Raux.
Require Import Fcore_defs.
Require Import Fcore_generic_fmt.
Require Import Fcalc_ops.
Require Import Fcore_ulp.
Require Import Psatz.
Open Scope R_scope.
Section Double_round.
Variable beta :
radix.
Notation bpow e := (
bpow beta e).
Notation ln_beta x := (
ln_beta beta x).
Definition double_round_eq fexp1 fexp2 choice1 choice2 x :=
round beta fexp1 (
Znearest choice1) (
round beta fexp2 (
Znearest choice2)
x)
=
round beta fexp1 (
Znearest choice1)
x.
A little tactic to simplify terms of the form bpow a * bpow b.
Ltac bpow_simplify :=
repeat
match goal with
| |-
context [(
Fcore_Raux.bpow _ _ *
Fcore_Raux.bpow _ _)] =>
rewrite <-
bpow_plus
| |-
context [(?
X1 *
Fcore_Raux.bpow _ _ *
Fcore_Raux.bpow _ _)] =>
rewrite (
Rmult_assoc X1);
rewrite <-
bpow_plus
| |-
context [(?
X1 * (?
X2 *
Fcore_Raux.bpow _ _) *
Fcore_Raux.bpow _ _)] =>
rewrite <- (
Rmult_assoc X1 X2);
rewrite (
Rmult_assoc (
X1 *
X2));
rewrite <-
bpow_plus
end;
repeat
match goal with
| |-
context [(
Fcore_Raux.bpow _ ?
X)] =>
progress ring_simplify X
end;
change (
Fcore_Raux.bpow _ 0)
with 1;
repeat
match goal with
| |-
context [(
_ * 1)] =>
rewrite Rmult_1_r
end.
Definition midp (
fexp :
Z ->
Z) (
x :
R) :=
round beta fexp Zfloor x + / 2 *
ulp beta fexp x.
Definition midp' (
fexp :
Z ->
Z) (
x :
R) :=
round beta fexp Zceil x - / 2 *
ulp beta fexp x.
Lemma double_round_lt_mid_further_place' :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
forall x,
0 <
x ->
(
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x) - 1)%
Z ->
x <
bpow (
ln_beta x) - / 2 *
ulp beta fexp2 x ->
x <
midp fexp1 x - / 2 *
ulp beta fexp2 x ->
double_round_eq fexp1 fexp2 choice1 choice2 x.
Proof.
Lemma double_round_lt_mid_further_place :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
forall x,
0 <
x ->
(
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x) - 1)%
Z ->
(
fexp1 (
ln_beta x) <=
ln_beta x)%
Z ->
x <
midp fexp1 x - / 2 *
ulp beta fexp2 x ->
double_round_eq fexp1 fexp2 choice1 choice2 x.
Proof.
Lemma double_round_lt_mid_same_place :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
forall (
choice1 choice2 :
Z ->
bool),
forall x,
0 <
x ->
(
fexp2 (
ln_beta x) =
fexp1 (
ln_beta x))%
Z ->
x <
midp fexp1 x ->
double_round_eq fexp1 fexp2 choice1 choice2 x.
Proof.
Lemma double_round_lt_mid :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
forall x,
0 <
x ->
(
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x))%
Z ->
(
fexp1 (
ln_beta x) <=
ln_beta x)%
Z ->
x <
midp fexp1 x ->
((
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x) - 1)%
Z ->
x <
midp fexp1 x - / 2 *
ulp beta fexp2 x) ->
double_round_eq fexp1 fexp2 choice1 choice2 x.
Proof.
Lemma double_round_gt_mid_further_place' :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
forall x,
0 <
x ->
(
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x) - 1)%
Z ->
round beta fexp2 (
Znearest choice2)
x <
bpow (
ln_beta x) ->
midp'
fexp1 x + / 2 *
ulp beta fexp2 x <
x ->
double_round_eq fexp1 fexp2 choice1 choice2 x.
Proof.
Lemma double_round_gt_mid_further_place :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
forall x,
0 <
x ->
(
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x) - 1)%
Z ->
(
fexp1 (
ln_beta x) <=
ln_beta x)%
Z ->
midp'
fexp1 x + / 2 *
ulp beta fexp2 x <
x ->
double_round_eq fexp1 fexp2 choice1 choice2 x.
Proof.
Lemma double_round_gt_mid_same_place :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
forall (
choice1 choice2 :
Z ->
bool),
forall x,
0 <
x ->
(
fexp2 (
ln_beta x) =
fexp1 (
ln_beta x))%
Z ->
midp'
fexp1 x <
x ->
double_round_eq fexp1 fexp2 choice1 choice2 x.
Proof.
Lemma double_round_gt_mid :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
forall x,
0 <
x ->
(
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x))%
Z ->
(
fexp1 (
ln_beta x) <=
ln_beta x)%
Z ->
midp'
fexp1 x <
x ->
((
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x) - 1)%
Z ->
midp'
fexp1 x + / 2 *
ulp beta fexp2 x <
x) ->
double_round_eq fexp1 fexp2 choice1 choice2 x.
Proof.
Section Double_round_mult.
Lemma ln_beta_mult_disj :
forall x y,
x <> 0 ->
y <> 0 ->
((
ln_beta (
x *
y) = (
ln_beta x +
ln_beta y - 1)%
Z :>
Z)
\/ (
ln_beta (
x *
y) = (
ln_beta x +
ln_beta y)%
Z :>
Z)).
Proof.
Definition double_round_mult_hyp fexp1 fexp2 :=
(
forall ex ey, (
fexp2 (
ex +
ey) <=
fexp1 ex +
fexp1 ey)%
Z)
/\ (
forall ex ey, (
fexp2 (
ex +
ey - 1) <=
fexp1 ex +
fexp1 ey)%
Z).
Lemma double_round_mult_aux :
forall (
fexp1 fexp2 :
Z ->
Z),
double_round_mult_hyp fexp1 fexp2 ->
forall x y,
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
generic_format beta fexp2 (
x *
y).
Proof.
Variable rnd :
R ->
Z.
Context {
valid_rnd :
Valid_rnd rnd }.
Theorem double_round_mult :
forall (
fexp1 fexp2 :
Z ->
Z),
double_round_mult_hyp fexp1 fexp2 ->
forall x y,
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
round beta fexp1 rnd (
round beta fexp2 rnd (
x *
y))
=
round beta fexp1 rnd (
x *
y).
Proof.
Section Double_round_mult_FLX.
Require Import Fcore_FLX.
Variable prec :
Z.
Variable prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Theorem double_round_mult_FLX :
(2 *
prec <=
prec')%
Z ->
forall x y,
FLX_format beta prec x ->
FLX_format beta prec y ->
round beta (
FLX_exp prec)
rnd (
round beta (
FLX_exp prec')
rnd (
x *
y))
=
round beta (
FLX_exp prec)
rnd (
x *
y).
Proof.
End Double_round_mult_FLX.
Section Double_round_mult_FLT.
Require Import Fcore_FLX.
Require Import Fcore_FLT.
Variable emin prec :
Z.
Variable emin'
prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Theorem double_round_mult_FLT :
(
emin' <= 2 *
emin)%
Z -> (2 *
prec <=
prec')%
Z ->
forall x y,
FLT_format beta emin prec x ->
FLT_format beta emin prec y ->
round beta (
FLT_exp emin prec)
rnd
(
round beta (
FLT_exp emin'
prec')
rnd (
x *
y))
=
round beta (
FLT_exp emin prec)
rnd (
x *
y).
Proof.
End Double_round_mult_FLT.
Section Double_round_mult_FTZ.
Require Import Fcore_FTZ.
Variable emin prec :
Z.
Variable emin'
prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Theorem double_round_mult_FTZ :
(
emin' +
prec' <= 2 *
emin +
prec)%
Z ->
(2 *
prec <=
prec')%
Z ->
forall x y,
FTZ_format beta emin prec x ->
FTZ_format beta emin prec y ->
round beta (
FTZ_exp emin prec)
rnd
(
round beta (
FTZ_exp emin'
prec')
rnd (
x *
y))
=
round beta (
FTZ_exp emin prec)
rnd (
x *
y).
Proof.
End Double_round_mult_FTZ.
End Double_round_mult.
Section Double_round_plus.
Lemma ln_beta_plus_disj :
forall x y,
0 <
y ->
y <=
x ->
((
ln_beta (
x +
y) =
ln_beta x :>
Z)
\/ (
ln_beta (
x +
y) = (
ln_beta x + 1)%
Z :>
Z)).
Proof.
Lemma ln_beta_plus_separated :
forall fexp :
Z ->
Z,
forall x y,
0 <
x -> 0 <=
y ->
generic_format beta fexp x ->
(
ln_beta y <=
fexp (
ln_beta x))%
Z ->
(
ln_beta (
x +
y) =
ln_beta x :>
Z).
Proof.
Lemma ln_beta_minus_disj :
forall x y,
0 <
x -> 0 <
y ->
(
ln_beta y <=
ln_beta x - 2)%
Z ->
((
ln_beta (
x -
y) =
ln_beta x :>
Z)
\/ (
ln_beta (
x -
y) = (
ln_beta x - 1)%
Z :>
Z)).
Proof.
Lemma ln_beta_minus_separated :
forall fexp :
Z ->
Z,
Valid_exp fexp ->
forall x y,
0 <
x -> 0 <
y ->
y <
x ->
bpow (
ln_beta x - 1) <
x ->
generic_format beta fexp x -> (
ln_beta y <=
fexp (
ln_beta x))%
Z ->
(
ln_beta (
x -
y) =
ln_beta x :>
Z).
Proof.
Definition double_round_plus_hyp fexp1 fexp2 :=
(
forall ex ey, (
fexp1 (
ex + 1) - 1 <=
ey)%
Z -> (
fexp2 ex <=
fexp1 ey)%
Z)
/\ (
forall ex ey, (
fexp1 (
ex - 1) + 1 <=
ey)%
Z -> (
fexp2 ex <=
fexp1 ey)%
Z)
/\ (
forall ex ey, (
fexp1 ex - 1 <=
ey)%
Z -> (
fexp2 ex <=
fexp1 ey)%
Z)
/\ (
forall ex ey, (
ex - 1 <=
ey)%
Z -> (
fexp2 ex <=
fexp1 ey)%
Z).
Lemma double_round_plus_aux0_aux_aux :
forall (
fexp1 fexp2 :
Z ->
Z),
forall x y,
(
fexp1 (
ln_beta x) <=
fexp1 (
ln_beta y))%
Z ->
(
fexp2 (
ln_beta (
x +
y))%
Z <=
fexp1 (
ln_beta x))%
Z ->
(
fexp2 (
ln_beta (
x +
y))%
Z <=
fexp1 (
ln_beta y))%
Z ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
generic_format beta fexp2 (
x +
y).
Proof.
Lemma double_round_plus_aux0_aux :
forall (
fexp1 fexp2 :
Z ->
Z),
forall x y,
(
fexp2 (
ln_beta (
x +
y))%
Z <=
fexp1 (
ln_beta x))%
Z ->
(
fexp2 (
ln_beta (
x +
y))%
Z <=
fexp1 (
ln_beta y))%
Z ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
generic_format beta fexp2 (
x +
y).
Proof.
Lemma double_round_plus_aux0 :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
double_round_plus_hyp fexp1 fexp2 ->
forall x y,
(0 <
x)%
R -> (0 <
y)%
R -> (
y <=
x)%
R ->
(
fexp1 (
ln_beta x) - 1 <=
ln_beta y)%
Z ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
generic_format beta fexp2 (
x +
y).
Proof.
Lemma double_round_plus_aux1_aux :
forall k, (0 <
k)%
Z ->
forall (
fexp :
Z ->
Z),
forall x y,
0 <
x -> 0 <
y ->
(
ln_beta y <=
fexp (
ln_beta x) -
k)%
Z ->
(
ln_beta (
x +
y) =
ln_beta x :>
Z) ->
generic_format beta fexp x ->
0 < (
x +
y) -
round beta fexp Zfloor (
x +
y) <
bpow (
fexp (
ln_beta x) -
k).
Proof.
Lemma double_round_plus_aux1 :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_hyp fexp1 fexp2 ->
forall x y,
0 <
x -> 0 <
y ->
(
ln_beta y <=
fexp1 (
ln_beta x) - 2)%
Z ->
generic_format beta fexp1 x ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x +
y).
Proof.
assert (
Hbeta : (2 <=
beta)%
Z).
{
destruct beta as (
beta_val,
beta_prop).
now apply Zle_bool_imp_le. }
intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Px Py Hly Fx.
assert (
Lxy :
ln_beta (
x +
y) =
ln_beta x :>
Z);
[
now apply (
ln_beta_plus_separated fexp1); [|
apply Rlt_le| |
omega]|].
destruct Hexp as (
_,(
_,(
_,
Hexp4))).
assert (
Hf2 : (
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x))%
Z);
[
now apply Hexp4;
omega|].
assert (
Bpow2 :
bpow (- 2) <= / 2 * / 2).
{
unfold Fcore_Raux.bpow,
Z.pow_pos;
simpl.
rewrite <-
Rinv_mult_distr; [|
lra|
lra].
apply Rinv_le; [
lra|].
change 4
with (
Z2R (2 * 2));
apply Z2R_le;
rewrite Zmult_1_r.
now apply Zmult_le_compat;
omega. }
assert (
P2 : (0 < 2)%
Z)
by omega.
unfold double_round_eq.
apply double_round_lt_mid.
-
exact Vfexp1.
-
exact Vfexp2.
-
lra.
-
now rewrite Lxy.
-
rewrite Lxy.
assert (
fexp1 (
ln_beta x) <
ln_beta x)%
Z; [|
omega].
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|].
-
unfold midp.
apply (
Rplus_lt_reg_r (-
round beta fexp1 Zfloor (
x +
y))).
apply (
Rlt_le_trans _ _ _ (
proj2 (
double_round_plus_aux1_aux 2
P2 fexp1 x y Px
Py Hly Lxy Fx))).
ring_simplify.
rewrite ulp_neq_0;[
idtac|
now apply Rgt_not_eq,
Rplus_lt_0_compat].
unfold canonic_exp;
rewrite Lxy.
apply (
Rmult_le_reg_r (
bpow (-
fexp1 (
ln_beta x)))); [
now apply bpow_gt_0|].
bpow_simplify.
apply (
Rle_trans _ _ _ Bpow2).
rewrite <- (
Rmult_1_r (/ 2))
at 3.
apply Rmult_le_compat_l;
lra.
-
rewrite ulp_neq_0;[
idtac|
now apply Rgt_not_eq,
Rplus_lt_0_compat].
unfold round,
F2R,
scaled_mantissa,
canonic_exp;
simpl;
rewrite Lxy.
intro Hf2'.
apply (
Rmult_lt_reg_r (
bpow (-
fexp1 (
ln_beta x))));
[
now apply bpow_gt_0|].
apply (
Rmult_lt_reg_r (
bpow (
fexp1 (
ln_beta x)))); [
now apply bpow_gt_0|].
bpow_simplify.
apply (
Rplus_lt_reg_r (-
round beta fexp1 Zfloor (
x +
y))).
unfold midp;
ring_simplify.
apply (
Rlt_le_trans _ _ _ (
proj2 (
double_round_plus_aux1_aux 2
P2 fexp1 x y Px
Py Hly Lxy Fx))).
apply (
Rmult_le_reg_r (
bpow (-
fexp1 (
ln_beta x)))); [
now apply bpow_gt_0|].
rewrite ulp_neq_0;[
idtac|
now apply Rgt_not_eq,
Rplus_lt_0_compat].
unfold canonic_exp;
rewrite Lxy,
Rmult_minus_distr_r;
bpow_simplify.
apply (
Rle_trans _ _ _ Bpow2).
rewrite <- (
Rmult_1_r (/ 2))
at 3;
rewrite <-
Rmult_minus_distr_l.
apply Rmult_le_compat_l; [
lra|].
apply (
Rplus_le_reg_r (- 1));
ring_simplify.
replace (
_ -
_)
with (- (/ 2))
by lra.
apply Ropp_le_contravar.
{
apply Rle_trans with (
bpow (- 1)).
-
apply bpow_le;
omega.
-
unfold Fcore_Raux.bpow,
Z.pow_pos;
simpl.
apply Rinv_le; [
lra|].
change 2
with (
Z2R 2);
apply Z2R_le;
omega. }
Qed.
Lemma double_round_plus_aux2 :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_hyp fexp1 fexp2 ->
forall x y,
0 <
x -> 0 <
y ->
y <=
x ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x +
y).
Proof.
Lemma double_round_plus_aux :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_hyp fexp1 fexp2 ->
forall x y,
0 <=
x -> 0 <=
y ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x +
y).
Proof.
Lemma double_round_minus_aux0_aux :
forall (
fexp1 fexp2 :
Z ->
Z),
forall x y,
(
fexp2 (
ln_beta (
x -
y))%
Z <=
fexp1 (
ln_beta x))%
Z ->
(
fexp2 (
ln_beta (
x -
y))%
Z <=
fexp1 (
ln_beta y))%
Z ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
generic_format beta fexp2 (
x -
y).
Proof.
Lemma double_round_minus_aux0 :
forall (
fexp1 fexp2 :
Z ->
Z),
double_round_plus_hyp fexp1 fexp2 ->
forall x y,
0 <
y ->
y <
x ->
(
fexp1 (
ln_beta x) - 1 <=
ln_beta y)%
Z ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
generic_format beta fexp2 (
x -
y).
Proof.
Lemma double_round_minus_aux1 :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
double_round_plus_hyp fexp1 fexp2 ->
forall x y,
0 <
y ->
y <
x ->
(
ln_beta y <=
fexp1 (
ln_beta x) - 2)%
Z ->
(
fexp1 (
ln_beta (
x -
y)) - 1 <=
ln_beta y)%
Z ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
generic_format beta fexp2 (
x -
y).
Proof.
Lemma double_round_minus_aux2_aux :
forall (
fexp :
Z ->
Z),
Valid_exp fexp ->
forall x y,
0 <
y ->
y <
x ->
(
ln_beta y <=
fexp (
ln_beta x) - 1)%
Z ->
generic_format beta fexp x ->
generic_format beta fexp y ->
round beta fexp Zceil (
x -
y) - (
x -
y) <=
y.
Proof.
intros fexp Vfexp x y Py Hxy Hly Fx Fy.
assert (
Px :=
Rlt_trans 0
y x Py Hxy).
revert Fx.
unfold generic_format,
F2R,
scaled_mantissa,
canonic_exp;
simpl.
set (
mx :=
Ztrunc (
x *
bpow (-
fexp (
ln_beta x)))).
intro Fx.
assert (
Hfx : (
fexp (
ln_beta x) <
ln_beta x)%
Z);
[
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|]|].
assert (
Hfy : (
fexp (
ln_beta y) <
ln_beta y)%
Z);
[
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|]|].
destruct (
Rlt_or_le (
bpow (
ln_beta x - 1))
x)
as [
Hx|
Hx].
-
assert (
Lxy :
ln_beta (
x -
y) =
ln_beta x :>
Z);
[
now apply (
ln_beta_minus_separated fexp); [| | | | | |
omega]|].
assert (
Rxy :
round beta fexp Zceil (
x -
y) =
x).
{
unfold round,
F2R,
scaled_mantissa,
canonic_exp;
simpl.
rewrite Lxy.
apply eq_sym;
rewrite Fx at 1;
apply eq_sym.
apply Rmult_eq_compat_r.
apply f_equal.
rewrite Fx at 1.
rewrite Rmult_minus_distr_r.
bpow_simplify.
apply Zceil_imp.
split.
-
unfold Zminus;
rewrite Z2R_plus.
apply Rplus_lt_compat_l.
apply Ropp_lt_contravar;
simpl.
apply (
Rmult_lt_reg_r (
bpow (
fexp (
ln_beta x))));
[
now apply bpow_gt_0|].
rewrite Rmult_1_l;
bpow_simplify.
apply Rlt_le_trans with (
bpow (
ln_beta y)).
+
rewrite <-
Rabs_right at 1; [|
now apply Rle_ge;
apply Rlt_le].
apply bpow_ln_beta_gt.
+
apply bpow_le.
omega.
-
rewrite <- (
Rplus_0_r (
Z2R _))
at 2.
apply Rplus_le_compat_l.
rewrite <-
Ropp_0;
apply Ropp_le_contravar.
rewrite <- (
Rmult_0_r y).
apply Rmult_le_compat_l; [
now apply Rlt_le|].
now apply bpow_ge_0. }
rewrite Rxy;
ring_simplify.
apply Rle_refl.
-
assert (
Xpow :
x =
bpow (
ln_beta x - 1)).
{
apply Rle_antisym; [
exact Hx|].
destruct (
ln_beta x)
as (
ex,
Hex);
simpl.
rewrite <- (
Rabs_right x); [|
now apply Rle_ge;
apply Rlt_le].
apply Hex.
now apply Rgt_not_eq. }
assert (
Lxy : (
ln_beta (
x -
y) =
ln_beta x - 1 :>
Z)%
Z).
{
apply Zle_antisym.
-
apply ln_beta_le_bpow.
+
apply Rminus_eq_contra.
now intro Hx';
rewrite Hx'
in Hxy;
apply (
Rlt_irrefl y).
+
rewrite Rabs_right;
lra.
-
apply (
ln_beta_minus_lb beta x y Px Py).
omega. }
assert (
Hfx1 : (
fexp (
ln_beta x - 1) <
ln_beta x - 1)%
Z);
[
now apply (
valid_exp_large fexp (
ln_beta y)); [|
omega]|].
assert (
Rxy :
round beta fexp Zceil (
x -
y) <=
x).
{
rewrite Xpow at 2.
unfold round,
F2R,
scaled_mantissa,
canonic_exp;
simpl.
rewrite Lxy.
apply (
Rmult_le_reg_r (
bpow (-
fexp (
ln_beta x - 1)%
Z)));
[
now apply bpow_gt_0|].
bpow_simplify.
rewrite <- (
Z2R_Zpower beta (
_ -
_ -
_)); [|
omega].
apply Z2R_le.
apply Zceil_glb.
rewrite Z2R_Zpower; [|
omega].
rewrite Xpow at 1.
rewrite Rmult_minus_distr_r.
bpow_simplify.
rewrite <- (
Rplus_0_r (
bpow _))
at 2.
apply Rplus_le_compat_l.
rewrite <-
Ropp_0;
apply Ropp_le_contravar.
rewrite <- (
Rmult_0_r y).
apply Rmult_le_compat_l; [
now apply Rlt_le|].
now apply bpow_ge_0. }
lra.
Qed.
Lemma double_round_minus_aux2 :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_hyp fexp1 fexp2 ->
forall x y,
0 <
y ->
y <
x ->
(
ln_beta y <=
fexp1 (
ln_beta x) - 2)%
Z ->
(
ln_beta y <=
fexp1 (
ln_beta (
x -
y)) - 2)%
Z ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x -
y).
Proof.
assert (
Hbeta : (2 <=
beta)%
Z).
{
destruct beta as (
beta_val,
beta_prop).
now apply Zle_bool_imp_le. }
intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Py Hxy Hly Hly'
Fx Fy.
assert (
Px :=
Rlt_trans 0
y x Py Hxy).
destruct Hexp as (
_,(
_,(
_,
Hexp4))).
assert (
Hf2 : (
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x))%
Z);
[
now apply Hexp4;
omega|].
assert (
Hfx : (
fexp1 (
ln_beta x) <
ln_beta x)%
Z);
[
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|]|].
assert (
Bpow2 :
bpow (- 2) <= / 2 * / 2).
{
unfold Fcore_Raux.bpow,
Z.pow_pos;
simpl.
rewrite <-
Rinv_mult_distr; [|
lra|
lra].
apply Rinv_le; [
lra|].
change 4
with (
Z2R (2 * 2));
apply Z2R_le;
rewrite Zmult_1_r.
now apply Zmult_le_compat;
omega. }
assert (
Ly :
y <
bpow (
ln_beta y)).
{
apply Rabs_lt_inv.
apply bpow_ln_beta_gt. }
unfold double_round_eq.
apply double_round_gt_mid.
-
exact Vfexp1.
-
exact Vfexp2.
-
lra.
-
apply Hexp4;
omega.
-
assert (
fexp1 (
ln_beta (
x -
y)) <
ln_beta (
x -
y))%
Z; [|
omega].
apply (
valid_exp_large fexp1 (
ln_beta x - 1)).
+
apply (
valid_exp_large fexp1 (
ln_beta y)); [|
omega].
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|].
+
now apply ln_beta_minus_lb; [| |
omega].
-
unfold midp'.
apply (
Rplus_lt_reg_r (/ 2 *
ulp beta fexp1 (
x -
y) - (
x -
y))).
ring_simplify.
replace (
_ +
_)
with (
round beta fexp1 Zceil (
x -
y) - (
x -
y))
by ring.
apply Rlt_le_trans with (
bpow (
fexp1 (
ln_beta (
x -
y)) - 2)).
+
apply Rle_lt_trans with y;
[
now apply double_round_minus_aux2_aux;
try assumption;
omega|].
apply (
Rlt_le_trans _ _ _ Ly).
now apply bpow_le.
+
rewrite ulp_neq_0;[
idtac|
now apply sym_not_eq,
Rlt_not_eq,
Rgt_minus].
unfold canonic_exp.
replace (
_ - 2)%
Z with (
fexp1 (
ln_beta (
x -
y)) - 1 - 1)%
Z by ring.
unfold Zminus at 1;
rewrite bpow_plus.
rewrite Rmult_comm.
apply Rmult_le_compat.
*
now apply bpow_ge_0.
*
now apply bpow_ge_0.
*
unfold Fcore_Raux.bpow,
Z.pow_pos;
simpl.
rewrite Zmult_1_r;
apply Rinv_le.
lra.
now change 2
with (
Z2R 2);
apply Z2R_le.
*
apply bpow_le;
omega.
-
intro Hf2'.
unfold midp'.
apply (
Rplus_lt_reg_r (/ 2 *
ulp beta fexp1 (
x -
y) - (
x -
y)
- / 2 *
ulp beta fexp2 (
x -
y))).
ring_simplify.
replace (
_ +
_)
with (
round beta fexp1 Zceil (
x -
y) - (
x -
y))
by ring.
apply Rle_lt_trans with y;
[
now apply double_round_minus_aux2_aux;
try assumption;
omega|].
apply (
Rlt_le_trans _ _ _ Ly).
apply Rle_trans with (
bpow (
fexp1 (
ln_beta (
x -
y)) - 2));
[
now apply bpow_le|].
replace (
_ - 2)%
Z with (
fexp1 (
ln_beta (
x -
y)) - 1 - 1)%
Z by ring.
unfold Zminus at 1;
rewrite bpow_plus.
rewrite <-
Rmult_minus_distr_l.
rewrite Rmult_comm;
apply Rmult_le_compat.
+
apply bpow_ge_0.
+
apply bpow_ge_0.
+
unfold Fcore_Raux.bpow,
Z.pow_pos;
simpl.
rewrite Zmult_1_r;
apply Rinv_le; [
lra|].
now change 2
with (
Z2R 2);
apply Z2R_le.
+
rewrite 2!
ulp_neq_0;
try now apply Rgt_not_eq,
Rgt_minus.
unfold canonic_exp.
apply (
Rplus_le_reg_r (
bpow (
fexp2 (
ln_beta (
x -
y)))));
ring_simplify.
apply Rle_trans with (2 *
bpow (
fexp1 (
ln_beta (
x -
y)) - 1)).
*
rewrite Rmult_plus_distr_r;
rewrite Rmult_1_l.
apply Rplus_le_compat_l.
now apply bpow_le.
*
unfold Zminus;
rewrite bpow_plus.
rewrite Rmult_comm;
rewrite Rmult_assoc.
rewrite <-
Rmult_1_r.
apply Rmult_le_compat_l; [
now apply bpow_ge_0|].
unfold Fcore_Raux.bpow,
Z.pow_pos;
simpl.
rewrite Zmult_1_r.
rewrite <- (
Rinv_r 2)
at 3; [|
lra].
rewrite Rmult_comm;
apply Rmult_le_compat_l; [
lra|].
apply Rinv_le; [
lra|].
now change 2
with (
Z2R 2);
apply Z2R_le.
Qed.
Lemma double_round_minus_aux3 :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_hyp fexp1 fexp2 ->
forall x y,
0 <
y ->
y <=
x ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x -
y).
Proof.
Lemma double_round_minus_aux :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_hyp fexp1 fexp2 ->
forall x y,
0 <=
x -> 0 <=
y ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x -
y).
Proof.
Lemma double_round_plus :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_hyp fexp1 fexp2 ->
forall x y,
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x +
y).
Proof.
Lemma double_round_minus :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_hyp fexp1 fexp2 ->
forall x y,
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x -
y).
Proof.
Section Double_round_plus_FLX.
Require Import Fcore_FLX.
Variable prec :
Z.
Variable prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FLX_double_round_plus_hyp :
(2 *
prec + 1 <=
prec')%
Z ->
double_round_plus_hyp (
FLX_exp prec) (
FLX_exp prec').
Proof.
Theorem double_round_plus_FLX :
forall choice1 choice2,
(2 *
prec + 1 <=
prec')%
Z ->
forall x y,
FLX_format beta prec x ->
FLX_format beta prec y ->
double_round_eq (
FLX_exp prec) (
FLX_exp prec')
choice1 choice2 (
x +
y).
Proof.
Theorem double_round_minus_FLX :
forall choice1 choice2,
(2 *
prec + 1 <=
prec')%
Z ->
forall x y,
FLX_format beta prec x ->
FLX_format beta prec y ->
double_round_eq (
FLX_exp prec) (
FLX_exp prec')
choice1 choice2 (
x -
y).
Proof.
End Double_round_plus_FLX.
Section Double_round_plus_FLT.
Require Import Fcore_FLX.
Require Import Fcore_FLT.
Variable emin prec :
Z.
Variable emin'
prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FLT_double_round_plus_hyp :
(
emin' <=
emin)%
Z -> (2 *
prec + 1 <=
prec')%
Z ->
double_round_plus_hyp (
FLT_exp emin prec) (
FLT_exp emin'
prec').
Proof.
Theorem double_round_plus_FLT :
forall choice1 choice2,
(
emin' <=
emin)%
Z -> (2 *
prec + 1 <=
prec')%
Z ->
forall x y,
FLT_format beta emin prec x ->
FLT_format beta emin prec y ->
double_round_eq (
FLT_exp emin prec) (
FLT_exp emin'
prec')
choice1 choice2 (
x +
y).
Proof.
Theorem double_round_minus_FLT :
forall choice1 choice2,
(
emin' <=
emin)%
Z -> (2 *
prec + 1 <=
prec')%
Z ->
forall x y,
FLT_format beta emin prec x ->
FLT_format beta emin prec y ->
double_round_eq (
FLT_exp emin prec) (
FLT_exp emin'
prec')
choice1 choice2 (
x -
y).
Proof.
End Double_round_plus_FLT.
Section Double_round_plus_FTZ.
Require Import Fcore_FLX.
Require Import Fcore_FTZ.
Variable emin prec :
Z.
Variable emin'
prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FTZ_double_round_plus_hyp :
(
emin' +
prec' <=
emin + 1)%
Z -> (2 *
prec + 1 <=
prec')%
Z ->
double_round_plus_hyp (
FTZ_exp emin prec) (
FTZ_exp emin'
prec').
Proof.
Theorem double_round_plus_FTZ :
forall choice1 choice2,
(
emin' +
prec' <=
emin + 1)%
Z -> (2 *
prec + 1 <=
prec')%
Z ->
forall x y,
FTZ_format beta emin prec x ->
FTZ_format beta emin prec y ->
double_round_eq (
FTZ_exp emin prec) (
FTZ_exp emin'
prec')
choice1 choice2 (
x +
y).
Proof.
Theorem double_round_minus_FTZ :
forall choice1 choice2,
(
emin' +
prec' <=
emin + 1)%
Z -> (2 *
prec + 1 <=
prec')%
Z ->
forall x y,
FTZ_format beta emin prec x ->
FTZ_format beta emin prec y ->
double_round_eq (
FTZ_exp emin prec) (
FTZ_exp emin'
prec')
choice1 choice2 (
x -
y).
Proof.
End Double_round_plus_FTZ.
Section Double_round_plus_beta_ge_3.
Definition double_round_plus_beta_ge_3_hyp fexp1 fexp2 :=
(
forall ex ey, (
fexp1 (
ex + 1) <=
ey)%
Z -> (
fexp2 ex <=
fexp1 ey)%
Z)
/\ (
forall ex ey, (
fexp1 (
ex - 1) + 1 <=
ey)%
Z -> (
fexp2 ex <=
fexp1 ey)%
Z)
/\ (
forall ex ey, (
fexp1 ex <=
ey)%
Z -> (
fexp2 ex <=
fexp1 ey)%
Z)
/\ (
forall ex ey, (
ex - 1 <=
ey)%
Z -> (
fexp2 ex <=
fexp1 ey)%
Z).
Lemma double_round_plus_beta_ge_3_aux0 :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
double_round_plus_beta_ge_3_hyp fexp1 fexp2 ->
forall x y,
(0 <
y)%
R -> (
y <=
x)%
R ->
(
fexp1 (
ln_beta x) <=
ln_beta y)%
Z ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
generic_format beta fexp2 (
x +
y).
Proof.
Lemma double_round_plus_beta_ge_3_aux1 :
(3 <=
beta)%
Z ->
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_beta_ge_3_hyp fexp1 fexp2 ->
forall x y,
0 <
x -> 0 <
y ->
(
ln_beta y <=
fexp1 (
ln_beta x) - 1)%
Z ->
generic_format beta fexp1 x ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x +
y).
Proof.
intros Hbeta fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Px Py Hly Fx.
assert (
Lxy :
ln_beta (
x +
y) =
ln_beta x :>
Z);
[
now apply (
ln_beta_plus_separated fexp1); [|
apply Rlt_le| |
omega]|].
destruct Hexp as (
_,(
_,(
_,
Hexp4))).
assert (
Hf2 : (
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x))%
Z);
[
now apply Hexp4;
omega|].
assert (
Bpow3 :
bpow (- 1) <= / 3).
{
unfold Fcore_Raux.bpow,
Z.pow_pos;
simpl.
rewrite Zmult_1_r.
apply Rinv_le; [
lra|].
now change 3
with (
Z2R 3);
apply Z2R_le. }
assert (
P1 : (0 < 1)%
Z)
by omega.
unfold double_round_eq.
apply double_round_lt_mid.
-
exact Vfexp1.
-
exact Vfexp2.
-
lra.
-
now rewrite Lxy.
-
rewrite Lxy.
assert (
fexp1 (
ln_beta x) <
ln_beta x)%
Z; [|
omega].
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|].
-
unfold midp.
apply (
Rplus_lt_reg_r (-
round beta fexp1 Zfloor (
x +
y))).
apply (
Rlt_le_trans _ _ _ (
proj2 (
double_round_plus_aux1_aux 1
P1 fexp1 x y Px
Py Hly Lxy Fx))).
ring_simplify.
rewrite ulp_neq_0;[
idtac|
now apply Rgt_not_eq,
Rplus_lt_0_compat].
unfold canonic_exp;
rewrite Lxy.
apply (
Rmult_le_reg_r (
bpow (-
fexp1 (
ln_beta x))));
[
now apply bpow_gt_0|].
bpow_simplify.
apply (
Rle_trans _ _ _ Bpow3);
lra.
-
rewrite ulp_neq_0;[
idtac|
now apply Rgt_not_eq,
Rplus_lt_0_compat].
unfold round,
F2R,
scaled_mantissa,
canonic_exp;
simpl;
rewrite Lxy.
intro Hf2'.
unfold midp.
apply (
Rplus_lt_reg_r (-
round beta fexp1 Zfloor (
x +
y)));
ring_simplify.
rewrite <-
Rmult_minus_distr_l.
apply (
Rlt_le_trans _ _ _ (
proj2 (
double_round_plus_aux1_aux 1
P1 fexp1 x y Px
Py Hly Lxy Fx))).
rewrite ulp_neq_0;[
idtac|
now apply Rgt_not_eq,
Rplus_lt_0_compat].
unfold canonic_exp;
rewrite Lxy.
apply (
Rmult_le_reg_r (
bpow (-
fexp1 (
ln_beta x))));
[
now apply bpow_gt_0|].
rewrite (
Rmult_assoc (/ 2)).
rewrite Rmult_minus_distr_r.
bpow_simplify.
apply (
Rle_trans _ _ _ Bpow3).
apply Rle_trans with (/ 2 * (2 / 3)); [
lra|].
apply Rmult_le_compat_l; [
lra|].
apply (
Rplus_le_reg_r (- 1));
ring_simplify.
replace (
_ -
_)
with (- (/ 3))
by lra.
apply Ropp_le_contravar.
now apply Rle_trans with (
bpow (- 1)); [
apply bpow_le;
omega|].
Qed.
Lemma double_round_plus_beta_ge_3_aux2 :
(3 <=
beta)%
Z ->
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_beta_ge_3_hyp fexp1 fexp2 ->
forall x y,
0 <
y ->
y <=
x ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x +
y).
Proof.
Lemma double_round_plus_beta_ge_3_aux :
(3 <=
beta)%
Z ->
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_beta_ge_3_hyp fexp1 fexp2 ->
forall x y,
0 <=
x -> 0 <=
y ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x +
y).
Proof.
Lemma double_round_minus_beta_ge_3_aux0 :
forall (
fexp1 fexp2 :
Z ->
Z),
double_round_plus_beta_ge_3_hyp fexp1 fexp2 ->
forall x y,
0 <
y ->
y <
x ->
(
fexp1 (
ln_beta x) <=
ln_beta y)%
Z ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
generic_format beta fexp2 (
x -
y).
Proof.
Lemma double_round_minus_beta_ge_3_aux1 :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
double_round_plus_beta_ge_3_hyp fexp1 fexp2 ->
forall x y,
0 <
y ->
y <
x ->
(
ln_beta y <=
fexp1 (
ln_beta x) - 1)%
Z ->
(
fexp1 (
ln_beta (
x -
y)) <=
ln_beta y)%
Z ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
generic_format beta fexp2 (
x -
y).
Proof.
Lemma double_round_minus_beta_ge_3_aux2 :
(3 <=
beta)%
Z ->
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_beta_ge_3_hyp fexp1 fexp2 ->
forall x y,
0 <
y ->
y <
x ->
(
ln_beta y <=
fexp1 (
ln_beta x) - 1)%
Z ->
(
ln_beta y <=
fexp1 (
ln_beta (
x -
y)) - 1)%
Z ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x -
y).
Proof.
Lemma double_round_minus_beta_ge_3_aux3 :
(3 <=
beta)%
Z ->
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_beta_ge_3_hyp fexp1 fexp2 ->
forall x y,
0 <
y ->
y <=
x ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x -
y).
Proof.
Lemma double_round_minus_beta_ge_3_aux :
(3 <=
beta)%
Z ->
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_beta_ge_3_hyp fexp1 fexp2 ->
forall x y,
0 <=
x -> 0 <=
y ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x -
y).
Proof.
Lemma double_round_plus_beta_ge_3 :
(3 <=
beta)%
Z ->
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_beta_ge_3_hyp fexp1 fexp2 ->
forall x y,
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x +
y).
Proof.
Lemma double_round_minus_beta_ge_3 :
(3 <=
beta)%
Z ->
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_plus_beta_ge_3_hyp fexp1 fexp2 ->
forall x y,
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x -
y).
Proof.
Section Double_round_plus_beta_ge_3_FLX.
Require Import Fcore_FLX.
Variable prec :
Z.
Variable prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FLX_double_round_plus_beta_ge_3_hyp :
(2 *
prec <=
prec')%
Z ->
double_round_plus_beta_ge_3_hyp (
FLX_exp prec) (
FLX_exp prec').
Proof.
Theorem double_round_plus_beta_ge_3_FLX :
(3 <=
beta)%
Z ->
forall choice1 choice2,
(2 *
prec <=
prec')%
Z ->
forall x y,
FLX_format beta prec x ->
FLX_format beta prec y ->
double_round_eq (
FLX_exp prec) (
FLX_exp prec')
choice1 choice2 (
x +
y).
Proof.
Theorem double_round_minus_beta_ge_3_FLX :
(3 <=
beta)%
Z ->
forall choice1 choice2,
(2 *
prec <=
prec')%
Z ->
forall x y,
FLX_format beta prec x ->
FLX_format beta prec y ->
double_round_eq (
FLX_exp prec) (
FLX_exp prec')
choice1 choice2 (
x -
y).
Proof.
End Double_round_plus_beta_ge_3_FLX.
Section Double_round_plus_beta_ge_3_FLT.
Require Import Fcore_FLX.
Require Import Fcore_FLT.
Variable emin prec :
Z.
Variable emin'
prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FLT_double_round_plus_beta_ge_3_hyp :
(
emin' <=
emin)%
Z -> (2 *
prec <=
prec')%
Z ->
double_round_plus_beta_ge_3_hyp (
FLT_exp emin prec) (
FLT_exp emin'
prec').
Proof.
Theorem double_round_plus_beta_ge_3_FLT :
(3 <=
beta)%
Z ->
forall choice1 choice2,
(
emin' <=
emin)%
Z -> (2 *
prec <=
prec')%
Z ->
forall x y,
FLT_format beta emin prec x ->
FLT_format beta emin prec y ->
double_round_eq (
FLT_exp emin prec) (
FLT_exp emin'
prec')
choice1 choice2 (
x +
y).
Proof.
Theorem double_round_minus_beta_ge_3_FLT :
(3 <=
beta)%
Z ->
forall choice1 choice2,
(
emin' <=
emin)%
Z -> (2 *
prec <=
prec')%
Z ->
forall x y,
FLT_format beta emin prec x ->
FLT_format beta emin prec y ->
double_round_eq (
FLT_exp emin prec) (
FLT_exp emin'
prec')
choice1 choice2 (
x -
y).
Proof.
End Double_round_plus_beta_ge_3_FLT.
Section Double_round_plus_beta_ge_3_FTZ.
Require Import Fcore_FLX.
Require Import Fcore_FTZ.
Variable emin prec :
Z.
Variable emin'
prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FTZ_double_round_plus_beta_ge_3_hyp :
(
emin' +
prec' <=
emin + 1)%
Z -> (2 *
prec <=
prec')%
Z ->
double_round_plus_beta_ge_3_hyp (
FTZ_exp emin prec) (
FTZ_exp emin'
prec').
Proof.
Theorem double_round_plus_beta_ge_3_FTZ :
(3 <=
beta)%
Z ->
forall choice1 choice2,
(
emin' +
prec' <=
emin + 1)%
Z -> (2 *
prec <=
prec')%
Z ->
forall x y,
FTZ_format beta emin prec x ->
FTZ_format beta emin prec y ->
double_round_eq (
FTZ_exp emin prec) (
FTZ_exp emin'
prec')
choice1 choice2 (
x +
y).
Proof.
Theorem double_round_minus_beta_ge_3_FTZ :
(3 <=
beta)%
Z ->
forall choice1 choice2,
(
emin' +
prec' <=
emin + 1)%
Z -> (2 *
prec <=
prec')%
Z ->
forall x y,
FTZ_format beta emin prec x ->
FTZ_format beta emin prec y ->
double_round_eq (
FTZ_exp emin prec) (
FTZ_exp emin'
prec')
choice1 choice2 (
x -
y).
Proof.
End Double_round_plus_beta_ge_3_FTZ.
End Double_round_plus_beta_ge_3.
End Double_round_plus.
Lemma double_round_mid_cases :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
forall x,
0 <
x ->
(
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x) - 1)%
Z ->
(
fexp1 (
ln_beta x) <=
ln_beta x)%
Z ->
(
Rabs (
x -
midp fexp1 x) <= / 2 * (
ulp beta fexp2 x) ->
double_round_eq fexp1 fexp2 choice1 choice2 x) ->
double_round_eq fexp1 fexp2 choice1 choice2 x.
Proof.
Section Double_round_sqrt.
Definition double_round_sqrt_hyp fexp1 fexp2 :=
(
forall ex, (2 *
fexp1 ex <=
fexp1 (2 *
ex))%
Z)
/\ (
forall ex, (2 *
fexp1 ex <=
fexp1 (2 *
ex - 1))%
Z)
/\ (
forall ex, (
fexp1 (2 *
ex) < 2 *
ex)%
Z ->
(
fexp2 ex +
ex <= 2 *
fexp1 ex - 2)%
Z).
Lemma ln_beta_sqrt_disj :
forall x,
0 <
x ->
(
ln_beta x = 2 *
ln_beta (
sqrt x) - 1 :>
Z)%
Z
\/ (
ln_beta x = 2 *
ln_beta (
sqrt x) :>
Z)%
Z.
Proof.
Lemma double_round_sqrt_aux :
forall fexp1 fexp2 :
Z ->
Z,
Valid_exp fexp1 ->
Valid_exp fexp2 ->
double_round_sqrt_hyp fexp1 fexp2 ->
forall x,
0 <
x ->
(
fexp2 (
ln_beta (
sqrt x)) <=
fexp1 (
ln_beta (
sqrt x)) - 1)%
Z ->
generic_format beta fexp1 x ->
/ 2 *
ulp beta fexp2 (
sqrt x) <
Rabs (
sqrt x -
midp fexp1 (
sqrt x)).
Proof.
intros fexp1 fexp2 Vfexp1 Vfexp2 Hexp x Px Hf2 Fx.
assert (
Hbeta : (2 <=
beta)%
Z).
{
destruct beta as (
beta_val,
beta_prop).
now apply Zle_bool_imp_le. }
set (
a :=
round beta fexp1 Zfloor (
sqrt x)).
set (
u1 :=
bpow (
fexp1 (
ln_beta (
sqrt x)))).
set (
u2 :=
bpow (
fexp2 (
ln_beta (
sqrt x)))).
set (
b := / 2 * (
u1 -
u2)).
set (
b' := / 2 * (
u1 +
u2)).
unfold midp;
rewrite 2!
ulp_neq_0;
try now apply Rgt_not_eq,
sqrt_lt_R0.
apply Rnot_ge_lt;
intro H;
apply Rge_le in H.
assert (
Fa :
generic_format beta fexp1 a).
{
unfold a.
apply generic_format_round.
-
exact Vfexp1.
-
now apply valid_rnd_DN. }
revert Fa;
revert Fx.
unfold generic_format,
F2R,
scaled_mantissa,
canonic_exp;
simpl.
set (
mx :=
Ztrunc (
x *
bpow (-
fexp1 (
ln_beta x)))).
set (
ma :=
Ztrunc (
a *
bpow (-
fexp1 (
ln_beta a)))).
intros Fx Fa.
assert (
Nna : 0 <=
a).
{
rewrite <- (
round_0 beta fexp1 Zfloor).
unfold a;
apply round_le.
-
exact Vfexp1.
-
now apply valid_rnd_DN.
-
apply sqrt_pos. }
assert (
Phu1 : 0 < / 2 *
u1).
{
apply Rmult_lt_0_compat; [
lra|
apply bpow_gt_0]. }
assert (
Phu2 : 0 < / 2 *
u2).
{
apply Rmult_lt_0_compat; [
lra|
apply bpow_gt_0]. }
assert (
Pb : 0 <
b).
{
unfold b.
rewrite <- (
Rmult_0_r (/ 2)).
apply Rmult_lt_compat_l; [
lra|].
apply Rlt_Rminus.
unfold u2,
u1.
apply bpow_lt.
omega. }
assert (
Pb' : 0 <
b').
{
now unfold b';
rewrite Rmult_plus_distr_l;
apply Rplus_lt_0_compat. }
assert (
Hr :
sqrt x <=
a +
b').
{
unfold b';
apply (
Rplus_le_reg_r (- / 2 *
u1 -
a));
ring_simplify.
replace (
_ -
_)
with (
sqrt x - (
a + / 2 *
u1))
by ring.
now apply Rabs_le_inv. }
assert (
Hl :
a +
b <=
sqrt x).
{
unfold b;
apply (
Rplus_le_reg_r (- / 2 *
u1 -
a));
ring_simplify.
replace (
_ +
sqrt _)
with (
sqrt x - (
a + / 2 *
u1))
by ring.
rewrite Ropp_mult_distr_l_reverse.
now apply Rabs_le_inv in H;
destruct H. }
assert (
Hf1 : (2 *
fexp1 (
ln_beta (
sqrt x)) <=
fexp1 (
ln_beta (
x)))%
Z);
[
destruct (
ln_beta_sqrt_disj x Px)
as [
H'|
H'];
rewrite H';
apply Hexp|].
assert (
Hlx : (
fexp1 (2 *
ln_beta (
sqrt x)) < 2 *
ln_beta (
sqrt x))%
Z).
{
destruct (
ln_beta_sqrt_disj x Px)
as [
Hlx|
Hlx].
-
apply (
valid_exp_large fexp1 (
ln_beta x)); [|
omega].
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|].
-
rewrite <-
Hlx.
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|]. }
assert (
Hsl :
a *
a +
u1 *
a -
u2 *
a +
b *
b <=
x).
{
replace (
_ +
_)
with ((
a +
b) * (
a +
b)); [|
now unfold b;
field].
rewrite <-
sqrt_def; [|
now apply Rlt_le].
assert (
H' : 0 <=
a +
b); [
now apply Rlt_le,
Rplus_le_lt_0_compat|].
now apply Rmult_le_compat. }
assert (
Hsr :
x <=
a *
a +
u1 *
a +
u2 *
a +
b' *
b').
{
replace (
_ +
_)
with ((
a +
b') * (
a +
b')); [|
now unfold b';
field].
rewrite <- (
sqrt_def x); [|
now apply Rlt_le].
assert (
H' : 0 <=
sqrt x); [
now apply sqrt_pos|].
now apply Rmult_le_compat. }
destruct (
Req_dec a 0)
as [
Za|
Nza].
-
apply (
Rlt_irrefl 0).
apply Rlt_le_trans with (
b *
b); [
now apply Rmult_lt_0_compat|].
apply Rle_trans with x.
+
revert Hsl;
unfold Rminus;
rewrite Za;
do 3
rewrite Rmult_0_r.
now rewrite Ropp_0;
do 3
rewrite Rplus_0_l.
+
rewrite Fx.
apply (
Rmult_le_reg_r (
bpow (-
fexp1 (
ln_beta x))));
[
now apply bpow_gt_0|].
rewrite Rmult_0_l;
bpow_simplify.
unfold mx.
rewrite Ztrunc_floor;
[|
now apply Rmult_le_pos; [
apply Rlt_le|
apply bpow_ge_0]].
apply Req_le.
change 0
with (
Z2R 0);
apply f_equal.
apply Zfloor_imp.
split; [
now apply Rmult_le_pos; [
apply Rlt_le|
apply bpow_ge_0]|
simpl].
apply (
Rmult_lt_reg_r (
bpow (
fexp1 (
ln_beta x)))); [
now apply bpow_gt_0|].
rewrite Rmult_1_l;
bpow_simplify.
apply Rlt_le_trans with (
bpow (2 *
fexp1 (
ln_beta (
sqrt x))));
[|
now apply bpow_le].
change 2%
Z with (1 + 1)%
Z;
rewrite Zmult_plus_distr_l;
rewrite Zmult_1_l.
rewrite bpow_plus.
rewrite <- (
sqrt_def x)
at 1; [|
now apply Rlt_le].
assert (
sqrt x <
bpow (
fexp1 (
ln_beta (
sqrt x))));
[|
now apply Rmult_lt_compat; [
apply sqrt_pos|
apply sqrt_pos| |]].
apply (
Rle_lt_trans _ _ _ Hr);
rewrite Za;
rewrite Rplus_0_l.
unfold b';
change (
bpow _)
with u1.
apply Rlt_le_trans with (/ 2 * (
u1 +
u1)); [|
lra].
apply Rmult_lt_compat_l; [
lra|];
apply Rplus_lt_compat_l.
unfold u2,
u1,
ulp,
canonic_exp;
apply bpow_lt;
omega.
-
assert (
Pa : 0 <
a); [
lra|].
assert (
Hla : (
ln_beta a =
ln_beta (
sqrt x) :>
Z)).
{
unfold a;
apply ln_beta_DN.
-
exact Vfexp1.
-
now fold a. }
assert (
Hl' : 0 < - (
u2 *
a) +
b *
b).
{
apply (
Rplus_lt_reg_r (
u2 *
a));
ring_simplify.
unfold b;
ring_simplify.
apply (
Rplus_lt_reg_r (/ 2 *
u2 *
u1));
field_simplify.
replace (
_ / 2)
with (
u2 * (
a + / 2 *
u1))
by field.
replace (
_ / 8)
with (/ 4 * (
u2 ^ 2 +
u1 ^ 2))
by field.
apply Rlt_le_trans with (
u2 *
bpow (
ln_beta (
sqrt x))).
-
apply Rmult_lt_compat_l; [
now unfold u2,
ulp;
apply bpow_gt_0|].
unfold u1;
rewrite <-
Hla.
apply Rlt_le_trans with (
a +
bpow (
fexp1 (
ln_beta a))).
+
apply Rplus_lt_compat_l.
rewrite <- (
Rmult_1_l (
bpow _))
at 2.
apply Rmult_lt_compat_r; [
apply bpow_gt_0|
lra].
+
apply Rle_trans with (
a+
ulp beta fexp1 a).
right;
now rewrite ulp_neq_0.
apply (
id_p_ulp_le_bpow _ _ _ _ Pa Fa).
apply Rabs_lt_inv,
bpow_ln_beta_gt.
-
apply Rle_trans with (
bpow (- 2) *
u1 ^ 2).
+
unfold pow;
rewrite Rmult_1_r.
unfold u1,
u2,
ulp,
canonic_exp;
bpow_simplify;
apply bpow_le.
now apply Hexp.
+
apply Rmult_le_compat.
*
apply bpow_ge_0.
*
apply pow2_ge_0.
*
unfold Fcore_Raux.bpow,
Z.pow_pos;
simpl;
rewrite Zmult_1_r.
apply Rinv_le; [
lra|].
change 4
with (
Z2R (2 * 2)%
Z);
apply Z2R_le,
Zmult_le_compat;
omega.
*
rewrite <- (
Rplus_0_l (
u1 ^ 2))
at 1;
apply Rplus_le_compat_r.
apply pow2_ge_0. }
assert (
Hr' :
x <=
a *
a +
u1 *
a).
{
rewrite Hla in Fa.
rewrite <-
Rmult_plus_distr_r.
unfold u1,
ulp,
canonic_exp.
rewrite <- (
Rmult_1_l (
bpow _));
rewrite Fa;
rewrite <-
Rmult_plus_distr_r.
rewrite <-
Rmult_assoc;
rewrite (
Rmult_comm _ (
Z2R ma)).
rewrite <- (
Rmult_assoc (
Z2R ma));
bpow_simplify.
apply (
Rmult_le_reg_r (
bpow (- 2 *
fexp1 (
ln_beta (
sqrt x)))));
[
now apply bpow_gt_0|
bpow_simplify].
rewrite Fx at 1;
bpow_simplify.
rewrite <-
Z2R_Zpower; [|
omega].
change 1
with (
Z2R 1);
rewrite <-
Z2R_plus;
do 2
rewrite <-
Z2R_mult.
apply Z2R_le,
Zlt_succ_le,
lt_Z2R.
unfold Z.succ;
rewrite Z2R_plus;
do 2
rewrite Z2R_mult;
rewrite Z2R_plus.
rewrite Z2R_Zpower; [|
omega].
apply (
Rmult_lt_reg_r (
bpow (2 *
fexp1 (
ln_beta (
sqrt x)))));
[
now apply bpow_gt_0|
bpow_simplify].
rewrite <-
Fx.
change 2%
Z with (1 + 1)%
Z;
rewrite Zmult_plus_distr_l;
rewrite Zmult_1_l.
rewrite bpow_plus;
simpl.
replace (
_ *
_)
with (
a *
a +
u1 *
a +
u1 *
u1);
[|
unfold u1,
ulp,
canonic_exp;
rewrite Fa;
ring].
apply (
Rle_lt_trans _ _ _ Hsr).
rewrite Rplus_assoc;
apply Rplus_lt_compat_l.
apply (
Rplus_lt_reg_r (-
b' *
b' + / 2 *
u1 *
u2));
ring_simplify.
replace (
_ +
_)
with ((
a + / 2 *
u1) *
u2)
by ring.
apply Rlt_le_trans with (
bpow (
ln_beta (
sqrt x)) *
u2).
-
apply Rmult_lt_compat_r; [
now unfold u2,
ulp;
apply bpow_gt_0|].
apply Rlt_le_trans with (
a +
u1); [
lra|].
unfold u1;
fold (
canonic_exp beta fexp1 (
sqrt x)).
rewrite <-
canonic_exp_DN; [|
exact Vfexp1|
exact Pa];
fold a.
rewrite <-
ulp_neq_0;
trivial.
apply id_p_ulp_le_bpow.
+
exact Pa.
+
now apply round_DN_pt.
+
apply Rle_lt_trans with (
sqrt x).
*
now apply round_DN_pt.
*
apply Rabs_lt_inv.
apply bpow_ln_beta_gt.
-
apply Rle_trans with (/ 2 *
u1 ^ 2).
+
apply Rle_trans with (
bpow (- 2) *
u1 ^ 2).
*
unfold pow;
rewrite Rmult_1_r.
unfold u2,
u1,
ulp,
canonic_exp.
bpow_simplify.
apply bpow_le.
rewrite Zplus_comm.
now apply Hexp.
*
apply Rmult_le_compat_r; [
now apply pow2_ge_0|].
unfold Fcore_Raux.bpow;
simpl;
unfold Z.pow_pos;
simpl.
rewrite Zmult_1_r.
apply Rinv_le; [
lra|].
change 2
with (
Z2R 2);
apply Z2R_le.
rewrite <- (
Zmult_1_l 2).
apply Zmult_le_compat;
omega.
+
assert (
u2 ^ 2 <
u1 ^ 2); [|
unfold b';
lra].
unfold pow;
do 2
rewrite Rmult_1_r.
assert (
H' : 0 <=
u2); [
unfold u2,
ulp;
apply bpow_ge_0|].
assert (
u2 <
u1); [|
now apply Rmult_lt_compat].
unfold u1,
u2,
ulp,
canonic_exp;
apply bpow_lt;
omega. }
apply (
Rlt_irrefl (
a *
a +
u1 *
a)).
apply Rlt_le_trans with (
a *
a +
u1 *
a -
u2 *
a +
b *
b).
+
rewrite <- (
Rplus_0_r (
a *
a +
_))
at 1.
unfold Rminus;
rewrite (
Rplus_assoc _ _ (
b *
b)).
now apply Rplus_lt_compat_l.
+
now apply Rle_trans with x.
Qed.
Lemma double_round_sqrt :
forall fexp1 fexp2 :
Z ->
Z,
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_sqrt_hyp fexp1 fexp2 ->
forall x,
generic_format beta fexp1 x ->
double_round_eq fexp1 fexp2 choice1 choice2 (
sqrt x).
Proof.
Section Double_round_sqrt_FLX.
Require Import Fcore_FLX.
Variable prec :
Z.
Variable prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FLX_double_round_sqrt_hyp :
(2 *
prec + 2 <=
prec')%
Z ->
double_round_sqrt_hyp (
FLX_exp prec) (
FLX_exp prec').
Proof.
Theorem double_round_sqrt_FLX :
forall choice1 choice2,
(2 *
prec + 2 <=
prec')%
Z ->
forall x,
FLX_format beta prec x ->
double_round_eq (
FLX_exp prec) (
FLX_exp prec')
choice1 choice2 (
sqrt x).
Proof.
End Double_round_sqrt_FLX.
Section Double_round_sqrt_FLT.
Require Import Fcore_FLX.
Require Import Fcore_FLT.
Variable emin prec :
Z.
Variable emin'
prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FLT_double_round_sqrt_hyp :
(
emin <= 0)%
Z ->
((
emin' <=
emin -
prec - 2)%
Z
\/ (2 *
emin' <=
emin - 4 *
prec - 2)%
Z) ->
(2 *
prec + 2 <=
prec')%
Z ->
double_round_sqrt_hyp (
FLT_exp emin prec) (
FLT_exp emin'
prec').
Proof.
Theorem double_round_sqrt_FLT :
forall choice1 choice2,
(
emin <= 0)%
Z ->
((
emin' <=
emin -
prec - 2)%
Z
\/ (2 *
emin' <=
emin - 4 *
prec - 2)%
Z) ->
(2 *
prec + 2 <=
prec')%
Z ->
forall x,
FLT_format beta emin prec x ->
double_round_eq (
FLT_exp emin prec) (
FLT_exp emin'
prec')
choice1 choice2 (
sqrt x).
Proof.
End Double_round_sqrt_FLT.
Section Double_round_sqrt_FTZ.
Require Import Fcore_FLX.
Require Import Fcore_FTZ.
Variable emin prec :
Z.
Variable emin'
prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FTZ_double_round_sqrt_hyp :
(2 * (
emin' +
prec') <=
emin +
prec <= 1)%
Z ->
(2 *
prec + 2 <=
prec')%
Z ->
double_round_sqrt_hyp (
FTZ_exp emin prec) (
FTZ_exp emin'
prec').
Proof.
Theorem double_round_sqrt_FTZ :
(4 <=
beta)%
Z ->
forall choice1 choice2,
(2 * (
emin' +
prec') <=
emin +
prec <= 1)%
Z ->
(2 *
prec + 2 <=
prec')%
Z ->
forall x,
FTZ_format beta emin prec x ->
double_round_eq (
FTZ_exp emin prec) (
FTZ_exp emin'
prec')
choice1 choice2 (
sqrt x).
Proof.
End Double_round_sqrt_FTZ.
Section Double_round_sqrt_beta_ge_4.
Definition double_round_sqrt_beta_ge_4_hyp fexp1 fexp2 :=
(
forall ex, (2 *
fexp1 ex <=
fexp1 (2 *
ex))%
Z)
/\ (
forall ex, (2 *
fexp1 ex <=
fexp1 (2 *
ex - 1))%
Z)
/\ (
forall ex, (
fexp1 (2 *
ex) < 2 *
ex)%
Z ->
(
fexp2 ex +
ex <= 2 *
fexp1 ex - 1)%
Z).
Lemma double_round_sqrt_beta_ge_4_aux :
(4 <=
beta)%
Z ->
forall fexp1 fexp2 :
Z ->
Z,
Valid_exp fexp1 ->
Valid_exp fexp2 ->
double_round_sqrt_beta_ge_4_hyp fexp1 fexp2 ->
forall x,
0 <
x ->
(
fexp2 (
ln_beta (
sqrt x)) <=
fexp1 (
ln_beta (
sqrt x)) - 1)%
Z ->
generic_format beta fexp1 x ->
/ 2 *
ulp beta fexp2 (
sqrt x) <
Rabs (
sqrt x -
midp fexp1 (
sqrt x)).
Proof.
intros Hbeta fexp1 fexp2 Vfexp1 Vfexp2 Hexp x Px Hf2 Fx.
set (
a :=
round beta fexp1 Zfloor (
sqrt x)).
set (
u1 :=
bpow (
fexp1 (
ln_beta (
sqrt x)))).
set (
u2 :=
bpow (
fexp2 (
ln_beta (
sqrt x)))).
set (
b := / 2 * (
u1 -
u2)).
set (
b' := / 2 * (
u1 +
u2)).
unfold midp;
rewrite 2!
ulp_neq_0;
try now apply Rgt_not_eq,
sqrt_lt_R0.
apply Rnot_ge_lt;
intro H;
apply Rge_le in H.
assert (
Fa :
generic_format beta fexp1 a).
{
unfold a.
apply generic_format_round.
-
exact Vfexp1.
-
now apply valid_rnd_DN. }
revert Fa;
revert Fx.
unfold generic_format,
F2R,
scaled_mantissa,
canonic_exp;
simpl.
set (
mx :=
Ztrunc (
x *
bpow (-
fexp1 (
ln_beta x)))).
set (
ma :=
Ztrunc (
a *
bpow (-
fexp1 (
ln_beta a)))).
intros Fx Fa.
assert (
Nna : 0 <=
a).
{
rewrite <- (
round_0 beta fexp1 Zfloor).
unfold a;
apply round_le.
-
exact Vfexp1.
-
now apply valid_rnd_DN.
-
apply sqrt_pos. }
assert (
Phu1 : 0 < / 2 *
u1).
{
apply Rmult_lt_0_compat; [
lra|
apply bpow_gt_0]. }
assert (
Phu2 : 0 < / 2 *
u2).
{
apply Rmult_lt_0_compat; [
lra|
apply bpow_gt_0]. }
assert (
Pb : 0 <
b).
{
unfold b.
rewrite <- (
Rmult_0_r (/ 2)).
apply Rmult_lt_compat_l; [
lra|].
apply Rlt_Rminus.
unfold u2,
u1,
ulp,
canonic_exp.
apply bpow_lt.
omega. }
assert (
Pb' : 0 <
b').
{
now unfold b';
rewrite Rmult_plus_distr_l;
apply Rplus_lt_0_compat. }
assert (
Hr :
sqrt x <=
a +
b').
{
unfold b';
apply (
Rplus_le_reg_r (- / 2 *
u1 -
a));
ring_simplify.
replace (
_ -
_)
with (
sqrt x - (
a + / 2 *
u1))
by ring.
now apply Rabs_le_inv. }
assert (
Hl :
a +
b <=
sqrt x).
{
unfold b;
apply (
Rplus_le_reg_r (- / 2 *
u1 -
a));
ring_simplify.
replace (
_ +
sqrt _)
with (
sqrt x - (
a + / 2 *
u1))
by ring.
rewrite Ropp_mult_distr_l_reverse.
now apply Rabs_le_inv in H;
destruct H. }
assert (
Hf1 : (2 *
fexp1 (
ln_beta (
sqrt x)) <=
fexp1 (
ln_beta (
x)))%
Z);
[
destruct (
ln_beta_sqrt_disj x Px)
as [
H'|
H'];
rewrite H';
apply Hexp|].
assert (
Hlx : (
fexp1 (2 *
ln_beta (
sqrt x)) < 2 *
ln_beta (
sqrt x))%
Z).
{
destruct (
ln_beta_sqrt_disj x Px)
as [
Hlx|
Hlx].
-
apply (
valid_exp_large fexp1 (
ln_beta x)); [|
omega].
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|].
-
rewrite <-
Hlx.
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|]. }
assert (
Hsl :
a *
a +
u1 *
a -
u2 *
a +
b *
b <=
x).
{
replace (
_ +
_)
with ((
a +
b) * (
a +
b)); [|
now unfold b;
field].
rewrite <-
sqrt_def; [|
now apply Rlt_le].
assert (
H' : 0 <=
a +
b); [
now apply Rlt_le,
Rplus_le_lt_0_compat|].
now apply Rmult_le_compat. }
assert (
Hsr :
x <=
a *
a +
u1 *
a +
u2 *
a +
b' *
b').
{
replace (
_ +
_)
with ((
a +
b') * (
a +
b')); [|
now unfold b';
field].
rewrite <- (
sqrt_def x); [|
now apply Rlt_le].
assert (
H' : 0 <=
sqrt x); [
now apply sqrt_pos|].
now apply Rmult_le_compat. }
destruct (
Req_dec a 0)
as [
Za|
Nza].
-
apply (
Rlt_irrefl 0).
apply Rlt_le_trans with (
b *
b); [
now apply Rmult_lt_0_compat|].
apply Rle_trans with x.
+
revert Hsl;
unfold Rminus;
rewrite Za;
do 3
rewrite Rmult_0_r.
now rewrite Ropp_0;
do 3
rewrite Rplus_0_l.
+
rewrite Fx.
apply (
Rmult_le_reg_r (
bpow (-
fexp1 (
ln_beta x))));
[
now apply bpow_gt_0|].
rewrite Rmult_0_l;
bpow_simplify.
unfold mx.
rewrite Ztrunc_floor;
[|
now apply Rmult_le_pos; [
apply Rlt_le|
apply bpow_ge_0]].
apply Req_le.
change 0
with (
Z2R 0);
apply f_equal.
apply Zfloor_imp.
split; [
now apply Rmult_le_pos; [
apply Rlt_le|
apply bpow_ge_0]|
simpl].
apply (
Rmult_lt_reg_r (
bpow (
fexp1 (
ln_beta x)))); [
now apply bpow_gt_0|].
rewrite Rmult_1_l;
bpow_simplify.
apply Rlt_le_trans with (
bpow (2 *
fexp1 (
ln_beta (
sqrt x))));
[|
now apply bpow_le].
change 2%
Z with (1 + 1)%
Z;
rewrite Zmult_plus_distr_l;
rewrite Zmult_1_l.
rewrite bpow_plus.
rewrite <- (
sqrt_def x)
at 1; [|
now apply Rlt_le].
assert (
sqrt x <
bpow (
fexp1 (
ln_beta (
sqrt x))));
[|
now apply Rmult_lt_compat; [
apply sqrt_pos|
apply sqrt_pos| |]].
apply (
Rle_lt_trans _ _ _ Hr);
rewrite Za;
rewrite Rplus_0_l.
unfold b';
change (
bpow _)
with u1.
apply Rlt_le_trans with (/ 2 * (
u1 +
u1)); [|
lra].
apply Rmult_lt_compat_l; [
lra|];
apply Rplus_lt_compat_l.
unfold u2,
u1,
ulp,
canonic_exp;
apply bpow_lt;
omega.
-
assert (
Pa : 0 <
a); [
lra|].
assert (
Hla : (
ln_beta a =
ln_beta (
sqrt x) :>
Z)).
{
unfold a;
apply ln_beta_DN.
-
exact Vfexp1.
-
now fold a. }
assert (
Hl' : 0 < - (
u2 *
a) +
b *
b).
{
apply (
Rplus_lt_reg_r (
u2 *
a));
ring_simplify.
unfold b;
ring_simplify.
apply (
Rplus_lt_reg_r (/ 2 *
u2 *
u1));
field_simplify.
replace (
_ / 2)
with (
u2 * (
a + / 2 *
u1))
by field.
replace (
_ / 8)
with (/ 4 * (
u2 ^ 2 +
u1 ^ 2))
by field.
apply Rlt_le_trans with (
u2 *
bpow (
ln_beta (
sqrt x))).
-
apply Rmult_lt_compat_l; [
now unfold u2,
ulp;
apply bpow_gt_0|].
unfold u1;
rewrite <-
Hla.
apply Rlt_le_trans with (
a +
ulp beta fexp1 a).
+
apply Rplus_lt_compat_l.
rewrite <- (
Rmult_1_l (
ulp _ _ _)).
rewrite ulp_neq_0;
trivial.
apply Rmult_lt_compat_r; [
apply bpow_gt_0|
lra].
+
apply (
id_p_ulp_le_bpow _ _ _ _ Pa Fa).
apply Rabs_lt_inv,
bpow_ln_beta_gt.
-
apply Rle_trans with (
bpow (- 1) *
u1 ^ 2).
+
unfold pow;
rewrite Rmult_1_r.
unfold u1,
u2,
ulp,
canonic_exp;
bpow_simplify;
apply bpow_le.
now apply Hexp.
+
apply Rmult_le_compat.
*
apply bpow_ge_0.
*
apply pow2_ge_0.
*
unfold Fcore_Raux.bpow,
Z.pow_pos;
simpl;
rewrite Zmult_1_r.
apply Rinv_le; [
lra|].
now change 4
with (
Z2R 4);
apply Z2R_le.
*
rewrite <- (
Rplus_0_l (
u1 ^ 2))
at 1;
apply Rplus_le_compat_r.
apply pow2_ge_0. }
assert (
Hr' :
x <=
a *
a +
u1 *
a).
{
rewrite Hla in Fa.
rewrite <-
Rmult_plus_distr_r.
unfold u1,
ulp,
canonic_exp.
rewrite <- (
Rmult_1_l (
bpow _));
rewrite Fa;
rewrite <-
Rmult_plus_distr_r.
rewrite <-
Rmult_assoc;
rewrite (
Rmult_comm _ (
Z2R ma)).
rewrite <- (
Rmult_assoc (
Z2R ma));
bpow_simplify.
apply (
Rmult_le_reg_r (
bpow (- 2 *
fexp1 (
ln_beta (
sqrt x)))));
[
now apply bpow_gt_0|
bpow_simplify].
rewrite Fx at 1;
bpow_simplify.
rewrite <-
Z2R_Zpower; [|
omega].
change 1
with (
Z2R 1);
rewrite <-
Z2R_plus;
do 2
rewrite <-
Z2R_mult.
apply Z2R_le,
Zlt_succ_le,
lt_Z2R.
unfold Z.succ;
rewrite Z2R_plus;
do 2
rewrite Z2R_mult;
rewrite Z2R_plus.
rewrite Z2R_Zpower; [|
omega].
apply (
Rmult_lt_reg_r (
bpow (2 *
fexp1 (
ln_beta (
sqrt x)))));
[
now apply bpow_gt_0|
bpow_simplify].
rewrite <-
Fx.
change 2%
Z with (1 + 1)%
Z;
rewrite Zmult_plus_distr_l;
rewrite Zmult_1_l.
rewrite bpow_plus;
simpl.
replace (
_ *
_)
with (
a *
a +
u1 *
a +
u1 *
u1);
[|
unfold u1,
ulp,
canonic_exp;
rewrite Fa;
ring].
apply (
Rle_lt_trans _ _ _ Hsr).
rewrite Rplus_assoc;
apply Rplus_lt_compat_l.
apply (
Rplus_lt_reg_r (-
b' *
b' + / 2 *
u1 *
u2));
ring_simplify.
replace (
_ +
_)
with ((
a + / 2 *
u1) *
u2)
by ring.
apply Rlt_le_trans with (
bpow (
ln_beta (
sqrt x)) *
u2).
-
apply Rmult_lt_compat_r; [
now unfold u2,
ulp;
apply bpow_gt_0|].
apply Rlt_le_trans with (
a +
u1); [
lra|].
unfold u1;
fold (
canonic_exp beta fexp1 (
sqrt x)).
rewrite <-
canonic_exp_DN; [|
exact Vfexp1|
exact Pa];
fold a.
rewrite <-
ulp_neq_0;
trivial.
apply id_p_ulp_le_bpow.
+
exact Pa.
+
now apply round_DN_pt.
+
apply Rle_lt_trans with (
sqrt x).
*
now apply round_DN_pt.
*
apply Rabs_lt_inv.
apply bpow_ln_beta_gt.
-
apply Rle_trans with (/ 2 *
u1 ^ 2).
+
apply Rle_trans with (
bpow (- 1) *
u1 ^ 2).
*
unfold pow;
rewrite Rmult_1_r.
unfold u2,
u1,
ulp,
canonic_exp.
bpow_simplify.
apply bpow_le.
rewrite Zplus_comm.
now apply Hexp.
*
apply Rmult_le_compat_r; [
now apply pow2_ge_0|].
unfold Fcore_Raux.bpow;
simpl;
unfold Z.pow_pos;
simpl.
rewrite Zmult_1_r.
apply Rinv_le; [
lra|].
change 2
with (
Z2R 2);
apply Z2R_le;
omega.
+
assert (
u2 ^ 2 <
u1 ^ 2); [|
unfold b';
lra].
unfold pow;
do 2
rewrite Rmult_1_r.
assert (
H' : 0 <=
u2); [
unfold u2,
ulp;
apply bpow_ge_0|].
assert (
u2 <
u1); [|
now apply Rmult_lt_compat].
unfold u1,
u2,
ulp,
canonic_exp;
apply bpow_lt;
omega. }
apply (
Rlt_irrefl (
a *
a +
u1 *
a)).
apply Rlt_le_trans with (
a *
a +
u1 *
a -
u2 *
a +
b *
b).
+
rewrite <- (
Rplus_0_r (
a *
a +
_))
at 1.
unfold Rminus;
rewrite (
Rplus_assoc _ _ (
b *
b)).
now apply Rplus_lt_compat_l.
+
now apply Rle_trans with x.
Qed.
Lemma double_round_sqrt_beta_ge_4 :
(4 <=
beta)%
Z ->
forall fexp1 fexp2 :
Z ->
Z,
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_sqrt_beta_ge_4_hyp fexp1 fexp2 ->
forall x,
generic_format beta fexp1 x ->
double_round_eq fexp1 fexp2 choice1 choice2 (
sqrt x).
Proof.
Section Double_round_sqrt_beta_ge_4_FLX.
Require Import Fcore_FLX.
Variable prec :
Z.
Variable prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FLX_double_round_sqrt_beta_ge_4_hyp :
(2 *
prec + 1 <=
prec')%
Z ->
double_round_sqrt_beta_ge_4_hyp (
FLX_exp prec) (
FLX_exp prec').
Proof.
Theorem double_round_sqrt_beta_ge_4_FLX :
(4 <=
beta)%
Z ->
forall choice1 choice2,
(2 *
prec + 1 <=
prec')%
Z ->
forall x,
FLX_format beta prec x ->
double_round_eq (
FLX_exp prec) (
FLX_exp prec')
choice1 choice2 (
sqrt x).
Proof.
End Double_round_sqrt_beta_ge_4_FLX.
Section Double_round_sqrt_beta_ge_4_FLT.
Require Import Fcore_FLX.
Require Import Fcore_FLT.
Variable emin prec :
Z.
Variable emin'
prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FLT_double_round_sqrt_beta_ge_4_hyp :
(
emin <= 0)%
Z ->
((
emin' <=
emin -
prec - 1)%
Z
\/ (2 *
emin' <=
emin - 4 *
prec)%
Z) ->
(2 *
prec + 1 <=
prec')%
Z ->
double_round_sqrt_beta_ge_4_hyp (
FLT_exp emin prec) (
FLT_exp emin'
prec').
Proof.
Theorem double_round_sqrt_beta_ge_4_FLT :
(4 <=
beta)%
Z ->
forall choice1 choice2,
(
emin <= 0)%
Z ->
((
emin' <=
emin -
prec - 1)%
Z
\/ (2 *
emin' <=
emin - 4 *
prec)%
Z) ->
(2 *
prec + 1 <=
prec')%
Z ->
forall x,
FLT_format beta emin prec x ->
double_round_eq (
FLT_exp emin prec) (
FLT_exp emin'
prec')
choice1 choice2 (
sqrt x).
Proof.
End Double_round_sqrt_beta_ge_4_FLT.
Section Double_round_sqrt_beta_ge_4_FTZ.
Require Import Fcore_FLX.
Require Import Fcore_FTZ.
Variable emin prec :
Z.
Variable emin'
prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FTZ_double_round_sqrt_beta_ge_4_hyp :
(2 * (
emin' +
prec') <=
emin +
prec <= 1)%
Z ->
(2 *
prec + 1 <=
prec')%
Z ->
double_round_sqrt_beta_ge_4_hyp (
FTZ_exp emin prec) (
FTZ_exp emin'
prec').
Proof.
Theorem double_round_sqrt_beta_ge_4_FTZ :
(4 <=
beta)%
Z ->
forall choice1 choice2,
(2 * (
emin' +
prec') <=
emin +
prec <= 1)%
Z ->
(2 *
prec + 1 <=
prec')%
Z ->
forall x,
FTZ_format beta emin prec x ->
double_round_eq (
FTZ_exp emin prec) (
FTZ_exp emin'
prec')
choice1 choice2 (
sqrt x).
Proof.
End Double_round_sqrt_beta_ge_4_FTZ.
End Double_round_sqrt_beta_ge_4.
End Double_round_sqrt.
Section Double_round_div.
Lemma double_round_eq_mid_beta_even :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
(
exists n, (
beta = 2 *
n :>
Z)%
Z) ->
forall x,
0 <
x ->
(
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x) - 1)%
Z ->
(
fexp1 (
ln_beta x) <=
ln_beta x)%
Z ->
x =
midp fexp1 x ->
double_round_eq fexp1 fexp2 choice1 choice2 x.
Proof.
Lemma double_round_really_zero :
forall (
fexp1 fexp2 :
Z ->
Z),
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
forall x,
0 <
x ->
(
ln_beta x <=
fexp1 (
ln_beta x) - 2)%
Z ->
double_round_eq fexp1 fexp2 choice1 choice2 x.
Proof.
Lemma double_round_zero :
forall fexp1 fexp2 :
Z ->
Z,
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
forall x,
0 <
x ->
(
fexp1 (
ln_beta x) =
ln_beta x + 1 :>
Z)%
Z ->
x <
bpow (
ln_beta x) - / 2 *
ulp beta fexp2 x ->
double_round_eq fexp1 fexp2 choice1 choice2 x.
Proof.
Lemma double_round_all_mid_cases :
forall fexp1 fexp2 :
Z ->
Z,
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
forall x,
0 <
x ->
(
fexp2 (
ln_beta x) <=
fexp1 (
ln_beta x) - 1)%
Z ->
((
fexp1 (
ln_beta x) =
ln_beta x + 1 :>
Z)%
Z ->
bpow (
ln_beta x) - / 2 *
ulp beta fexp2 x <=
x ->
double_round_eq fexp1 fexp2 choice1 choice2 x) ->
((
fexp1 (
ln_beta x) <=
ln_beta x)%
Z ->
midp fexp1 x - / 2 *
ulp beta fexp2 x <=
x <
midp fexp1 x ->
double_round_eq fexp1 fexp2 choice1 choice2 x) ->
((
fexp1 (
ln_beta x) <=
ln_beta x)%
Z ->
x =
midp fexp1 x ->
double_round_eq fexp1 fexp2 choice1 choice2 x) ->
((
fexp1 (
ln_beta x) <=
ln_beta x)%
Z ->
midp fexp1 x <
x <=
midp fexp1 x + / 2 *
ulp beta fexp2 x ->
double_round_eq fexp1 fexp2 choice1 choice2 x) ->
double_round_eq fexp1 fexp2 choice1 choice2 x.
Proof.
Lemma ln_beta_div_disj :
forall x y :
R,
0 <
x -> 0 <
y ->
((
ln_beta (
x /
y) =
ln_beta x -
ln_beta y :>
Z)%
Z
\/ (
ln_beta (
x /
y) =
ln_beta x -
ln_beta y + 1 :>
Z)%
Z).
Proof.
Definition double_round_div_hyp fexp1 fexp2 :=
(
forall ex, (
fexp2 ex <=
fexp1 ex - 1)%
Z)
/\ (
forall ex ey, (
fexp1 ex <
ex)%
Z -> (
fexp1 ey <
ey)%
Z ->
(
fexp1 (
ex -
ey) <=
ex -
ey + 1)%
Z ->
(
fexp2 (
ex -
ey) <=
fexp1 ex -
ey)%
Z)
/\ (
forall ex ey, (
fexp1 ex <
ex)%
Z -> (
fexp1 ey <
ey)%
Z ->
(
fexp1 (
ex -
ey + 1) <=
ex -
ey + 1 + 1)%
Z ->
(
fexp2 (
ex -
ey + 1) <=
fexp1 ex -
ey)%
Z)
/\ (
forall ex ey, (
fexp1 ex <
ex)%
Z -> (
fexp1 ey <
ey)%
Z ->
(
fexp1 (
ex -
ey) <=
ex -
ey)%
Z ->
(
fexp2 (
ex -
ey) <=
fexp1 (
ex -
ey)
+
fexp1 ey -
ey)%
Z)
/\ (
forall ex ey, (
fexp1 ex <
ex)%
Z -> (
fexp1 ey <
ey)%
Z ->
(
fexp1 (
ex -
ey) =
ex -
ey + 1)%
Z ->
(
fexp2 (
ex -
ey) <=
ex -
ey -
ey +
fexp1 ey)%
Z).
Lemma double_round_div_aux0 :
forall fexp1 fexp2 :
Z ->
Z,
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_div_hyp fexp1 fexp2 ->
forall x y,
0 <
x -> 0 <
y ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
fexp1 (
ln_beta (
x /
y)) = (
ln_beta (
x /
y) + 1)%
Z ->
~ (
bpow (
ln_beta (
x /
y)) - / 2 *
ulp beta fexp2 (
x /
y) <=
x /
y).
Proof.
intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Px Py Fx Fy Hf1.
assert (
Hfx : (
fexp1 (
ln_beta x) <
ln_beta x)%
Z);
[
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|]|].
assert (
Hfy : (
fexp1 (
ln_beta y) <
ln_beta y)%
Z);
[
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|]|].
set (
p :=
bpow (
ln_beta (
x /
y))).
set (
u2 :=
bpow (
fexp2 (
ln_beta (
x /
y)))).
revert Fx Fy.
unfold generic_format,
F2R,
scaled_mantissa,
canonic_exp;
simpl.
set (
mx :=
Ztrunc (
x *
bpow (-
fexp1 (
ln_beta x)))).
set (
my :=
Ztrunc (
y *
bpow (-
fexp1 (
ln_beta y)))).
intros Fx Fy.
rewrite ulp_neq_0.
2:
apply Rmult_integral_contrapositive_currified; [
now apply Rgt_not_eq|
idtac].
2:
now apply Rinv_neq_0_compat,
Rgt_not_eq.
intro Hl.
assert (
Hr :
x /
y <
p);
[
now apply Rabs_lt_inv;
apply bpow_ln_beta_gt|].
apply (
Rlt_irrefl (
p - / 2 *
u2)).
apply (
Rle_lt_trans _ _ _ Hl).
apply (
Rmult_lt_reg_r y _ _ Py).
unfold Rdiv;
rewrite Rmult_assoc.
rewrite Rinv_l; [|
now apply Rgt_not_eq];
rewrite Rmult_1_r.
destruct (
Zle_or_lt Z0 (
fexp1 (
ln_beta x) -
ln_beta (
x /
y)
-
fexp1 (
ln_beta y))%
Z)
as [
He|
He].
-
apply Rle_lt_trans with (
p *
y -
p *
bpow (
fexp1 (
ln_beta y))).
+
rewrite Fx;
rewrite Fy at 1.
rewrite <-
Rmult_assoc.
rewrite (
Rmult_comm p).
unfold p;
bpow_simplify.
apply (
Rmult_le_reg_r (
bpow (-
ln_beta (
x /
y) -
fexp1 (
ln_beta y))));
[
now apply bpow_gt_0|].
rewrite Rmult_minus_distr_r.
bpow_simplify.
rewrite <-
Z2R_Zpower; [|
exact He].
rewrite <-
Z2R_mult.
change 1
with (
Z2R 1);
rewrite <-
Z2R_minus.
apply Z2R_le.
apply (
Zplus_le_reg_r _ _ 1);
ring_simplify.
apply Zlt_le_succ.
apply lt_Z2R.
rewrite Z2R_mult.
rewrite Z2R_Zpower; [|
exact He].
apply (
Rmult_lt_reg_r (
bpow (
fexp1 (
ln_beta y) +
ln_beta (
x /
y))));
[
now apply bpow_gt_0|].
bpow_simplify.
rewrite <-
Fx.
rewrite bpow_plus.
rewrite <-
Rmult_assoc;
rewrite <-
Fy.
fold p.
apply (
Rmult_lt_reg_r (/
y)); [
now apply Rinv_0_lt_compat|].
field_simplify;
lra.
+
rewrite Rmult_minus_distr_r.
unfold Rminus;
apply Rplus_lt_compat_l.
apply Ropp_lt_contravar.
apply Rlt_le_trans with (
u2 *
bpow (
ln_beta y)).
*
rewrite <- (
Rmult_1_l (
u2 *
_)).
rewrite Rmult_assoc.
{
apply Rmult_lt_compat.
-
lra.
-
now apply Rmult_le_pos; [
apply bpow_ge_0|
apply Rlt_le].
-
lra.
-
apply Rmult_lt_compat_l; [
now apply bpow_gt_0|].
apply Rabs_lt_inv.
apply bpow_ln_beta_gt. }
*
unfold u2,
p,
ulp,
canonic_exp;
bpow_simplify;
apply bpow_le.
apply (
Zplus_le_reg_r _ _ (-
ln_beta y));
ring_simplify.
rewrite (
Zplus_comm (-
_));
fold (
Zminus (
ln_beta (
x /
y)) (
ln_beta y)).
destruct (
ln_beta_div_disj x y Px Py)
as [
Hxy|
Hxy];
rewrite Hxy;
[
now apply Hexp; [| |
rewrite <-
Hxy]|].
replace (
_ -
_ + 1)%
Z with ((
ln_beta x + 1) -
ln_beta y)%
Z by ring.
apply Hexp.
{
now assert (
fexp1 (
ln_beta x + 1) <=
ln_beta x)%
Z;
[
apply valid_exp|
omega]. }
{
assumption. }
replace (
_ + 1 -
_)%
Z with (
ln_beta x -
ln_beta y + 1)%
Z by ring.
now rewrite <-
Hxy.
-
apply Rle_lt_trans with (
p *
y -
bpow (
fexp1 (
ln_beta x))).
+
rewrite Fx at 1;
rewrite Fy at 1.
apply (
Rmult_le_reg_r (
bpow (-
fexp1 (
ln_beta x))));
[
now apply bpow_gt_0|].
rewrite Rmult_minus_distr_r.
bpow_simplify.
rewrite (
Rmult_comm p).
unfold p;
bpow_simplify.
rewrite <-
Z2R_Zpower; [|
omega].
rewrite <-
Z2R_mult.
change 1
with (
Z2R 1);
rewrite <-
Z2R_minus.
apply Z2R_le.
apply (
Zplus_le_reg_r _ _ 1);
ring_simplify.
apply Zlt_le_succ.
apply lt_Z2R.
rewrite Z2R_mult.
rewrite Z2R_Zpower; [|
omega].
apply (
Rmult_lt_reg_r (
bpow (
fexp1 (
ln_beta x))));
[
now apply bpow_gt_0|
bpow_simplify].
rewrite <-
Fx.
rewrite Zplus_comm;
rewrite bpow_plus.
rewrite <-
Rmult_assoc;
rewrite <-
Fy.
fold p.
apply (
Rmult_lt_reg_r (/
y)); [
now apply Rinv_0_lt_compat|].
field_simplify;
lra.
+
rewrite Rmult_minus_distr_r.
unfold Rminus;
apply Rplus_lt_compat_l.
apply Ropp_lt_contravar.
apply Rlt_le_trans with (
u2 *
bpow (
ln_beta y)).
*
rewrite <- (
Rmult_1_l (
u2 *
_)).
rewrite Rmult_assoc.
{
apply Rmult_lt_compat.
-
lra.
-
now apply Rmult_le_pos; [
apply bpow_ge_0|
apply Rlt_le].
-
lra.
-
apply Rmult_lt_compat_l; [
now apply bpow_gt_0|].
apply Rabs_lt_inv.
apply bpow_ln_beta_gt. }
*
unfold u2,
p,
ulp,
canonic_exp;
bpow_simplify;
apply bpow_le.
apply (
Zplus_le_reg_r _ _ (-
ln_beta y));
ring_simplify.
rewrite (
Zplus_comm (-
_));
fold (
Zminus (
ln_beta (
x /
y)) (
ln_beta y)).
destruct (
ln_beta_div_disj x y Px Py)
as [
Hxy|
Hxy];
rewrite Hxy;
apply Hexp;
try assumption;
rewrite <-
Hxy;
rewrite Hf1;
apply Zle_refl.
Qed.
Lemma double_round_div_aux1 :
forall fexp1 fexp2 :
Z ->
Z,
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_div_hyp fexp1 fexp2 ->
forall x y,
0 <
x -> 0 <
y ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
(
fexp1 (
ln_beta (
x /
y)) <=
ln_beta (
x /
y))%
Z ->
~ (
midp fexp1 (
x /
y) - / 2 *
ulp beta fexp2 (
x /
y)
<=
x /
y
<
midp fexp1 (
x /
y)).
Proof.
intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Px Py Fx Fy Hf1.
assert (
Hfx : (
fexp1 (
ln_beta x) <
ln_beta x)%
Z);
[
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|]|].
assert (
Hfy : (
fexp1 (
ln_beta y) <
ln_beta y)%
Z);
[
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|]|].
assert (
S : (
x /
y <> 0)%
R).
apply Rmult_integral_contrapositive_currified; [
now apply Rgt_not_eq|
idtac].
now apply Rinv_neq_0_compat,
Rgt_not_eq.
cut (~ (/ 2 * (
ulp beta fexp1 (
x /
y) -
ulp beta fexp2 (
x /
y))
<=
x /
y -
round beta fexp1 Zfloor (
x /
y)
< / 2 *
ulp beta fexp1 (
x /
y))).
{
intro H;
intro H';
apply H;
split.
-
apply (
Rplus_le_reg_l (
round beta fexp1 Zfloor (
x /
y))).
ring_simplify.
apply H'.
-
apply (
Rplus_lt_reg_l (
round beta fexp1 Zfloor (
x /
y))).
ring_simplify.
apply H'. }
set (
u1 :=
bpow (
fexp1 (
ln_beta (
x /
y)))).
set (
u2 :=
bpow (
fexp2 (
ln_beta (
x /
y)))).
set (
x' :=
round beta fexp1 Zfloor (
x /
y)).
rewrite 2!
ulp_neq_0;
trivial.
revert Fx Fy.
unfold generic_format,
F2R,
scaled_mantissa,
canonic_exp;
simpl.
set (
mx :=
Ztrunc (
x *
bpow (-
fexp1 (
ln_beta x)))).
set (
my :=
Ztrunc (
y *
bpow (-
fexp1 (
ln_beta y)))).
intros Fx Fy.
intro Hlr.
apply (
Rlt_irrefl (/ 2 * (
u1 -
u2))).
apply (
Rle_lt_trans _ _ _ (
proj1 Hlr)).
apply (
Rplus_lt_reg_r x');
ring_simplify.
apply (
Rmult_lt_reg_r y _ _ Py).
unfold Rdiv;
rewrite Rmult_assoc.
rewrite Rinv_l; [|
now apply Rgt_not_eq];
rewrite Rmult_1_r.
rewrite Rmult_minus_distr_r;
rewrite Rmult_plus_distr_r.
apply (
Rmult_lt_reg_l 2); [
lra|].
rewrite Rmult_minus_distr_l;
rewrite Rmult_plus_distr_l.
do 5
rewrite <-
Rmult_assoc.
rewrite Rinv_r; [|
lra];
do 2
rewrite Rmult_1_l.
destruct (
Zle_or_lt Z0 (
fexp1 (
ln_beta x) -
fexp1 (
ln_beta (
x /
y))
-
fexp1 (
ln_beta y))%
Z)
as [
He|
He].
-
apply Rle_lt_trans with (2 *
x' *
y +
u1 *
y
-
bpow (
fexp1 (
ln_beta (
x /
y))
+
fexp1 (
ln_beta y))).
+
rewrite Fx at 1;
rewrite Fy at 1 2.
apply (
Rmult_le_reg_r (
bpow (-
fexp1 (
ln_beta (
x /
y))
-
fexp1 (
ln_beta y))));
[
now apply bpow_gt_0|].
rewrite Rmult_minus_distr_r;
rewrite (
Rmult_plus_distr_r (
_ *
_ *
_)).
bpow_simplify.
replace (2 *
x' *
_ *
_)
with (2 *
Z2R my *
x' *
bpow (-
fexp1 (
ln_beta (
x /
y))))
by ring.
rewrite (
Rmult_comm u1).
unfold x',
u1,
round,
F2R,
ulp,
scaled_mantissa,
canonic_exp;
simpl.
bpow_simplify.
rewrite <-
Z2R_Zpower; [|
exact He].
change 2
with (
Z2R 2).
do 4
rewrite <-
Z2R_mult.
rewrite <-
Z2R_plus.
change 1
with (
Z2R 1);
rewrite <-
Z2R_minus.
apply Z2R_le.
apply (
Zplus_le_reg_r _ _ 1);
ring_simplify.
apply Zlt_le_succ.
apply lt_Z2R.
rewrite Z2R_plus.
do 4
rewrite Z2R_mult;
simpl.
rewrite Z2R_Zpower; [|
exact He].
apply (
Rmult_lt_reg_r (
bpow (
fexp1 (
ln_beta (
x /
y))
+
fexp1 (
ln_beta y))));
[
now apply bpow_gt_0|
bpow_simplify].
rewrite Rmult_assoc.
rewrite <-
Fx.
rewrite (
Rmult_plus_distr_r _ _ (
Fcore_Raux.bpow _ _)).
rewrite Rmult_assoc.
rewrite bpow_plus.
rewrite <- (
Rmult_assoc (
Z2R (
Zfloor _))).
change (
Z2R (
Zfloor _) *
_)
with x'.
do 2
rewrite (
Rmult_comm _ (
bpow (
fexp1 (
ln_beta y)))).
rewrite Rmult_assoc.
do 2
rewrite <- (
Rmult_assoc (
Z2R my)).
rewrite <-
Fy.
change (
bpow _)
with u1.
apply (
Rmult_lt_reg_l (/ 2)); [
lra|].
rewrite Rmult_plus_distr_l.
do 4
rewrite <-
Rmult_assoc.
rewrite Rinv_l; [|
lra];
do 2
rewrite Rmult_1_l.
apply (
Rplus_lt_reg_r (-
y *
x'));
ring_simplify.
apply (
Rmult_lt_reg_l (/
y)); [
now apply Rinv_0_lt_compat|].
rewrite Rmult_minus_distr_l.
do 3
rewrite <-
Rmult_assoc.
rewrite Rinv_l; [|
now apply Rgt_not_eq];
do 2
rewrite Rmult_1_l.
now rewrite Rmult_comm.
+
apply Rplus_lt_compat_l.
apply Ropp_lt_contravar.
apply Rlt_le_trans with (
u2 *
bpow (
ln_beta y)).
* {
apply Rmult_lt_compat_l.
-
apply bpow_gt_0.
-
apply Rabs_lt_inv.
apply bpow_ln_beta_gt. }
*
unfold u2,
ulp,
canonic_exp;
bpow_simplify;
apply bpow_le.
apply (
Zplus_le_reg_r _ _ (-
ln_beta y));
ring_simplify.
rewrite <-
Zplus_assoc;
rewrite (
Zplus_comm (-
_)).
destruct (
ln_beta_div_disj x y Px Py)
as [
Hxy|
Hxy];
rewrite Hxy;
[
now apply Hexp; [| |
rewrite <-
Hxy]|].
replace (
_ -
_ + 1)%
Z with ((
ln_beta x + 1) -
ln_beta y)%
Z by ring.
apply Hexp.
{
now assert (
fexp1 (
ln_beta x + 1) <=
ln_beta x)%
Z;
[
apply valid_exp|
omega]. }
{
assumption. }
replace (
_ + 1 -
_)%
Z with (
ln_beta x -
ln_beta y + 1)%
Z by ring.
now rewrite <-
Hxy.
-
apply Rle_lt_trans with (2 *
x' *
y +
u1 *
y -
bpow (
fexp1 (
ln_beta x))).
+
rewrite Fx at 1;
rewrite Fy at 1 2.
apply (
Rmult_le_reg_r (
bpow (-
fexp1 (
ln_beta x))));
[
now apply bpow_gt_0|].
rewrite Rmult_minus_distr_r;
rewrite (
Rmult_plus_distr_r (
_ *
_ *
_)).
bpow_simplify.
replace (2 *
x' *
_ *
_)
with (2 *
Z2R my *
x' *
bpow (
fexp1 (
ln_beta y) -
fexp1 (
ln_beta x)))
by ring.
rewrite (
Rmult_comm u1).
unfold x',
u1,
round,
F2R,
ulp,
scaled_mantissa,
canonic_exp;
simpl.
bpow_simplify.
rewrite <- (
Z2R_Zpower _ (
_ -
_)%
Z); [|
omega].
change 2
with (
Z2R 2).
do 5
rewrite <-
Z2R_mult.
rewrite <-
Z2R_plus.
change 1
with (
Z2R 1);
rewrite <-
Z2R_minus.
apply Z2R_le.
apply (
Zplus_le_reg_r _ _ 1);
ring_simplify.
apply Zlt_le_succ.
apply lt_Z2R.
rewrite Z2R_plus.
do 5
rewrite Z2R_mult;
simpl.
rewrite Z2R_Zpower; [|
omega].
apply (
Rmult_lt_reg_r (
bpow (
fexp1 (
ln_beta x))));
[
now apply bpow_gt_0|].
rewrite Rmult_assoc.
rewrite <-
Fx.
rewrite (
Rmult_plus_distr_r _ _ (
Fcore_Raux.bpow _ _)).
bpow_simplify.
rewrite Rmult_assoc.
rewrite bpow_plus.
rewrite <- (
Rmult_assoc (
Z2R (
Zfloor _))).
change (
Z2R (
Zfloor _) *
_)
with x'.
do 2
rewrite (
Rmult_comm _ (
bpow (
fexp1 (
ln_beta y)))).
rewrite Rmult_assoc.
do 2
rewrite <- (
Rmult_assoc (
Z2R my)).
rewrite <-
Fy.
change (
bpow _)
with u1.
apply (
Rmult_lt_reg_l (/ 2)); [
lra|].
rewrite Rmult_plus_distr_l.
do 4
rewrite <-
Rmult_assoc.
rewrite Rinv_l; [|
lra];
do 2
rewrite Rmult_1_l.
apply (
Rplus_lt_reg_r (-
y *
x'));
ring_simplify.
apply (
Rmult_lt_reg_l (/
y)); [
now apply Rinv_0_lt_compat|].
rewrite Rmult_minus_distr_l.
do 3
rewrite <-
Rmult_assoc.
rewrite Rinv_l; [|
now apply Rgt_not_eq];
do 2
rewrite Rmult_1_l.
now rewrite Rmult_comm.
+
apply Rplus_lt_compat_l.
apply Ropp_lt_contravar.
apply Rlt_le_trans with (
u2 *
bpow (
ln_beta y)).
* {
apply Rmult_lt_compat_l.
-
apply bpow_gt_0.
-
apply Rabs_lt_inv.
apply bpow_ln_beta_gt. }
*
unfold u2,
ulp,
canonic_exp;
bpow_simplify;
apply bpow_le.
apply (
Zplus_le_reg_r _ _ (-
ln_beta y));
ring_simplify.
rewrite (
Zplus_comm (-
_)).
destruct (
ln_beta_div_disj x y Px Py)
as [
Hxy|
Hxy];
rewrite Hxy;
apply Hexp;
try assumption;
rewrite <-
Hxy;
omega.
Qed.
Lemma double_round_div_aux2 :
forall fexp1 fexp2 :
Z ->
Z,
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
double_round_div_hyp fexp1 fexp2 ->
forall x y,
0 <
x -> 0 <
y ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
(
fexp1 (
ln_beta (
x /
y)) <=
ln_beta (
x /
y))%
Z ->
~ (
midp fexp1 (
x /
y)
<
x /
y
<=
midp fexp1 (
x /
y) + / 2 *
ulp beta fexp2 (
x /
y)).
Proof.
intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Px Py Fx Fy Hf1.
assert (
Hfx : (
fexp1 (
ln_beta x) <
ln_beta x)%
Z);
[
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|]|].
assert (
Hfy : (
fexp1 (
ln_beta y) <
ln_beta y)%
Z);
[
now apply ln_beta_generic_gt; [|
apply Rgt_not_eq|]|].
cut (~ (/ 2 *
ulp beta fexp1 (
x /
y)
<
x /
y -
round beta fexp1 Zfloor (
x /
y)
<= / 2 * (
ulp beta fexp1 (
x /
y) +
ulp beta fexp2 (
x /
y)))).
{
intro H;
intro H';
apply H;
split.
-
apply (
Rplus_lt_reg_l (
round beta fexp1 Zfloor (
x /
y))).
ring_simplify.
apply H'.
-
apply (
Rplus_le_reg_l (
round beta fexp1 Zfloor (
x /
y))).
ring_simplify.
apply H'. }
set (
u1 :=
bpow (
fexp1 (
ln_beta (
x /
y)))).
set (
u2 :=
bpow (
fexp2 (
ln_beta (
x /
y)))).
set (
x' :=
round beta fexp1 Zfloor (
x /
y)).
assert (
S : (
x /
y <> 0)%
R).
apply Rmult_integral_contrapositive_currified; [
now apply Rgt_not_eq|
idtac].
now apply Rinv_neq_0_compat,
Rgt_not_eq.
rewrite 2!
ulp_neq_0;
trivial.
revert Fx Fy.
unfold generic_format,
F2R,
scaled_mantissa,
canonic_exp;
simpl.
set (
mx :=
Ztrunc (
x *
bpow (-
fexp1 (
ln_beta x)))).
set (
my :=
Ztrunc (
y *
bpow (-
fexp1 (
ln_beta y)))).
intros Fx Fy.
intro Hlr.
apply (
Rlt_irrefl (/ 2 * (
u1 +
u2))).
apply Rlt_le_trans with (
x /
y -
x'); [|
now apply Hlr].
apply (
Rplus_lt_reg_r x');
ring_simplify.
apply (
Rmult_lt_reg_r y _ _ Py).
unfold Rdiv;
rewrite Rmult_assoc.
rewrite Rinv_l; [|
now apply Rgt_not_eq];
rewrite Rmult_1_r.
do 2
rewrite Rmult_plus_distr_r.
apply (
Rmult_lt_reg_l 2); [
lra|].
do 2
rewrite Rmult_plus_distr_l.
do 5
rewrite <-
Rmult_assoc.
rewrite Rinv_r; [|
lra];
do 2
rewrite Rmult_1_l.
destruct (
Zle_or_lt Z0 (
fexp1 (
ln_beta x) -
fexp1 (
ln_beta (
x /
y))
-
fexp1 (
ln_beta y))%
Z)
as [
He|
He].
-
apply Rlt_le_trans with (
u1 *
y +
bpow (
fexp1 (
ln_beta (
x /
y))
+
fexp1 (
ln_beta y))
+ 2 *
x' *
y).
+
apply Rplus_lt_compat_r,
Rplus_lt_compat_l.
apply Rlt_le_trans with (
u2 *
bpow (
ln_beta y)).
* {
apply Rmult_lt_compat_l.
-
apply bpow_gt_0.
-
apply Rabs_lt_inv.
apply bpow_ln_beta_gt. }
*
unfold u2,
ulp,
canonic_exp;
bpow_simplify;
apply bpow_le.
apply (
Zplus_le_reg_r _ _ (-
ln_beta y));
ring_simplify.
rewrite <-
Zplus_assoc;
rewrite (
Zplus_comm (-
_)).
destruct (
ln_beta_div_disj x y Px Py)
as [
Hxy|
Hxy];
rewrite Hxy;
[
now apply Hexp; [| |
rewrite <-
Hxy]|].
replace (
_ -
_ + 1)%
Z with ((
ln_beta x + 1) -
ln_beta y)%
Z by ring.
apply Hexp.
{
now assert (
fexp1 (
ln_beta x + 1) <=
ln_beta x)%
Z;
[
apply valid_exp|
omega]. }
{
assumption. }
replace (
_ + 1 -
_)%
Z with (
ln_beta x -
ln_beta y + 1)%
Z by ring.
now rewrite <-
Hxy.
+
apply Rge_le;
rewrite Fx at 1;
apply Rle_ge.
replace (
u1 *
y)
with (
u1 * (
Z2R my *
bpow (
fexp1 (
ln_beta y))));
[|
now apply eq_sym;
rewrite Fy at 1].
replace (2 *
x' *
y)
with (2 *
x' * (
Z2R my *
bpow (
fexp1 (
ln_beta y))));
[|
now apply eq_sym;
rewrite Fy at 1].
apply (
Rmult_le_reg_r (
bpow (-
fexp1 (
ln_beta (
x /
y))
-
fexp1 (
ln_beta y))));
[
now apply bpow_gt_0|].
do 2
rewrite Rmult_plus_distr_r.
bpow_simplify.
rewrite (
Rmult_comm u1).
unfold u1,
ulp,
canonic_exp;
bpow_simplify.
rewrite (
Rmult_assoc 2).
rewrite (
Rmult_comm x').
rewrite (
Rmult_assoc 2).
unfold x',
round,
F2R,
scaled_mantissa,
canonic_exp;
simpl.
bpow_simplify.
rewrite <- (
Z2R_Zpower _ (
_ -
_)%
Z); [|
exact He].
change 2
with (
Z2R 2).
do 4
rewrite <-
Z2R_mult.
change 1
with (
Z2R 1);
do 2
rewrite <-
Z2R_plus.
apply Z2R_le.
rewrite Zplus_comm,
Zplus_assoc.
apply Zlt_le_succ.
apply lt_Z2R.
rewrite Z2R_plus.
do 4
rewrite Z2R_mult;
simpl.
rewrite Z2R_Zpower; [|
exact He].
apply (
Rmult_lt_reg_r (
bpow (
fexp1 (
ln_beta y))));
[
now apply bpow_gt_0|].
rewrite Rmult_plus_distr_r.
rewrite (
Rmult_comm _ (
Z2R _)).
do 2
rewrite Rmult_assoc.
rewrite <-
Fy.
bpow_simplify.
unfold Zminus;
rewrite bpow_plus.
rewrite (
Rmult_assoc _ (
Z2R mx)).
rewrite <- (
Rmult_assoc (
Z2R mx)).
rewrite <-
Fx.
apply (
Rmult_lt_reg_r (
bpow (
fexp1 (
ln_beta (
x /
y)))));
[
now apply bpow_gt_0|].
rewrite Rmult_plus_distr_r.
bpow_simplify.
rewrite (
Rmult_comm _ y).
do 2
rewrite Rmult_assoc.
change (
Z2R _ *
_)
with x'.
change (
bpow _)
with u1.
apply (
Rmult_lt_reg_l (/ 2)); [
lra|].
rewrite Rmult_plus_distr_l.
do 4
rewrite <-
Rmult_assoc.
rewrite Rinv_l; [|
lra];
do 2
rewrite Rmult_1_l.
apply (
Rplus_lt_reg_r (-
y *
x'));
ring_simplify.
apply (
Rmult_lt_reg_l (/
y)); [
now apply Rinv_0_lt_compat|].
rewrite Rmult_plus_distr_l.
do 3
rewrite <-
Rmult_assoc.
rewrite Ropp_mult_distr_r_reverse.
rewrite Ropp_mult_distr_l_reverse.
rewrite Rinv_l; [|
now apply Rgt_not_eq];
do 2
rewrite Rmult_1_l.
rewrite (
Rmult_comm (/
y)).
now rewrite (
Rplus_comm (-
x')).
-
apply Rlt_le_trans with (2 *
x' *
y +
u1 *
y +
bpow (
fexp1 (
ln_beta x))).
+
rewrite Rplus_comm,
Rplus_assoc;
do 2
apply Rplus_lt_compat_l.
apply Rlt_le_trans with (
u2 *
bpow (
ln_beta y)).
*
apply Rmult_lt_compat_l.
now apply bpow_gt_0.
now apply Rabs_lt_inv;
apply bpow_ln_beta_gt.
*
unfold u2,
ulp,
canonic_exp;
bpow_simplify;
apply bpow_le.
apply (
Zplus_le_reg_r _ _ (-
ln_beta y));
ring_simplify.
rewrite (
Zplus_comm (-
_)).
destruct (
ln_beta_div_disj x y Px Py)
as [
Hxy|
Hxy];
rewrite Hxy;
apply Hexp;
try assumption;
rewrite <-
Hxy;
omega.
+
apply Rge_le;
rewrite Fx at 1;
apply Rle_ge.
rewrite Fy at 1 2.
apply (
Rmult_le_reg_r (
bpow (-
fexp1 (
ln_beta x))));
[
now apply bpow_gt_0|].
do 2
rewrite Rmult_plus_distr_r.
bpow_simplify.
replace (2 *
x' *
_ *
_)
with (2 *
Z2R my *
x' *
bpow (
fexp1 (
ln_beta y) -
fexp1 (
ln_beta x)))
by ring.
rewrite (
Rmult_comm u1).
unfold x',
u1,
round,
F2R,
ulp,
scaled_mantissa,
canonic_exp;
simpl.
bpow_simplify.
rewrite <- (
Z2R_Zpower _ (
_ -
_)%
Z); [|
omega].
change 2
with (
Z2R 2).
do 5
rewrite <-
Z2R_mult.
change 1
with (
Z2R 1);
do 2
rewrite <-
Z2R_plus.
apply Z2R_le.
apply Zlt_le_succ.
apply lt_Z2R.
rewrite Z2R_plus.
do 5
rewrite Z2R_mult;
simpl.
rewrite Z2R_Zpower; [|
omega].
apply (
Rmult_lt_reg_r (
bpow (
fexp1 (
ln_beta x))));
[
now apply bpow_gt_0|].
rewrite (
Rmult_assoc _ (
Z2R mx)).
rewrite <-
Fx.
rewrite Rmult_plus_distr_r.
bpow_simplify.
rewrite bpow_plus.
rewrite Rmult_assoc.
rewrite <- (
Rmult_assoc (
Z2R _)).
change (
Z2R _ *
bpow _)
with x'.
do 2
rewrite (
Rmult_comm _ (
bpow (
fexp1 (
ln_beta y)))).
rewrite Rmult_assoc.
do 2
rewrite <- (
Rmult_assoc (
Z2R my)).
rewrite <-
Fy.
change (
bpow _)
with u1.
apply (
Rmult_lt_reg_l (/ 2)); [
lra|].
rewrite Rmult_plus_distr_l.
do 4
rewrite <-
Rmult_assoc.
rewrite Rinv_l; [|
lra];
do 2
rewrite Rmult_1_l.
apply (
Rplus_lt_reg_r (-
y *
x'));
ring_simplify.
apply (
Rmult_lt_reg_l (/
y)); [
now apply Rinv_0_lt_compat|].
rewrite Rmult_plus_distr_l.
do 3
rewrite <-
Rmult_assoc.
rewrite Ropp_mult_distr_r_reverse.
rewrite Ropp_mult_distr_l_reverse.
rewrite Rinv_l; [|
now apply Rgt_not_eq];
do 2
rewrite Rmult_1_l.
rewrite (
Rmult_comm (/
y)).
now rewrite (
Rplus_comm (-
x')).
Qed.
Lemma double_round_div_aux :
forall fexp1 fexp2 :
Z ->
Z,
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
(
exists n, (
beta = 2 *
n :>
Z)%
Z) ->
double_round_div_hyp fexp1 fexp2 ->
forall x y,
0 <
x -> 0 <
y ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x /
y).
Proof.
Lemma double_round_div :
forall fexp1 fexp2 :
Z ->
Z,
Valid_exp fexp1 ->
Valid_exp fexp2 ->
forall (
choice1 choice2 :
Z ->
bool),
(
exists n, (
beta = 2 *
n :>
Z)%
Z) ->
double_round_div_hyp fexp1 fexp2 ->
forall x y,
y <> 0 ->
generic_format beta fexp1 x ->
generic_format beta fexp1 y ->
double_round_eq fexp1 fexp2 choice1 choice2 (
x /
y).
Proof.
Section Double_round_div_FLX.
Require Import Fcore_FLX.
Variable prec :
Z.
Variable prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FLX_double_round_div_hyp :
(2 *
prec <=
prec')%
Z ->
double_round_div_hyp (
FLX_exp prec) (
FLX_exp prec').
Proof.
Theorem double_round_div_FLX :
forall choice1 choice2,
(
exists n, (
beta = 2 *
n :>
Z)%
Z) ->
(2 *
prec <=
prec')%
Z ->
forall x y,
y <> 0 ->
FLX_format beta prec x ->
FLX_format beta prec y ->
double_round_eq (
FLX_exp prec) (
FLX_exp prec')
choice1 choice2 (
x /
y).
Proof.
End Double_round_div_FLX.
Section Double_round_div_FLT.
Require Import Fcore_FLX.
Require Import Fcore_FLT.
Variable emin prec :
Z.
Variable emin'
prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FLT_double_round_div_hyp :
(
emin' <=
emin -
prec - 2)%
Z ->
(2 *
prec <=
prec')%
Z ->
double_round_div_hyp (
FLT_exp emin prec) (
FLT_exp emin'
prec').
Proof.
Theorem double_round_div_FLT :
forall choice1 choice2,
(
exists n, (
beta = 2 *
n :>
Z)%
Z) ->
(
emin' <=
emin -
prec - 2)%
Z ->
(2 *
prec <=
prec')%
Z ->
forall x y,
y <> 0 ->
FLT_format beta emin prec x ->
FLT_format beta emin prec y ->
double_round_eq (
FLT_exp emin prec) (
FLT_exp emin'
prec')
choice1 choice2 (
x /
y).
Proof.
End Double_round_div_FLT.
Section Double_round_div_FTZ.
Require Import Fcore_FLX.
Require Import Fcore_FTZ.
Variable emin prec :
Z.
Variable emin'
prec' :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Context {
prec_gt_0_' :
Prec_gt_0 prec' }.
Lemma FTZ_double_round_div_hyp :
(
emin' +
prec' <=
emin - 1)%
Z ->
(2 *
prec <=
prec')%
Z ->
double_round_div_hyp (
FTZ_exp emin prec) (
FTZ_exp emin'
prec').
Proof.
Theorem double_round_div_FTZ :
forall choice1 choice2,
(
exists n, (
beta = 2 *
n :>
Z)%
Z) ->
(
emin' +
prec' <=
emin - 1)%
Z ->
(2 *
prec <=
prec')%
Z ->
forall x y,
y <> 0 ->
FTZ_format beta emin prec x ->
FTZ_format beta emin prec y ->
double_round_eq (
FTZ_exp emin prec) (
FTZ_exp emin'
prec')
choice1 choice2 (
x /
y).
Proof.
End Double_round_div_FTZ.
End Double_round_div.
End Double_round.