Module Fprop_Sterbenz
Sterbenz conditions for exact subtraction
Require Import Fcore_Raux.
Require Import Fcore_defs.
Require Import Fcore_generic_fmt.
Require Import Fcalc_ops.
Section Fprop_Sterbenz.
Variable beta :
radix.
Notation bpow e := (
bpow beta e).
Variable fexp :
Z ->
Z.
Context {
valid_exp :
Valid_exp fexp }.
Context {
monotone_exp :
Monotone_exp fexp }.
Notation format := (
generic_format beta fexp).
Theorem generic_format_plus :
forall x y,
format x ->
format y ->
(
Rabs (
x +
y) <
bpow (
Zmin (
ln_beta beta x) (
ln_beta beta y)))%
R ->
format (
x +
y)%
R.
Proof.
intros x y Fx Fy Hxy.
destruct (
Req_dec (
x +
y) 0)
as [
Zxy|
Zxy].
rewrite Zxy.
apply generic_format_0.
destruct (
Req_dec x R0)
as [
Zx|
Zx].
now rewrite Zx,
Rplus_0_l.
destruct (
Req_dec y R0)
as [
Zy|
Zy].
now rewrite Zy,
Rplus_0_r.
revert Hxy.
destruct (
ln_beta beta x)
as (
ex,
Ex).
simpl.
specialize (
Ex Zx).
destruct (
ln_beta beta y)
as (
ey,
Ey).
simpl.
specialize (
Ey Zy).
intros Hxy.
set (
fx :=
Float beta (
Ztrunc (
scaled_mantissa beta fexp x)) (
fexp ex)).
assert (
Hx:
x =
F2R fx).
rewrite Fx at 1.
unfold canonic_exp.
now rewrite ln_beta_unique with (1 :=
Ex).
set (
fy :=
Float beta (
Ztrunc (
scaled_mantissa beta fexp y)) (
fexp ey)).
assert (
Hy:
y =
F2R fy).
rewrite Fy at 1.
unfold canonic_exp.
now rewrite ln_beta_unique with (1 :=
Ey).
rewrite Hx,
Hy.
rewrite <-
F2R_plus.
apply generic_format_F2R.
intros _.
case_eq (
Fplus beta fx fy).
intros mxy exy Pxy.
rewrite <-
Pxy,
F2R_plus, <-
Hx, <-
Hy.
unfold canonic_exp.
replace exy with (
fexp (
Zmin ex ey)).
apply monotone_exp.
now apply ln_beta_le_bpow.
replace exy with (
Fexp (
Fplus beta fx fy))
by exact (
f_equal Fexp Pxy).
rewrite Fexp_Fplus.
simpl.
clear -
monotone_exp.
apply sym_eq.
destruct (
Zmin_spec ex ey)
as [(
H1,
H2)|(
H1,
H2)] ;
rewrite H2.
apply Zmin_l.
now apply monotone_exp.
apply Zmin_r.
apply monotone_exp.
apply Zlt_le_weak.
now apply Zgt_lt.
Qed.
Theorem generic_format_plus_weak :
forall x y,
format x ->
format y ->
(
Rabs (
x +
y) <=
Rmin (
Rabs x) (
Rabs y))%
R ->
format (
x +
y)%
R.
Proof.
Lemma sterbenz_aux :
forall x y,
format x ->
format y ->
(
y <=
x <= 2 *
y)%
R ->
format (
x -
y)%
R.
Proof.
Theorem sterbenz :
forall x y,
format x ->
format y ->
(
y / 2 <=
x <= 2 *
y)%
R ->
format (
x -
y)%
R.
Proof.
End Fprop_Sterbenz.