aboutsummaryrefslogtreecommitdiffstats
path: root/flocq/Core/Fcore_FTZ.v
diff options
context:
space:
mode:
authorGuillaume Melquiond <guillaume.melquiond@inria.fr>2015-09-22 15:41:50 +0200
committerGuillaume Melquiond <guillaume.melquiond@inria.fr>2015-09-22 15:41:50 +0200
commit0af966a42eb60e9af43f9a450d924758a83946c6 (patch)
tree2bef73e80a8a80da1c47deee66dc825feac45bdd /flocq/Core/Fcore_FTZ.v
parentc212ab7a8adea516db72f17d818393629dbde1b3 (diff)
downloadcompcert-kvx-0af966a42eb60e9af43f9a450d924758a83946c6.tar.gz
compcert-kvx-0af966a42eb60e9af43f9a450d924758a83946c6.zip
Upgrade to Flocq 2.5.0.
Note: this version of Flocq is compatible with both Coq 8.4 and 8.5.
Diffstat (limited to 'flocq/Core/Fcore_FTZ.v')
-rw-r--r--flocq/Core/Fcore_FTZ.v17
1 files changed, 16 insertions, 1 deletions
diff --git a/flocq/Core/Fcore_FTZ.v b/flocq/Core/Fcore_FTZ.v
index 5f3e5337..2ebc7851 100644
--- a/flocq/Core/Fcore_FTZ.v
+++ b/flocq/Core/Fcore_FTZ.v
@@ -23,6 +23,7 @@ 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.
@@ -182,7 +183,6 @@ Theorem FTZ_format_FLXN :
(bpow (emin + prec - 1) <= Rabs x)%R ->
FLXN_format beta prec x -> FTZ_format x.
Proof.
-clear prec_gt_0_.
intros x Hx Fx.
apply FTZ_format_generic.
apply generic_format_FLXN in Fx.
@@ -195,6 +195,21 @@ apply Zle_refl.
omega.
Qed.
+Theorem ulp_FTZ_0: ulp beta FTZ_exp 0 = bpow (emin+prec-1).
+Proof with auto with typeclass_instances.
+unfold ulp; rewrite Req_bool_true; trivial.
+case (negligible_exp_spec FTZ_exp).
+intros T; specialize (T (emin-1)%Z); contradict T.
+apply Zle_not_lt; unfold FTZ_exp; unfold Prec_gt_0 in prec_gt_0_.
+rewrite Zlt_bool_true; omega.
+assert (V:(FTZ_exp (emin+prec-1) = emin+prec-1)%Z).
+unfold FTZ_exp; rewrite Zlt_bool_true; omega.
+intros n H2; rewrite <-V.
+apply f_equal, fexp_negligible_exp_eq...
+omega.
+Qed.
+
+
Section FTZ_round.
(** Rounding with FTZ *)