Floating-point format with abrupt underflow
Require Import Fcore_Raux.
Require Import Fcore_defs.
Require Import Fcore_rnd.
Require Import Fcore_generic_fmt.
Require Import Fcore_float_prop.
Require Import Fcore_ulp.
Require Import Fcore_FLX.
Section RND_FTZ.
Variable beta :
radix.
Notation bpow e := (
bpow beta e).
Variable emin prec :
Z.
Context {
prec_gt_0_ :
Prec_gt_0 prec }.
Definition FTZ_format (
x :
R) :=
exists f :
float beta,
x =
F2R f /\ (
x <>
R0 ->
Zpower beta (
prec - 1) <=
Zabs (
Fnum f) <
Zpower beta prec)%
Z /\
(
emin <=
Fexp f)%
Z.
Definition FTZ_exp e :=
if Zlt_bool (
e -
prec)
emin then (
emin +
prec - 1)%
Z else (
e -
prec)%
Z.
Properties of the FTZ format
Global Instance FTZ_exp_valid :
Valid_exp FTZ_exp.
Proof.
Theorem FLXN_format_FTZ :
forall x,
FTZ_format x ->
FLXN_format beta prec x.
Proof.
intros x ((
xm,
xe), (
Hx1, (
Hx2,
Hx3))).
eexists.
apply (
conj Hx1 Hx2).
Qed.
Theorem generic_format_FTZ :
forall x,
FTZ_format x ->
generic_format beta FTZ_exp x.
Proof.
Theorem FTZ_format_generic :
forall x,
generic_format beta FTZ_exp x ->
FTZ_format x.
Proof.
Theorem FTZ_format_satisfies_any :
satisfies_any FTZ_format.
Proof.
Theorem FTZ_format_FLXN :
forall x :
R,
(
bpow (
emin +
prec - 1) <=
Rabs x)%
R ->
FLXN_format beta prec x ->
FTZ_format x.
Proof.
Theorem ulp_FTZ_0:
ulp beta FTZ_exp 0 =
bpow (
emin+
prec-1).
Proof with
Section FTZ_round.
Rounding with FTZ
Variable rnd :
R ->
Z.
Context {
valid_rnd :
Valid_rnd rnd }.
Definition Zrnd_FTZ x :=
if Rle_bool R1 (
Rabs x)
then rnd x else Z0.
Global Instance valid_rnd_FTZ :
Valid_rnd Zrnd_FTZ.
Proof with
Theorem round_FTZ_FLX :
forall x :
R,
(
bpow (
emin +
prec - 1) <=
Rabs x)%
R ->
round beta FTZ_exp Zrnd_FTZ x =
round beta (
FLX_exp prec)
rnd x.
Proof.
Theorem round_FTZ_small :
forall x :
R,
(
Rabs x <
bpow (
emin +
prec - 1))%
R ->
round beta FTZ_exp Zrnd_FTZ x =
R0.
Proof with
End FTZ_round.
End RND_FTZ.