Unit in the Last Place: our definition using fexp and its properties, successor and predecessor
Require Import Fcore_Raux.
Require Import Fcore_defs.
Require Import Fcore_rnd.
Require Import Fcore_generic_fmt.
Require Import Fcore_float_prop.
Section Fcore_ulp.
Variable beta :
radix.
Notation bpow e := (
bpow beta e).
Variable fexp :
Z ->
Z.
Context {
valid_exp :
Valid_exp fexp }.
Definition ulp x :=
bpow (
canonic_exp beta fexp x).
Notation F := (
generic_format beta fexp).
Theorem ulp_opp :
forall x,
ulp (-
x) =
ulp x.
Proof.
Theorem ulp_abs :
forall x,
ulp (
Rabs x) =
ulp x.
Proof.
Theorem ulp_le_id:
forall x,
(0 <
x)%
R ->
F x ->
(
ulp x <=
x)%
R.
Proof.
Theorem ulp_le_abs:
forall x,
(
x <> 0)%
R ->
F x ->
(
ulp x <=
Rabs x)%
R.
Proof.
Theorem ulp_DN_UP :
forall x, ~
F x ->
round beta fexp Zceil x = (
round beta fexp Zfloor x +
ulp x)%
R.
Proof.
The successor of x is x + ulp x
Theorem succ_le_bpow :
forall x e, (0 <
x)%
R ->
F x ->
(
x <
bpow e)%
R ->
(
x +
ulp x <=
bpow e)%
R.
Proof.
Theorem ln_beta_succ :
forall x, (0 <
x)%
R ->
F x ->
forall eps, (0 <=
eps <
ulp x)%
R ->
ln_beta beta (
x +
eps) =
ln_beta beta x :>
Z.
Proof.
Theorem round_DN_succ :
forall x, (0 <
x)%
R ->
F x ->
forall eps, (0 <=
eps <
ulp x)%
R ->
round beta fexp Zfloor (
x +
eps) =
x.
Proof.
Theorem generic_format_succ :
forall x, (0 <
x)%
R ->
F x ->
F (
x +
ulp x).
Proof.
Theorem round_UP_succ :
forall x, (0 <
x)%
R ->
F x ->
forall eps, (0 <
eps <=
ulp x)%
R ->
round beta fexp Zceil (
x +
eps) = (
x +
ulp x)%
R.
Proof with
Theorem succ_le_lt:
forall x y,
F x ->
F y ->
(0 <
x <
y)%
R ->
(
x +
ulp x <=
y)%
R.
Proof with
Error of a rounding, expressed in number of ulps
Theorem ulp_error :
forall rnd {
Zrnd :
Valid_rnd rnd }
x,
(
Rabs (
round beta fexp rnd x -
x) <
ulp x)%
R.
Proof with
Theorem ulp_half_error :
forall choice x,
(
Rabs (
round beta fexp (
Znearest choice)
x -
x) <= /2 *
ulp x)%
R.
Proof with
Theorem ulp_le :
forall {
Hm :
Monotone_exp fexp },
forall x y:
R,
(0 <
x)%
R -> (
x <=
y)%
R ->
(
ulp x <=
ulp y)%
R.
Proof.
Theorem ulp_bpow :
forall e,
ulp (
bpow e) =
bpow (
fexp (
e + 1)).
intros e.
unfold ulp.
apply f_equal.
apply canonic_exp_fexp.
rewrite Rabs_pos_eq.
split.
ring_simplify (
e + 1 - 1)%
Z.
apply Rle_refl.
apply bpow_lt.
apply Zlt_succ.
apply bpow_ge_0.
Qed.
Proof.
Proof.
Proof.
Proof.
Proof.
Proof.
Proof.
Proof.
Proof.
Proof.
Proof.
Proof.
Proof.
Proof.