aboutsummaryrefslogtreecommitdiffstats
path: root/flocq/Core
diff options
context:
space:
mode:
authorGuillaume Melquiond <guillaume.melquiond@inria.fr>2019-02-13 18:53:17 +0100
committerXavier Leroy <xavierleroy@users.noreply.github.com>2019-03-27 11:38:25 +0100
commit0f919eb26c68d3882e612a1b3a9df45bee6d3624 (patch)
treeb8bcf57e06d761be09b8d2cf2f80741acb1e4949 /flocq/Core
parentd5c0b4054c8490bda3b3d191724c58d5d4002e58 (diff)
downloadcompcert-kvx-0f919eb26c68d3882e612a1b3a9df45bee6d3624.tar.gz
compcert-kvx-0f919eb26c68d3882e612a1b3a9df45bee6d3624.zip
Upgrade embedded version of Flocq to 3.1.
Main changes to CompCert outside of Flocq are as follows: - Minimal supported version of Coq is now 8.7, due to Flocq requirements. - Most modifications are due to Z2R being dropped in favor of IZR and to the way Flocq now handles NaNs. - CompCert now correctly handles NaNs for the Risc-V architecture (hopefully).
Diffstat (limited to 'flocq/Core')
-rw-r--r--flocq/Core/Core.v (renamed from flocq/Core/Fcore.v)16
-rw-r--r--flocq/Core/Defs.v (renamed from flocq/Core/Fcore_defs.v)36
-rw-r--r--flocq/Core/Digits.v (renamed from flocq/Core/Fcore_digits.v)211
-rw-r--r--flocq/Core/FIX.v (renamed from flocq/Core/Fcore_FIX.v)30
-rw-r--r--flocq/Core/FLT.v (renamed from flocq/Core/Fcore_FLT.v)182
-rw-r--r--flocq/Core/FLX.v362
-rw-r--r--flocq/Core/FTZ.v (renamed from flocq/Core/Fcore_FTZ.v)109
-rw-r--r--flocq/Core/Fcore_FLX.v271
-rw-r--r--flocq/Core/Float_prop.v (renamed from flocq/Core/Fcore_float_prop.v)228
-rw-r--r--flocq/Core/Generic_fmt.v (renamed from flocq/Core/Fcore_generic_fmt.v)793
-rw-r--r--flocq/Core/Raux.v (renamed from flocq/Core/Fcore_Raux.v)964
-rw-r--r--flocq/Core/Round_NE.v (renamed from flocq/Core/Fcore_rnd_ne.v)185
-rw-r--r--flocq/Core/Round_pred.v (renamed from flocq/Core/Fcore_rnd.v)176
-rw-r--r--flocq/Core/Ulp.v (renamed from flocq/Core/Fcore_ulp.v)925
-rw-r--r--flocq/Core/Zaux.v (renamed from flocq/Core/Fcore_Zaux.v)238
15 files changed, 2434 insertions, 2292 deletions
diff --git a/flocq/Core/Fcore.v b/flocq/Core/Core.v
index 2a5a5f02..78a140e1 100644
--- a/flocq/Core/Fcore.v
+++ b/flocq/Core/Core.v
@@ -2,9 +2,9 @@
This file is part of the Flocq formalization of floating-point
arithmetic in Coq: http://flocq.gforge.inria.fr/
-Copyright (C) 2010-2013 Sylvie Boldo
+Copyright (C) 2010-2018 Sylvie Boldo
#<br />#
-Copyright (C) 2010-2013 Guillaume Melquiond
+Copyright (C) 2010-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
@@ -18,13 +18,5 @@ COPYING file for more details.
*)
(** To ease the import *)
-Require Export Fcore_Raux.
-Require Export Fcore_defs.
-Require Export Fcore_float_prop.
-Require Export Fcore_rnd.
-Require Export Fcore_generic_fmt.
-Require Export Fcore_rnd_ne.
-Require Export Fcore_FIX.
-Require Export Fcore_FLX.
-Require Export Fcore_FLT.
-Require Export Fcore_ulp.
+Require Export Raux Defs Float_prop Round_pred Generic_fmt Round_NE.
+Require Export FIX FLX FLT Ulp.
diff --git a/flocq/Core/Fcore_defs.v b/flocq/Core/Defs.v
index 01b868ab..f5c6f33b 100644
--- a/flocq/Core/Fcore_defs.v
+++ b/flocq/Core/Defs.v
@@ -2,9 +2,9 @@
This file is part of the Flocq formalization of floating-point
arithmetic in Coq: http://flocq.gforge.inria.fr/
-Copyright (C) 2010-2013 Sylvie Boldo
+Copyright (C) 2009-2018 Sylvie Boldo
#<br />#
-Copyright (C) 2010-2013 Guillaume Melquiond
+Copyright (C) 2009-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
@@ -18,20 +18,20 @@ COPYING file for more details.
*)
(** * Basic definitions: float and rounding property *)
-Require Import Fcore_Raux.
+Require Import Raux.
Section Def.
(** Definition of a floating-point number *)
Record float (beta : radix) := Float { Fnum : Z ; Fexp : Z }.
-Arguments Fnum {beta} f.
-Arguments Fexp {beta} f.
+Arguments Fnum {beta}.
+Arguments Fexp {beta}.
Variable beta : radix.
Definition F2R (f : float beta) :=
- (Z2R (Fnum f) * bpow beta (Fexp f))%R.
+ (IZR (Fnum f) * bpow beta (Fexp f))%R.
(** Requirements on a rounding mode *)
Definition round_pred_total (P : R -> R -> Prop) :=
@@ -46,9 +46,9 @@ Definition round_pred (P : R -> R -> Prop) :=
End Def.
-Arguments Fnum {beta} f.
-Arguments Fexp {beta} f.
-Arguments F2R {beta} f.
+Arguments Fnum {beta}.
+Arguments Fexp {beta}.
+Arguments F2R {beta}.
Section RND.
@@ -57,45 +57,27 @@ Definition Rnd_DN_pt (F : R -> Prop) (x f : R) :=
F f /\ (f <= x)%R /\
forall g : R, F g -> (g <= x)%R -> (g <= f)%R.
-Definition Rnd_DN (F : R -> Prop) (rnd : R -> R) :=
- forall x : R, Rnd_DN_pt F x (rnd x).
-
(** property of being a round toward +inf *)
Definition Rnd_UP_pt (F : R -> Prop) (x f : R) :=
F f /\ (x <= f)%R /\
forall g : R, F g -> (x <= g)%R -> (f <= g)%R.
-Definition Rnd_UP (F : R -> Prop) (rnd : R -> R) :=
- forall x : R, Rnd_UP_pt F x (rnd x).
-
(** property of being a round toward zero *)
Definition Rnd_ZR_pt (F : R -> Prop) (x f : R) :=
( (0 <= x)%R -> Rnd_DN_pt F x f ) /\
( (x <= 0)%R -> Rnd_UP_pt F x f ).
-Definition Rnd_ZR (F : R -> Prop) (rnd : R -> R) :=
- forall x : R, Rnd_ZR_pt F x (rnd x).
-
(** property of being a round to nearest *)
Definition Rnd_N_pt (F : R -> Prop) (x f : R) :=
F f /\
forall g : R, F g -> (Rabs (f - x) <= Rabs (g - x))%R.
-Definition Rnd_N (F : R -> Prop) (rnd : R -> R) :=
- forall x : R, Rnd_N_pt F x (rnd x).
-
Definition Rnd_NG_pt (F : R -> Prop) (P : R -> R -> Prop) (x f : R) :=
Rnd_N_pt F x f /\
( P x f \/ forall f2 : R, Rnd_N_pt F x f2 -> f2 = f ).
-Definition Rnd_NG (F : R -> Prop) (P : R -> R -> Prop) (rnd : R -> R) :=
- forall x : R, Rnd_NG_pt F P x (rnd x).
-
Definition Rnd_NA_pt (F : R -> Prop) (x f : R) :=
Rnd_N_pt F x f /\
forall f2 : R, Rnd_N_pt F x f2 -> (Rabs f2 <= Rabs f)%R.
-Definition Rnd_NA (F : R -> Prop) (rnd : R -> R) :=
- forall x : R, Rnd_NA_pt F x (rnd x).
-
End RND.
diff --git a/flocq/Core/Fcore_digits.v b/flocq/Core/Digits.v
index 53743035..bed2e20a 100644
--- a/flocq/Core/Fcore_digits.v
+++ b/flocq/Core/Digits.v
@@ -2,9 +2,9 @@
This file is part of the Flocq formalization of floating-point
arithmetic in Coq: http://flocq.gforge.inria.fr/
-Copyright (C) 2011-2013 Sylvie Boldo
+Copyright (C) 2011-2018 Sylvie Boldo
#<br />#
-Copyright (C) 2011-2013 Guillaume Melquiond
+Copyright (C) 2011-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
@@ -17,9 +17,8 @@ MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
COPYING file for more details.
*)
-Require Import ZArith.
-Require Import Zquot.
-Require Import Fcore_Zaux.
+Require Import ZArith Zquot.
+Require Import Zaux.
(** Number of bits (radix 2) of a positive integer.
@@ -74,7 +73,7 @@ Qed.
Theorem Zdigit_opp :
forall n k,
- Zdigit (-n) k = Zopp (Zdigit n k).
+ Zdigit (-n) k = Z.opp (Zdigit n k).
Proof.
intros n k.
unfold Zdigit.
@@ -89,11 +88,11 @@ Theorem Zdigit_ge_Zpower_pos :
Proof.
intros e n Hn k Hk.
unfold Zdigit.
-rewrite Zquot_small.
+rewrite Z.quot_small.
apply Zrem_0_l.
split.
apply Hn.
-apply Zlt_le_trans with (1 := proj2 Hn).
+apply Z.lt_le_trans with (1 := proj2 Hn).
replace k with (e + (k - e))%Z by ring.
rewrite Zpower_plus.
rewrite <- (Zmult_1_r (beta ^ e)) at 1.
@@ -102,8 +101,8 @@ apply (Zlt_le_succ 0).
apply Zpower_gt_0.
now apply Zle_minus_le_0.
apply Zlt_le_weak.
-now apply Zle_lt_trans with n.
-generalize (Zle_lt_trans _ _ _ (proj1 Hn) (proj2 Hn)).
+now apply Z.le_lt_trans with n.
+generalize (Z.le_lt_trans _ _ _ (proj1 Hn) (proj2 Hn)).
clear.
now destruct e as [|e|e].
now apply Zle_minus_le_0.
@@ -111,7 +110,7 @@ Qed.
Theorem Zdigit_ge_Zpower :
forall e n,
- (Zabs n < Zpower beta e)%Z ->
+ (Z.abs n < Zpower beta e)%Z ->
forall k, (e <= k)%Z -> Zdigit n k = Z0.
Proof.
intros e [|n|n] Hn k.
@@ -119,10 +118,10 @@ easy.
apply Zdigit_ge_Zpower_pos.
now split.
intros He.
-change (Zneg n) with (Zopp (Zpos n)).
+change (Zneg n) with (Z.opp (Zpos n)).
rewrite Zdigit_opp.
rewrite Zdigit_ge_Zpower_pos with (2 := He).
-apply Zopp_0.
+apply Z.opp_0.
now split.
Qed.
@@ -134,17 +133,17 @@ Proof.
intros e n He (Hn1,Hn2).
unfold Zdigit.
rewrite <- ZOdiv_mod_mult.
-rewrite Zrem_small.
+rewrite Z.rem_small.
intros H.
apply Zle_not_lt with (1 := Hn1).
rewrite (Z.quot_rem' n (beta ^ e)).
rewrite H, Zmult_0_r, Zplus_0_l.
apply Zrem_lt_pos_pos.
-apply Zle_trans with (2 := Hn1).
+apply Z.le_trans with (2 := Hn1).
apply Zpower_ge_0.
now apply Zpower_gt_0.
split.
-apply Zle_trans with (2 := Hn1).
+apply Z.le_trans with (2 := Hn1).
apply Zpower_ge_0.
replace (beta ^ e * beta)%Z with (beta ^ (e + 1))%Z.
exact Hn2.
@@ -154,12 +153,12 @@ Qed.
Theorem Zdigit_not_0 :
forall e n, (0 <= e)%Z ->
- (Zpower beta e <= Zabs n < Zpower beta (e + 1))%Z ->
+ (Zpower beta e <= Z.abs n < Zpower beta (e + 1))%Z ->
Zdigit n e <> Z0.
Proof.
intros e n He Hn.
destruct (Zle_or_lt 0 n) as [Hn'|Hn'].
-rewrite (Zabs_eq _ Hn') in Hn.
+rewrite (Z.abs_eq _ Hn') in Hn.
now apply Zdigit_not_0_pos.
intros H.
rewrite (Zabs_non_eq n) in Hn by now apply Zlt_le_weak.
@@ -245,8 +244,8 @@ intros n k k' Hk.
unfold Zdigit.
rewrite ZOdiv_small_abs.
apply Zrem_0_l.
-apply Zlt_le_trans with (Zpower beta k').
-rewrite <- (Zabs_eq (beta ^ k')) at 2 by apply Zpower_ge_0.
+apply Z.lt_le_trans with (Zpower beta k').
+rewrite <- (Z.abs_eq (beta ^ k')) at 2 by apply Zpower_ge_0.
apply Zrem_lt.
apply Zgt_not_eq.
now apply Zpower_gt_0.
@@ -266,7 +265,7 @@ Proof.
intros n.
induction k.
apply sym_eq.
-apply Zrem_1_r.
+apply Z.rem_1_r.
simpl Zsum_digit.
rewrite IHk.
unfold Zdigit.
@@ -284,65 +283,35 @@ apply Zle_0_nat.
easy.
Qed.
-Theorem Zpower_gt_id :
- forall n, (n < Zpower beta n)%Z.
-Proof.
-intros [|n|n] ; try easy.
-simpl.
-rewrite Zpower_pos_nat.
-rewrite Zpos_eq_Z_of_nat_o_nat_of_P.
-induction (nat_of_P n).
-easy.
-rewrite inj_S.
-change (Zpower_nat beta (S n0)) with (beta * Zpower_nat beta n0)%Z.
-unfold Zsucc.
-apply Zlt_le_trans with (beta * (Z_of_nat n0 + 1))%Z.
-clear.
-apply Zlt_0_minus_lt.
-replace (beta * (Z_of_nat n0 + 1) - (Z_of_nat n0 + 1))%Z with ((beta - 1) * (Z_of_nat n0 + 1))%Z by ring.
-apply Zmult_lt_0_compat.
-cut (2 <= beta)%Z. omega.
-apply Zle_bool_imp_le.
-apply beta.
-apply (Zle_lt_succ 0).
-apply Zle_0_nat.
-apply Zmult_le_compat_l.
-now apply Zlt_le_succ.
-apply Zle_trans with 2%Z.
-easy.
-apply Zle_bool_imp_le.
-apply beta.
-Qed.
-
Theorem Zdigit_ext :
forall n1 n2,
(forall k, (0 <= k)%Z -> Zdigit n1 k = Zdigit n2 k) ->
n1 = n2.
Proof.
intros n1 n2 H.
-rewrite <- (ZOmod_small_abs n1 (Zpower beta (Zmax (Zabs n1) (Zabs n2)))).
-rewrite <- (ZOmod_small_abs n2 (Zpower beta (Zmax (Zabs n1) (Zabs n2)))) at 2.
-replace (Zmax (Zabs n1) (Zabs n2)) with (Z_of_nat (Zabs_nat (Zmax (Zabs n1) (Zabs n2)))).
+rewrite <- (ZOmod_small_abs n1 (Zpower beta (Z.max (Z.abs n1) (Z.abs n2)))).
+rewrite <- (ZOmod_small_abs n2 (Zpower beta (Z.max (Z.abs n1) (Z.abs n2)))) at 2.
+replace (Z.max (Z.abs n1) (Z.abs n2)) with (Z_of_nat (Z.abs_nat (Z.max (Z.abs n1) (Z.abs n2)))).
rewrite <- 2!Zsum_digit_digit.
-induction (Zabs_nat (Zmax (Zabs n1) (Zabs n2))).
+induction (Z.abs_nat (Z.max (Z.abs n1) (Z.abs n2))).
easy.
simpl.
rewrite H, IHn.
apply refl_equal.
apply Zle_0_nat.
rewrite inj_Zabs_nat.
-apply Zabs_eq.
-apply Zle_trans with (Zabs n1).
+apply Z.abs_eq.
+apply Z.le_trans with (Z.abs n1).
apply Zabs_pos.
-apply Zle_max_l.
-apply Zlt_le_trans with (Zpower beta (Zabs n2)).
+apply Z.le_max_l.
+apply Z.lt_le_trans with (Zpower beta (Z.abs n2)).
apply Zpower_gt_id.
apply Zpower_le.
-apply Zle_max_r.
-apply Zlt_le_trans with (Zpower beta (Zabs n1)).
+apply Z.le_max_r.
+apply Z.lt_le_trans with (Zpower beta (Z.abs n1)).
apply Zpower_gt_id.
apply Zpower_le.
-apply Zle_max_l.
+apply Z.le_max_l.
Qed.
Theorem ZOmod_plus_pow_digit :
@@ -354,11 +323,11 @@ intros u v n Huv Hd.
destruct (Zle_or_lt 0 n) as [Hn|Hn].
rewrite Zplus_rem with (1 := Huv).
apply ZOmod_small_abs.
-generalize (Zle_refl n).
-pattern n at -2 ; rewrite <- Zabs_eq with (1 := Hn).
+generalize (Z.le_refl n).
+pattern n at -2 ; rewrite <- Z.abs_eq with (1 := Hn).
rewrite <- (inj_Zabs_nat n).
-induction (Zabs_nat n) as [|p IHp].
-now rewrite 2!Zrem_1_r.
+induction (Z.abs_nat n) as [|p IHp].
+now rewrite 2!Z.rem_1_r.
rewrite <- 2!Zsum_digit_digit.
simpl Zsum_digit.
rewrite inj_S.
@@ -367,39 +336,39 @@ replace (Zsum_digit (Zdigit u) p + Zdigit u (Z_of_nat p) * beta ^ Z_of_nat p +
(Zsum_digit (Zdigit v) p + Zdigit v (Z_of_nat p) * beta ^ Z_of_nat p))%Z with
(Zsum_digit (Zdigit u) p + Zsum_digit (Zdigit v) p +
(Zdigit u (Z_of_nat p) + Zdigit v (Z_of_nat p)) * beta ^ Z_of_nat p)%Z by ring.
-apply (Zle_lt_trans _ _ _ (Zabs_triangle _ _)).
-replace (beta ^ Zsucc (Z_of_nat p))%Z with (beta ^ Z_of_nat p + (beta - 1) * beta ^ Z_of_nat p)%Z.
+apply (Z.le_lt_trans _ _ _ (Z.abs_triangle _ _)).
+replace (beta ^ Z.succ (Z_of_nat p))%Z with (beta ^ Z_of_nat p + (beta - 1) * beta ^ Z_of_nat p)%Z.
apply Zplus_lt_le_compat.
rewrite 2!Zsum_digit_digit.
apply IHp.
now apply Zle_succ_le.
rewrite Zabs_Zmult.
-rewrite (Zabs_eq (beta ^ Z_of_nat p)) by apply Zpower_ge_0.
+rewrite (Z.abs_eq (beta ^ Z_of_nat p)) by apply Zpower_ge_0.
apply Zmult_le_compat_r. 2: apply Zpower_ge_0.
apply Zlt_succ_le.
-assert (forall u v, Zabs (Zdigit u v) < Zsucc (beta - 1))%Z.
+assert (forall u v, Z.abs (Zdigit u v) < Z.succ (beta - 1))%Z.
clear ; intros n k.
assert (0 < beta)%Z.
-apply Zlt_le_trans with 2%Z.
+apply Z.lt_le_trans with 2%Z.
apply refl_equal.
apply Zle_bool_imp_le.
apply beta.
-replace (Zsucc (beta - 1)) with (Zabs beta).
+replace (Z.succ (beta - 1)) with (Z.abs beta).
apply Zrem_lt.
now apply Zgt_not_eq.
-rewrite Zabs_eq.
+rewrite Z.abs_eq.
apply Zsucc_pred.
now apply Zlt_le_weak.
assert (0 <= Z_of_nat p < n)%Z.
split.
apply Zle_0_nat.
-apply Zgt_lt.
+apply Z.gt_lt.
now apply Zle_succ_gt.
destruct (Hd (Z_of_nat p) H0) as [K|K] ; rewrite K.
apply H.
rewrite Zplus_0_r.
apply H.
-unfold Zsucc.
+unfold Z.succ.
ring_simplify.
rewrite Zpower_plus.
change (beta ^1)%Z with (beta * 1)%Z.
@@ -422,7 +391,7 @@ rewrite <- ZOmod_plus_pow_digit by assumption.
apply f_equal.
destruct (Zle_or_lt 0 n) as [Hn|Hn].
apply ZOdiv_small_abs.
-rewrite <- Zabs_eq.
+rewrite <- Z.abs_eq.
apply Zrem_lt.
apply Zgt_not_eq.
now apply Zpower_gt_0.
@@ -562,7 +531,7 @@ rewrite Zle_bool_true.
rewrite Zdigit_mod_pow by apply Hk.
rewrite Zdigit_scale by apply Hk.
unfold Zminus.
-now rewrite Zopp_involutive, Zplus_comm.
+now rewrite Z.opp_involutive, Zplus_comm.
omega.
Qed.
@@ -608,13 +577,13 @@ Qed.
Theorem Zslice_slice :
forall n k1 k2 k1' k2', (0 <= k1' <= k2)%Z ->
- Zslice (Zslice n k1 k2) k1' k2' = Zslice n (k1 + k1') (Zmin (k2 - k1') k2').
+ Zslice (Zslice n k1 k2) k1' k2' = Zslice n (k1 + k1') (Z.min (k2 - k1') k2').
Proof.
intros n k1 k2 k1' k2' Hk1'.
destruct (Zle_or_lt 0 k2') as [Hk2'|Hk2'].
apply Zdigit_ext.
intros k Hk.
-destruct (Zle_or_lt (Zmin (k2 - k1') k2') k) as [Hk'|Hk'].
+destruct (Zle_or_lt (Z.min (k2 - k1') k2') k) as [Hk'|Hk'].
rewrite (Zdigit_slice_out n (k1 + k1')) with (1 := Hk').
destruct (Zle_or_lt k2' k) as [Hk''|Hk''].
now apply Zdigit_slice_out.
@@ -627,7 +596,7 @@ rewrite Zdigit_slice.
now rewrite Zplus_assoc.
zify ; omega.
unfold Zslice.
-rewrite Zmin_r.
+rewrite Z.min_r.
now rewrite Zle_bool_false.
omega.
Qed.
@@ -659,11 +628,11 @@ replace k1 with Z0 by omega.
case Zle_bool_spec ; intros Hk'.
replace k with Z0 by omega.
simpl.
-now rewrite Zquot_1_r.
-rewrite Zopp_involutive.
+now rewrite Z.quot_1_r.
+rewrite Z.opp_involutive.
apply Zmult_1_r.
rewrite Zle_bool_false by omega.
-rewrite 2!Zopp_involutive, Zplus_comm.
+rewrite 2!Z.opp_involutive, Zplus_comm.
rewrite Zpower_plus by assumption.
apply Zquot_Zquot.
Qed.
@@ -689,7 +658,7 @@ apply Zdigit_ext.
intros k' Hk'.
rewrite Zdigit_scale with (1 := Hk').
unfold Zminus.
-rewrite (Zplus_comm k'), Zopp_involutive.
+rewrite (Zplus_comm k'), Z.opp_involutive.
destruct (Zle_or_lt k2 k') as [Hk2|Hk2].
rewrite Zdigit_slice_out with (1 := Hk2).
apply sym_eq.
@@ -770,7 +739,7 @@ Definition Zdigits n :=
Theorem Zdigits_correct :
forall n,
- (Zpower beta (Zdigits n - 1) <= Zabs n < Zpower beta (Zdigits n))%Z.
+ (Zpower beta (Zdigits n - 1) <= Z.abs n < Zpower beta (Zdigits n))%Z.
Proof.
cut (forall p, Zpower beta (Zdigits (Zpos p) - 1) <= Zpos p < Zpower beta (Zdigits (Zpos p)))%Z.
intros H [|n|n] ; try exact (H n).
@@ -779,7 +748,7 @@ intros n.
simpl.
(* *)
assert (U: (Zpos n < Zpower beta (Z_of_nat (S (digits2_Pnat n))))%Z).
-apply Zlt_le_trans with (1 := proj2 (digits2_Pnat_correct n)).
+apply Z.lt_le_trans with (1 := proj2 (digits2_Pnat_correct n)).
rewrite Zpower_Zpower_nat.
rewrite Zabs_nat_Z_of_nat.
induction (S (digits2_Pnat n)).
@@ -797,7 +766,7 @@ apply Zle_0_nat.
(* *)
revert U.
rewrite inj_S.
-unfold Zsucc.
+unfold Z.succ.
generalize (digits2_Pnat n).
intros u U.
pattern (radix_val beta) at 2 4 ; replace (radix_val beta) with (Zpower beta 1) by apply Zmult_1_r.
@@ -805,12 +774,12 @@ assert (V: (Zpower beta (1 - 1) <= Zpos n)%Z).
now apply (Zlt_le_succ 0).
generalize (conj V U).
clear.
-generalize (Zle_refl 1).
+generalize (Z.le_refl 1).
generalize 1%Z at 2 3 5 6 7 9 10.
(* *)
induction u.
easy.
-rewrite inj_S; unfold Zsucc.
+rewrite inj_S; unfold Z.succ.
simpl Zdigits_aux.
intros v Hv U.
case Zlt_bool_spec ; intros K.
@@ -829,20 +798,20 @@ Qed.
Theorem Zdigits_unique :
forall n d,
- (Zpower beta (d - 1) <= Zabs n < Zpower beta d)%Z ->
+ (Zpower beta (d - 1) <= Z.abs n < Zpower beta d)%Z ->
Zdigits n = d.
Proof.
intros n d Hd.
assert (Hd' := Zdigits_correct n).
apply Zle_antisym.
apply (Zpower_lt_Zpower beta).
-now apply Zle_lt_trans with (Zabs n).
+now apply Z.le_lt_trans with (Z.abs n).
apply (Zpower_lt_Zpower beta).
-now apply Zle_lt_trans with (Zabs n).
+now apply Z.le_lt_trans with (Z.abs n).
Qed.
Theorem Zdigits_abs :
- forall n, Zdigits (Zabs n) = Zdigits n.
+ forall n, Zdigits (Z.abs n) = Zdigits n.
Proof.
now intros [|n|n].
Qed.
@@ -852,10 +821,10 @@ Theorem Zdigits_gt_0 :
Proof.
intros n Zn.
rewrite <- (Zdigits_abs n).
-assert (Hn: (0 < Zabs n)%Z).
+assert (Hn: (0 < Z.abs n)%Z).
destruct n ; [|easy|easy].
now elim Zn.
-destruct (Zabs n) as [|p|p] ; try easy ; clear.
+destruct (Z.abs n) as [|p|p] ; try easy ; clear.
simpl.
generalize 1%Z (radix_val beta) (refl_equal Lt : (0 < 1)%Z).
induction (digits2_Pnat p).
@@ -872,7 +841,7 @@ Theorem Zdigits_ge_0 :
forall n, (0 <= Zdigits n)%Z.
Proof.
intros n.
-destruct (Z_eq_dec n 0) as [H|H].
+destruct (Z.eq_dec n 0) as [H|H].
now rewrite H.
apply Zlt_le_weak.
now apply Zdigits_gt_0.
@@ -908,8 +877,8 @@ unfold Zslice.
rewrite Zle_bool_true with (1 := Hl).
destruct (Zdigits_correct (Z.rem (Zscale n (- k)) (Zpower beta l))) as (H1,H2).
apply Zpower_lt_Zpower with beta.
-apply Zle_lt_trans with (1 := H1).
-rewrite <- (Zabs_eq (beta ^ l)) at 2 by apply Zpower_ge_0.
+apply Z.le_lt_trans with (1 := H1).
+rewrite <- (Z.abs_eq (beta ^ l)) at 2 by apply Zpower_ge_0.
apply Zrem_lt.
apply Zgt_not_eq.
now apply Zpower_gt_0.
@@ -923,7 +892,7 @@ Proof.
intros m e Hm He.
assert (H := Zdigits_correct m).
apply Zdigits_unique.
-rewrite Z.abs_mul, Z.abs_pow, (Zabs_eq beta).
+rewrite Z.abs_mul, Z.abs_pow, (Z.abs_eq beta).
2: now apply Zlt_le_weak, radix_gt_0.
split.
replace (Zdigits m + e - 1)%Z with (Zdigits m - 1 + e)%Z by ring.
@@ -976,18 +945,18 @@ Qed.
Theorem Zpower_le_Zdigits :
forall e x,
(e < Zdigits x)%Z ->
- (Zpower beta e <= Zabs x)%Z.
+ (Zpower beta e <= Z.abs x)%Z.
Proof.
intros e x Hex.
destruct (Zdigits_correct x) as [H1 H2].
-apply Zle_trans with (2 := H1).
+apply Z.le_trans with (2 := H1).
apply Zpower_le.
clear -Hex ; omega.
Qed.
Theorem Zdigits_le_Zpower :
forall e x,
- (Zabs x < Zpower beta e)%Z ->
+ (Z.abs x < Zpower beta e)%Z ->
(Zdigits x <= e)%Z.
Proof.
intros e x.
@@ -998,17 +967,17 @@ Qed.
Theorem Zpower_gt_Zdigits :
forall e x,
(Zdigits x <= e)%Z ->
- (Zabs x < Zpower beta e)%Z.
+ (Z.abs x < Zpower beta e)%Z.
Proof.
intros e x Hex.
destruct (Zdigits_correct x) as [H1 H2].
-apply Zlt_le_trans with (1 := H2).
+apply Z.lt_le_trans with (1 := H2).
now apply Zpower_le.
Qed.
Theorem Zdigits_gt_Zpower :
forall e x,
- (Zpower beta e <= Zabs x)%Z ->
+ (Zpower beta e <= Z.abs x)%Z ->
(e < Zdigits x)%Z.
Proof.
intros e x Hex.
@@ -1029,17 +998,17 @@ Theorem Zdigits_mult_strong :
Proof.
intros x y Hx Hy.
apply Zdigits_le_Zpower.
-rewrite Zabs_eq.
-apply Zlt_le_trans with ((x + 1) * (y + 1))%Z.
+rewrite Z.abs_eq.
+apply Z.lt_le_trans with ((x + 1) * (y + 1))%Z.
ring_simplify.
-apply Zle_lt_succ, Zle_refl.
+apply Zle_lt_succ, Z.le_refl.
rewrite Zpower_plus by apply Zdigits_ge_0.
apply Zmult_le_compat.
apply Zlt_le_succ.
-rewrite <- (Zabs_eq x) at 1 by easy.
+rewrite <- (Z.abs_eq x) at 1 by easy.
apply Zdigits_correct.
apply Zlt_le_succ.
-rewrite <- (Zabs_eq y) at 1 by easy.
+rewrite <- (Z.abs_eq y) at 1 by easy.
apply Zdigits_correct.
clear -Hx ; omega.
clear -Hy ; omega.
@@ -1057,7 +1026,7 @@ intros x y.
rewrite <- Zdigits_abs.
rewrite <- (Zdigits_abs x).
rewrite <- (Zdigits_abs y).
-apply Zle_trans with (Zdigits (Zabs x + Zabs y + Zabs x * Zabs y)).
+apply Z.le_trans with (Zdigits (Z.abs x + Z.abs y + Z.abs x * Z.abs y)).
apply Zdigits_le.
apply Zabs_pos.
rewrite Zabs_Zmult.
@@ -1097,28 +1066,28 @@ intros m e Hm He.
assert (H := Zdigits_correct m).
apply Zdigits_unique.
destruct (Zle_lt_or_eq _ _ (proj2 He)) as [He'|He'].
- rewrite Zabs_eq in H by easy.
+ rewrite Z.abs_eq in H by easy.
destruct H as [H1 H2].
- rewrite Zabs_eq.
+ rewrite Z.abs_eq.
split.
replace (Zdigits m - e - 1)%Z with (Zdigits m - 1 - e)%Z by ring.
rewrite Z.pow_sub_r.
2: apply Zgt_not_eq, radix_gt_0.
2: clear -He He' ; omega.
apply Z_div_le with (2 := H1).
- now apply Zlt_gt, Zpower_gt_0.
+ now apply Z.lt_gt, Zpower_gt_0.
apply Zmult_lt_reg_r with (Zpower beta e).
now apply Zpower_gt_0.
- apply Zle_lt_trans with m.
+ apply Z.le_lt_trans with m.
rewrite Zmult_comm.
apply Z_mult_div_ge.
- now apply Zlt_gt, Zpower_gt_0.
+ now apply Z.lt_gt, Zpower_gt_0.
rewrite <- Zpower_plus.
now replace (Zdigits m - e + e)%Z with (Zdigits m) by ring.
now apply Zle_minus_le_0.
apply He.
apply Z_div_pos with (2 := Hm).
- now apply Zlt_gt, Zpower_gt_0.
+ now apply Z.lt_gt, Zpower_gt_0.
rewrite He'.
rewrite (Zeq_minus _ (Zdigits m)) by reflexivity.
simpl.
@@ -1126,7 +1095,7 @@ rewrite Zdiv_small.
easy.
split.
exact Hm.
-now rewrite <- (Zabs_eq m) at 1.
+now rewrite <- (Z.abs_eq m) at 1.
Qed.
End Fcore_digits.
@@ -1143,7 +1112,7 @@ intros m.
apply eq_sym, Zdigits_unique.
rewrite <- Zpower_nat_Z.
rewrite Nat2Z.inj_succ.
-change (_ - 1)%Z with (Zpred (Zsucc (Z.of_nat (digits2_Pnat m)))).
+change (_ - 1)%Z with (Z.pred (Z.succ (Z.of_nat (digits2_Pnat m)))).
rewrite <- Zpred_succ.
rewrite <- Zpower_nat_Z.
apply digits2_Pnat_correct.
@@ -1152,8 +1121,8 @@ Qed.
Fixpoint digits2_pos (n : positive) : positive :=
match n with
| xH => xH
- | xO p => Psucc (digits2_pos p)
- | xI p => Psucc (digits2_pos p)
+ | xO p => Pos.succ (digits2_pos p)
+ | xI p => Pos.succ (digits2_pos p)
end.
Theorem Zpos_digits2_pos :
diff --git a/flocq/Core/Fcore_FIX.v b/flocq/Core/FIX.v
index e224a64a..4e0a25e6 100644
--- a/flocq/Core/Fcore_FIX.v
+++ b/flocq/Core/FIX.v
@@ -2,9 +2,9 @@
This file is part of the Flocq formalization of floating-point
arithmetic in Coq: http://flocq.gforge.inria.fr/
-Copyright (C) 2010-2013 Sylvie Boldo
+Copyright (C) 2009-2018 Sylvie Boldo
#<br />#
-Copyright (C) 2010-2013 Guillaume Melquiond
+Copyright (C) 2009-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
@@ -18,12 +18,7 @@ COPYING file for more details.
*)
(** * Fixed-point format *)
-Require Import Fcore_Raux.
-Require Import Fcore_defs.
-Require Import Fcore_rnd.
-Require Import Fcore_generic_fmt.
-Require Import Fcore_ulp.
-Require Import Fcore_rnd_ne.
+Require Import Raux Defs Round_pred Generic_fmt Ulp Round_NE.
Section RND_FIX.
@@ -33,10 +28,9 @@ Notation bpow := (bpow beta).
Variable emin : Z.
-(* fixed-point format with exponent emin *)
-Definition FIX_format (x : R) :=
- exists f : float beta,
- x = F2R f /\ (Fexp f = emin)%Z.
+Inductive FIX_format (x : R) : Prop :=
+ FIX_spec (f : float beta) :
+ x = F2R f -> (Fexp f = emin)%Z -> FIX_format x.
Definition FIX_exp (e : Z) := emin.
@@ -49,16 +43,16 @@ unfold FIX_exp.
split ; intros H.
now apply Zlt_le_weak.
split.
-apply Zle_refl.
+apply Z.le_refl.
now intros _ _.
Qed.
Theorem generic_format_FIX :
forall x, FIX_format x -> generic_format beta FIX_exp x.
Proof.
-intros x ((xm, xe), (Hx1, Hx2)).
+intros x [[xm xe] Hx1 Hx2].
rewrite Hx1.
-now apply generic_format_canonic.
+now apply generic_format_canonical.
Qed.
Theorem FIX_format_generic :
@@ -82,10 +76,11 @@ Qed.
Global Instance FIX_exp_monotone : Monotone_exp FIX_exp.
Proof.
intros ex ey H.
-apply Zle_refl.
+apply Z.le_refl.
Qed.
-Theorem ulp_FIX: forall x, ulp beta FIX_exp x = bpow emin.
+Theorem ulp_FIX :
+ forall x, ulp beta FIX_exp x = bpow emin.
Proof.
intros x; unfold ulp.
case Req_bool_spec; intros Zx.
@@ -96,5 +91,4 @@ intros n _; reflexivity.
reflexivity.
Qed.
-
End RND_FIX.
diff --git a/flocq/Core/Fcore_FLT.v b/flocq/Core/FLT.v
index 2258b1d9..bd48d4b7 100644
--- a/flocq/Core/Fcore_FLT.v
+++ b/flocq/Core/FLT.v
@@ -2,9 +2,9 @@
This file is part of the Flocq formalization of floating-point
arithmetic in Coq: http://flocq.gforge.inria.fr/
-Copyright (C) 2010-2013 Sylvie Boldo
+Copyright (C) 2009-2018 Sylvie Boldo
#<br />#
-Copyright (C) 2010-2013 Guillaume Melquiond
+Copyright (C) 2009-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
@@ -18,15 +18,9 @@ COPYING file for more details.
*)
(** * Floating-point format with gradual 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_FLX.
-Require Import Fcore_FIX.
-Require Import Fcore_ulp.
-Require Import Fcore_rnd_ne.
+Require Import Raux Defs Round_pred Generic_fmt Float_prop.
+Require Import FLX FIX Ulp Round_NE.
+Require Import Psatz.
Section RND_FLT.
@@ -38,12 +32,12 @@ Variable emin prec : Z.
Context { prec_gt_0_ : Prec_gt_0 prec }.
-(* floating-point format with gradual underflow *)
-Definition FLT_format (x : R) :=
- exists f : float beta,
- x = F2R f /\ (Zabs (Fnum f) < Zpower beta prec)%Z /\ (emin <= Fexp f)%Z.
+Inductive FLT_format (x : R) : Prop :=
+ FLT_spec (f : float beta) :
+ x = F2R f -> (Z.abs (Fnum f) < Zpower beta prec)%Z ->
+ (emin <= Fexp f)%Z -> FLT_format x.
-Definition FLT_exp e := Zmax (e - prec) emin.
+Definition FLT_exp e := Z.max (e - prec) emin.
(** Properties of the FLT format *)
Global Instance FLT_exp_valid : Valid_exp FLT_exp.
@@ -59,17 +53,17 @@ Theorem generic_format_FLT :
forall x, FLT_format x -> generic_format beta FLT_exp x.
Proof.
clear prec_gt_0_.
-intros x ((mx, ex), (H1, (H2, H3))).
+intros x [[mx ex] H1 H2 H3].
simpl in H2, H3.
rewrite H1.
apply generic_format_F2R.
intros Zmx.
-unfold canonic_exp, FLT_exp.
-rewrite ln_beta_F2R with (1 := Zmx).
-apply Zmax_lub with (2 := H3).
+unfold cexp, FLT_exp.
+rewrite mag_F2R with (1 := Zmx).
+apply Z.max_lub with (2 := H3).
apply Zplus_le_reg_r with (prec - ex)%Z.
ring_simplify.
-now apply ln_beta_le_Zpower.
+now apply mag_le_Zpower.
Qed.
Theorem FLT_format_generic :
@@ -77,32 +71,32 @@ Theorem FLT_format_generic :
Proof.
intros x.
unfold generic_format.
-set (ex := canonic_exp beta FLT_exp x).
+set (ex := cexp beta FLT_exp x).
set (mx := Ztrunc (scaled_mantissa beta FLT_exp x)).
intros Hx.
rewrite Hx.
eexists ; repeat split ; simpl.
-apply lt_Z2R.
-rewrite Z2R_Zpower. 2: now apply Zlt_le_weak.
+apply lt_IZR.
+rewrite IZR_Zpower. 2: now apply Zlt_le_weak.
apply Rmult_lt_reg_r with (bpow ex).
apply bpow_gt_0.
rewrite <- bpow_plus.
-change (F2R (Float beta (Zabs mx) ex) < bpow (prec + ex))%R.
+change (F2R (Float beta (Z.abs mx) ex) < bpow (prec + ex))%R.
rewrite F2R_Zabs.
rewrite <- Hx.
destruct (Req_dec x 0) as [Hx0|Hx0].
rewrite Hx0, Rabs_R0.
apply bpow_gt_0.
-unfold canonic_exp in ex.
-destruct (ln_beta beta x) as (ex', He).
+unfold cexp in ex.
+destruct (mag beta x) as (ex', He).
simpl in ex.
specialize (He Hx0).
apply Rlt_le_trans with (1 := proj2 He).
apply bpow_le.
cut (ex' - prec <= ex)%Z. omega.
unfold ex, FLT_exp.
-apply Zle_max_l.
-apply Zle_max_r.
+apply Z.le_max_l.
+apply Z.le_max_r.
Qed.
@@ -128,18 +122,18 @@ apply FLT_format_generic.
apply generic_format_FLT.
Qed.
-Theorem canonic_exp_FLT_FLX :
+Theorem cexp_FLT_FLX :
forall x,
(bpow (emin + prec - 1) <= Rabs x)%R ->
- canonic_exp beta FLT_exp x = canonic_exp beta (FLX_exp prec) x.
+ cexp beta FLT_exp x = cexp beta (FLX_exp prec) x.
Proof.
intros x Hx.
assert (Hx0: x <> 0%R).
intros H1; rewrite H1, Rabs_R0 in Hx.
contradict Hx; apply Rlt_not_le, bpow_gt_0.
-unfold canonic_exp.
+unfold cexp.
apply Zmax_left.
-destruct (ln_beta beta x) as (ex, He).
+destruct (mag beta x) as (ex, He).
unfold FLX_exp. simpl.
specialize (He Hx0).
cut (emin + prec - 1 < ex)%Z. omega.
@@ -160,7 +154,7 @@ destruct (Req_dec x 0) as [Hx0|Hx0].
rewrite Hx0.
apply generic_format_0.
unfold generic_format, scaled_mantissa.
-now rewrite canonic_exp_FLT_FLX.
+now rewrite cexp_FLT_FLX.
Qed.
Theorem generic_format_FLX_FLT :
@@ -173,29 +167,30 @@ unfold generic_format in Hx; rewrite Hx.
apply generic_format_F2R.
intros _.
rewrite <- Hx.
-unfold canonic_exp, FLX_exp, FLT_exp.
-apply Zle_max_l.
+unfold cexp, FLX_exp, FLT_exp.
+apply Z.le_max_l.
Qed.
Theorem round_FLT_FLX : forall rnd x,
(bpow (emin + prec - 1) <= Rabs x)%R ->
round beta FLT_exp rnd x = round beta (FLX_exp prec) rnd x.
+Proof.
intros rnd x Hx.
unfold round, scaled_mantissa.
-rewrite canonic_exp_FLT_FLX ; trivial.
+rewrite cexp_FLT_FLX ; trivial.
Qed.
(** Links between FLT and FIX (underflow) *)
-Theorem canonic_exp_FLT_FIX :
+Theorem cexp_FLT_FIX :
forall x, x <> 0%R ->
(Rabs x < bpow (emin + prec))%R ->
- canonic_exp beta FLT_exp x = canonic_exp beta (FIX_exp emin) x.
+ cexp beta FLT_exp x = cexp beta (FIX_exp emin) x.
Proof.
intros x Hx0 Hx.
-unfold canonic_exp.
+unfold cexp.
apply Zmax_right.
unfold FIX_exp.
-destruct (ln_beta beta x) as (ex, Hex).
+destruct (mag beta x) as (ex, Hex).
simpl.
cut (ex - 1 < emin + prec)%Z. omega.
apply (lt_bpow beta).
@@ -214,7 +209,7 @@ rewrite Hx.
apply generic_format_F2R.
intros _.
rewrite <- Hx.
-apply Zle_max_r.
+apply Z.le_max_r.
Qed.
Theorem generic_format_FLT_FIX :
@@ -226,9 +221,37 @@ Proof with auto with typeclass_instances.
apply generic_inclusion_le...
intros e He.
unfold FIX_exp.
-apply Zmax_lub.
+apply Z.max_lub.
omega.
-apply Zle_refl.
+apply Z.le_refl.
+Qed.
+
+Lemma negligible_exp_FLT :
+ exists n, negligible_exp FLT_exp = Some n /\ (n <= emin)%Z.
+Proof.
+case (negligible_exp_spec FLT_exp).
+{ intro H; exfalso; specialize (H emin); revert H.
+ apply Zle_not_lt, Z.le_max_r. }
+intros n Hn; exists n; split; [now simpl|].
+destruct (Z.max_spec (n - prec) emin) as [(Hm, Hm')|(Hm, Hm')].
+{ now revert Hn; unfold FLT_exp; rewrite Hm'. }
+revert Hn prec_gt_0_; unfold FLT_exp, Prec_gt_0; rewrite Hm'; lia.
+Qed.
+
+Theorem generic_format_FLT_1 (Hemin : (emin <= 0)%Z) :
+ generic_format beta FLT_exp 1.
+Proof.
+unfold generic_format, scaled_mantissa, cexp, F2R; simpl.
+rewrite Rmult_1_l, (mag_unique beta 1 1).
+{ unfold FLT_exp.
+ destruct (Z.max_spec_le (1 - prec) emin) as [(H,Hm)|(H,Hm)]; rewrite Hm;
+ (rewrite <- IZR_Zpower; [|unfold Prec_gt_0 in prec_gt_0_; omega]);
+ (rewrite Ztrunc_IZR, IZR_Zpower, <-bpow_plus;
+ [|unfold Prec_gt_0 in prec_gt_0_; omega]);
+ now replace (_ + _)%Z with Z0 by ring. }
+rewrite Rabs_R1; simpl; split; [now right|].
+rewrite IZR_Zpower_pos; simpl; rewrite Rmult_1_r; apply IZR_lt.
+apply (Z.lt_le_trans _ 2); [omega|]; apply Zle_bool_imp_le, beta.
Qed.
Theorem ulp_FLT_small: forall x, (Rabs x < bpow (emin+prec))%R ->
@@ -240,7 +263,7 @@ unfold ulp; case Req_bool_spec; intros Hx2.
case (negligible_exp_spec FLT_exp).
intros T; specialize (T (emin-1)%Z); contradict T.
apply Zle_not_lt; unfold FLT_exp.
-apply Zle_trans with (2:=Z.le_max_r _ _); omega.
+apply Z.le_trans with (2:=Z.le_max_r _ _); omega.
assert (V:FLT_exp emin = emin).
unfold FLT_exp; apply Z.max_r.
unfold Prec_gt_0 in prec_gt_0_; omega.
@@ -248,10 +271,10 @@ intros n H2; rewrite <-V.
apply f_equal, fexp_negligible_exp_eq...
omega.
(* x <> 0 *)
-apply f_equal; unfold canonic_exp, FLT_exp.
+apply f_equal; unfold cexp, FLT_exp.
apply Z.max_r.
-assert (ln_beta beta x-1 < emin+prec)%Z;[idtac|omega].
-destruct (ln_beta beta x) as (e,He); simpl.
+assert (mag beta x-1 < emin+prec)%Z;[idtac|omega].
+destruct (mag beta x) as (e,He); simpl.
apply lt_bpow with beta.
apply Rle_lt_trans with (2:=Hx).
now apply He.
@@ -266,8 +289,8 @@ assert (Zx : (x <> 0)%R).
intros Z; contradict Hx; apply Rgt_not_le, Rlt_gt.
rewrite Z, Rabs_R0; apply bpow_gt_0.
rewrite ulp_neq_0 with (1 := Zx).
-unfold canonic_exp, FLT_exp.
-destruct (ln_beta beta x) as (e,He).
+unfold cexp, FLT_exp.
+destruct (mag beta x) as (e,He).
apply Rle_trans with (bpow (e-1)*bpow (1-prec))%R.
rewrite <- bpow_plus.
right; apply f_equal.
@@ -289,17 +312,68 @@ intros x; case (Req_dec x 0); intros Hx.
rewrite Hx, ulp_FLT_small, Rabs_R0, Rmult_0_l; try apply bpow_gt_0.
rewrite Rabs_R0; apply bpow_gt_0.
rewrite ulp_neq_0; try exact Hx.
-unfold canonic_exp, FLT_exp.
-apply Rlt_le_trans with (bpow (ln_beta beta x)*bpow (-prec))%R.
+unfold cexp, FLT_exp.
+apply Rlt_le_trans with (bpow (mag beta x)*bpow (-prec))%R.
apply Rmult_lt_compat_r.
apply bpow_gt_0.
-now apply bpow_ln_beta_gt.
+now apply bpow_mag_gt.
rewrite <- bpow_plus.
apply bpow_le.
apply Z.le_max_l.
Qed.
+Lemma ulp_FLT_exact_shift :
+ forall x e,
+ (x <> 0)%R ->
+ (emin + prec <= mag beta x)%Z ->
+ (emin + prec - mag beta x <= e)%Z ->
+ (ulp beta FLT_exp (x * bpow e) = ulp beta FLT_exp x * bpow e)%R.
+Proof.
+intros x e Nzx Hmx He.
+unfold ulp; rewrite Req_bool_false;
+ [|now intro H; apply Nzx, (Rmult_eq_reg_r (bpow e));
+ [rewrite Rmult_0_l|apply Rgt_not_eq, Rlt_gt, bpow_gt_0]].
+rewrite (Req_bool_false _ _ Nzx), <- bpow_plus; f_equal; unfold cexp, FLT_exp.
+rewrite (mag_mult_bpow _ _ _ Nzx), !Z.max_l; omega.
+Qed.
+
+Lemma succ_FLT_exact_shift_pos :
+ forall x e,
+ (0 < x)%R ->
+ (emin + prec <= mag beta x)%Z ->
+ (emin + prec - mag beta x <= e)%Z ->
+ (succ beta FLT_exp (x * bpow e) = succ beta FLT_exp x * bpow e)%R.
+Proof.
+intros x e Px Hmx He.
+rewrite succ_eq_pos; [|now apply Rlt_le, Rmult_lt_0_compat, bpow_gt_0].
+rewrite (succ_eq_pos _ _ _ (Rlt_le _ _ Px)).
+now rewrite Rmult_plus_distr_r; f_equal; apply ulp_FLT_exact_shift; [lra| |].
+Qed.
+Lemma succ_FLT_exact_shift :
+ forall x e,
+ (x <> 0)%R ->
+ (emin + prec + 1 <= mag beta x)%Z ->
+ (emin + prec - mag beta x + 1 <= e)%Z ->
+ (succ beta FLT_exp (x * bpow e) = succ beta FLT_exp x * bpow e)%R.
+Proof.
+intros x e Nzx Hmx He.
+destruct (Rle_or_lt 0 x) as [Px|Nx].
+{ now apply succ_FLT_exact_shift_pos; [lra|lia|lia]. }
+unfold succ.
+rewrite Rle_bool_false; [|assert (H := bpow_gt_0 beta e); nra].
+rewrite Rle_bool_false; [|now simpl].
+rewrite Ropp_mult_distr_l_reverse, <-Ropp_mult_distr_l_reverse; f_equal.
+unfold pred_pos.
+rewrite mag_mult_bpow; [|lra].
+replace (_ - 1)%Z with (mag beta (- x) - 1 + e)%Z; [|ring]; rewrite bpow_plus.
+unfold Req_bool; rewrite Rcompare_mult_r; [|now apply bpow_gt_0].
+fold (Req_bool (-x) (bpow (mag beta (-x) - 1))); case Req_bool.
+{ rewrite mag_opp; unfold FLT_exp; do 2 (rewrite Z.max_l; [|lia]).
+ replace (_ - _)%Z with (mag beta x - 1 - prec + e)%Z; [|ring].
+ rewrite bpow_plus; ring. }
+rewrite ulp_FLT_exact_shift; [ring|lra| |]; rewrite mag_opp; lia.
+Qed.
(** FLT is a nice format: it has a monotone exponent... *)
Global Instance FLT_exp_monotone : Monotone_exp FLT_exp.
@@ -310,7 +384,7 @@ zify ; omega.
Qed.
(** and it allows a rounding to nearest, ties to even. *)
-Hypothesis NE_prop : Zeven beta = false \/ (1 < prec)%Z.
+Hypothesis NE_prop : Z.even beta = false \/ (1 < prec)%Z.
Global Instance exists_NE_FLT : Exists_NE beta FLT_exp.
Proof.
diff --git a/flocq/Core/FLX.v b/flocq/Core/FLX.v
new file mode 100644
index 00000000..803d96ef
--- /dev/null
+++ b/flocq/Core/FLX.v
@@ -0,0 +1,362 @@
+(**
+This file is part of the Flocq formalization of floating-point
+arithmetic in Coq: http://flocq.gforge.inria.fr/
+
+Copyright (C) 2009-2018 Sylvie Boldo
+#<br />#
+Copyright (C) 2009-2018 Guillaume Melquiond
+
+This library is free software; you can redistribute it and/or
+modify it under the terms of the GNU Lesser General Public
+License as published by the Free Software Foundation; either
+version 3 of the License, or (at your option) any later version.
+
+This library is distributed in the hope that it will be useful,
+but WITHOUT ANY WARRANTY; without even the implied warranty of
+MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
+COPYING file for more details.
+*)
+
+(** * Floating-point format without underflow *)
+Require Import Raux Defs Round_pred Generic_fmt Float_prop.
+Require Import FIX Ulp Round_NE.
+Require Import Psatz.
+
+Section RND_FLX.
+
+Variable beta : radix.
+
+Notation bpow e := (bpow beta e).
+
+Variable prec : Z.
+
+Class Prec_gt_0 :=
+ prec_gt_0 : (0 < prec)%Z.
+
+Context { prec_gt_0_ : Prec_gt_0 }.
+
+Inductive FLX_format (x : R) : Prop :=
+ FLX_spec (f : float beta) :
+ x = F2R f -> (Z.abs (Fnum f) < Zpower beta prec)%Z -> FLX_format x.
+
+Definition FLX_exp (e : Z) := (e - prec)%Z.
+
+(** Properties of the FLX format *)
+
+Global Instance FLX_exp_valid : Valid_exp FLX_exp.
+Proof.
+intros k.
+unfold FLX_exp.
+generalize prec_gt_0.
+repeat split ; intros ; omega.
+Qed.
+
+Theorem FIX_format_FLX :
+ forall x e,
+ (bpow (e - 1) <= Rabs x <= bpow e)%R ->
+ FLX_format x ->
+ FIX_format beta (e - prec) x.
+Proof.
+clear prec_gt_0_.
+intros x e Hx [[xm xe] H1 H2].
+rewrite H1, (F2R_prec_normalize beta xm xe e prec).
+now eexists.
+exact H2.
+now rewrite <- H1.
+Qed.
+
+Theorem FLX_format_generic :
+ forall x, generic_format beta FLX_exp x -> FLX_format x.
+Proof.
+intros x H.
+rewrite H.
+eexists ; repeat split.
+simpl.
+apply lt_IZR.
+rewrite abs_IZR.
+rewrite <- scaled_mantissa_generic with (1 := H).
+rewrite <- scaled_mantissa_abs.
+apply Rmult_lt_reg_r with (bpow (cexp beta FLX_exp (Rabs x))).
+apply bpow_gt_0.
+rewrite scaled_mantissa_mult_bpow.
+rewrite IZR_Zpower, <- bpow_plus.
+2: now apply Zlt_le_weak.
+unfold cexp, FLX_exp.
+ring_simplify (prec + (mag beta (Rabs x) - prec))%Z.
+rewrite mag_abs.
+destruct (Req_dec x 0) as [Hx|Hx].
+rewrite Hx, Rabs_R0.
+apply bpow_gt_0.
+destruct (mag beta x) as (ex, Ex).
+now apply Ex.
+Qed.
+
+Theorem generic_format_FLX :
+ forall x, FLX_format x -> generic_format beta FLX_exp x.
+Proof.
+clear prec_gt_0_.
+intros x [[mx ex] H1 H2].
+simpl in H2.
+rewrite H1.
+apply generic_format_F2R.
+intros Zmx.
+unfold cexp, FLX_exp.
+rewrite mag_F2R with (1 := Zmx).
+apply Zplus_le_reg_r with (prec - ex)%Z.
+ring_simplify.
+now apply mag_le_Zpower.
+Qed.
+
+Theorem FLX_format_satisfies_any :
+ satisfies_any FLX_format.
+Proof.
+refine (satisfies_any_eq _ _ _ (generic_format_satisfies_any beta FLX_exp)).
+intros x.
+split.
+apply FLX_format_generic.
+apply generic_format_FLX.
+Qed.
+
+Theorem FLX_format_FIX :
+ forall x e,
+ (bpow (e - 1) <= Rabs x <= bpow e)%R ->
+ FIX_format beta (e - prec) x ->
+ FLX_format x.
+Proof with auto with typeclass_instances.
+intros x e Hx Fx.
+apply FLX_format_generic.
+apply generic_format_FIX in Fx.
+revert Fx.
+apply generic_inclusion with (e := e)...
+apply Z.le_refl.
+Qed.
+
+(** unbounded floating-point format with normal mantissas *)
+Inductive FLXN_format (x : R) : Prop :=
+ FLXN_spec (f : float beta) :
+ x = F2R f ->
+ (x <> 0%R -> Zpower beta (prec - 1) <= Z.abs (Fnum f) < Zpower beta prec)%Z ->
+ FLXN_format x.
+
+Theorem generic_format_FLXN :
+ forall x, FLXN_format x -> generic_format beta FLX_exp x.
+Proof.
+intros x [[xm ex] H1 H2].
+destruct (Req_dec x 0) as [Zx|Zx].
+rewrite Zx.
+apply generic_format_0.
+specialize (H2 Zx).
+apply generic_format_FLX.
+rewrite H1.
+eexists ; repeat split.
+apply H2.
+Qed.
+
+Theorem FLXN_format_generic :
+ forall x, generic_format beta FLX_exp x -> FLXN_format x.
+Proof.
+intros x Hx.
+rewrite Hx.
+simpl.
+eexists. easy.
+rewrite <- Hx.
+intros Zx.
+simpl.
+split.
+(* *)
+apply le_IZR.
+rewrite IZR_Zpower.
+2: now apply Zlt_0_le_0_pred.
+rewrite abs_IZR, <- scaled_mantissa_generic with (1 := Hx).
+apply Rmult_le_reg_r with (bpow (cexp beta FLX_exp x)).
+apply bpow_gt_0.
+rewrite <- bpow_plus.
+rewrite <- scaled_mantissa_abs.
+rewrite <- cexp_abs.
+rewrite scaled_mantissa_mult_bpow.
+unfold cexp, FLX_exp.
+rewrite mag_abs.
+ring_simplify (prec - 1 + (mag beta x - prec))%Z.
+destruct (mag beta x) as (ex,Ex).
+now apply Ex.
+(* *)
+apply lt_IZR.
+rewrite IZR_Zpower.
+2: now apply Zlt_le_weak.
+rewrite abs_IZR, <- scaled_mantissa_generic with (1 := Hx).
+apply Rmult_lt_reg_r with (bpow (cexp beta FLX_exp x)).
+apply bpow_gt_0.
+rewrite <- bpow_plus.
+rewrite <- scaled_mantissa_abs.
+rewrite <- cexp_abs.
+rewrite scaled_mantissa_mult_bpow.
+unfold cexp, FLX_exp.
+rewrite mag_abs.
+ring_simplify (prec + (mag beta x - prec))%Z.
+destruct (mag beta x) as (ex,Ex).
+now apply Ex.
+Qed.
+
+Theorem FLXN_format_satisfies_any :
+ satisfies_any FLXN_format.
+Proof.
+refine (satisfies_any_eq _ _ _ (generic_format_satisfies_any beta FLX_exp)).
+split ; intros H.
+now apply FLXN_format_generic.
+now apply generic_format_FLXN.
+Qed.
+
+Lemma negligible_exp_FLX :
+ negligible_exp FLX_exp = None.
+Proof.
+case (negligible_exp_spec FLX_exp).
+intros _; reflexivity.
+intros n H2; contradict H2.
+unfold FLX_exp; unfold Prec_gt_0 in prec_gt_0_; omega.
+Qed.
+
+Theorem generic_format_FLX_1 :
+ generic_format beta FLX_exp 1.
+Proof.
+unfold generic_format, scaled_mantissa, cexp, F2R; simpl.
+rewrite Rmult_1_l, (mag_unique beta 1 1).
+{ unfold FLX_exp.
+ rewrite <- IZR_Zpower; [|unfold Prec_gt_0 in prec_gt_0_; omega].
+ rewrite Ztrunc_IZR, IZR_Zpower; [|unfold Prec_gt_0 in prec_gt_0_; omega].
+ rewrite <- bpow_plus.
+ now replace (_ + _)%Z with Z0 by ring. }
+rewrite Rabs_R1; simpl; split; [now right|].
+unfold Z.pow_pos; simpl; rewrite Zmult_1_r; apply IZR_lt.
+assert (H := Zle_bool_imp_le _ _ (radix_prop beta)); omega.
+Qed.
+
+Theorem ulp_FLX_0: (ulp beta FLX_exp 0 = 0)%R.
+Proof.
+unfold ulp; rewrite Req_bool_true; trivial.
+rewrite negligible_exp_FLX; easy.
+Qed.
+
+Lemma ulp_FLX_1 : ulp beta FLX_exp 1 = bpow (-prec + 1).
+Proof.
+unfold ulp, FLX_exp, cexp; rewrite Req_bool_false; [|apply R1_neq_R0].
+rewrite mag_1; f_equal; ring.
+Qed.
+
+Lemma succ_FLX_1 : (succ beta FLX_exp 1 = 1 + bpow (-prec + 1))%R.
+Proof.
+now unfold succ; rewrite Rle_bool_true; [|apply Rle_0_1]; rewrite ulp_FLX_1.
+Qed.
+
+Theorem eq_0_round_0_FLX :
+ forall rnd {Vr: Valid_rnd rnd} x,
+ round beta FLX_exp rnd x = 0%R -> x = 0%R.
+Proof.
+intros rnd Hr x.
+apply eq_0_round_0_negligible_exp; try assumption.
+apply FLX_exp_valid.
+apply negligible_exp_FLX.
+Qed.
+
+Theorem gt_0_round_gt_0_FLX :
+ forall rnd {Vr: Valid_rnd rnd} x,
+ (0 < x)%R -> (0 < round beta FLX_exp rnd x)%R.
+Proof with auto with typeclass_instances.
+intros rnd Hr x Hx.
+assert (K: (0 <= round beta FLX_exp rnd x)%R).
+rewrite <- (round_0 beta FLX_exp rnd).
+apply round_le... now apply Rlt_le.
+destruct K; try easy.
+absurd (x = 0)%R.
+now apply Rgt_not_eq.
+apply eq_0_round_0_FLX with rnd...
+Qed.
+
+
+Theorem ulp_FLX_le :
+ forall x, (ulp beta FLX_exp x <= Rabs x * bpow (1-prec))%R.
+Proof.
+intros x; case (Req_dec x 0); intros Hx.
+rewrite Hx, ulp_FLX_0, Rabs_R0.
+right; ring.
+rewrite ulp_neq_0; try exact Hx.
+unfold cexp, FLX_exp.
+replace (mag beta x - prec)%Z with ((mag beta x - 1) + (1-prec))%Z by ring.
+rewrite bpow_plus.
+apply Rmult_le_compat_r.
+apply bpow_ge_0.
+now apply bpow_mag_le.
+Qed.
+
+Theorem ulp_FLX_ge :
+ forall x, (Rabs x * bpow (-prec) <= ulp beta FLX_exp x)%R.
+Proof.
+intros x; case (Req_dec x 0); intros Hx.
+rewrite Hx, ulp_FLX_0, Rabs_R0.
+right; ring.
+rewrite ulp_neq_0; try exact Hx.
+unfold cexp, FLX_exp.
+unfold Zminus; rewrite bpow_plus.
+apply Rmult_le_compat_r.
+apply bpow_ge_0.
+left; now apply bpow_mag_gt.
+Qed.
+
+Lemma ulp_FLX_exact_shift :
+ forall x e,
+ (ulp beta FLX_exp (x * bpow e) = ulp beta FLX_exp x * bpow e)%R.
+Proof.
+intros x e.
+destruct (Req_dec x 0) as [Hx|Hx].
+{ unfold ulp.
+ now rewrite !Req_bool_true, negligible_exp_FLX; rewrite ?Hx, ?Rmult_0_l. }
+unfold ulp; rewrite Req_bool_false;
+ [|now intro H; apply Hx, (Rmult_eq_reg_r (bpow e));
+ [rewrite Rmult_0_l|apply Rgt_not_eq, Rlt_gt, bpow_gt_0]].
+rewrite (Req_bool_false _ _ Hx), <- bpow_plus; f_equal; unfold cexp, FLX_exp.
+now rewrite mag_mult_bpow; [ring|].
+Qed.
+
+Lemma succ_FLX_exact_shift :
+ forall x e,
+ (succ beta FLX_exp (x * bpow e) = succ beta FLX_exp x * bpow e)%R.
+Proof.
+intros x e.
+destruct (Rle_or_lt 0 x) as [Px|Nx].
+{ rewrite succ_eq_pos; [|now apply Rmult_le_pos, bpow_ge_0].
+ rewrite (succ_eq_pos _ _ _ Px).
+ now rewrite Rmult_plus_distr_r; f_equal; apply ulp_FLX_exact_shift. }
+unfold succ.
+rewrite Rle_bool_false; [|assert (H := bpow_gt_0 beta e); nra].
+rewrite Rle_bool_false; [|now simpl].
+rewrite Ropp_mult_distr_l_reverse, <-Ropp_mult_distr_l_reverse; f_equal.
+unfold pred_pos.
+rewrite mag_mult_bpow; [|lra].
+replace (_ - 1)%Z with (mag beta (- x) - 1 + e)%Z; [|ring]; rewrite bpow_plus.
+unfold Req_bool; rewrite Rcompare_mult_r; [|now apply bpow_gt_0].
+fold (Req_bool (-x) (bpow (mag beta (-x) - 1))); case Req_bool.
+{ unfold FLX_exp.
+ replace (_ - _)%Z with (mag beta (- x) - 1 - prec + e)%Z; [|ring].
+ rewrite bpow_plus; ring. }
+rewrite ulp_FLX_exact_shift; ring.
+Qed.
+
+(** FLX is a nice format: it has a monotone exponent... *)
+Global Instance FLX_exp_monotone : Monotone_exp FLX_exp.
+Proof.
+intros ex ey Hxy.
+now apply Zplus_le_compat_r.
+Qed.
+
+(** and it allows a rounding to nearest, ties to even. *)
+Hypothesis NE_prop : Z.even beta = false \/ (1 < prec)%Z.
+
+Global Instance exists_NE_FLX : Exists_NE beta FLX_exp.
+Proof.
+destruct NE_prop as [H|H].
+now left.
+right.
+unfold FLX_exp.
+split ; omega.
+Qed.
+
+End RND_FLX.
diff --git a/flocq/Core/Fcore_FTZ.v b/flocq/Core/FTZ.v
index a2fab00b..1a93bcd9 100644
--- a/flocq/Core/Fcore_FTZ.v
+++ b/flocq/Core/FTZ.v
@@ -2,9 +2,9 @@
This file is part of the Flocq formalization of floating-point
arithmetic in Coq: http://flocq.gforge.inria.fr/
-Copyright (C) 2010-2013 Sylvie Boldo
+Copyright (C) 2009-2018 Sylvie Boldo
#<br />#
-Copyright (C) 2010-2013 Guillaume Melquiond
+Copyright (C) 2009-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
@@ -18,13 +18,8 @@ COPYING file for more details.
*)
(** * 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.
+Require Import Raux Defs Round_pred Generic_fmt.
+Require Import Float_prop Ulp FLX.
Section RND_FTZ.
@@ -36,11 +31,12 @@ Variable emin prec : Z.
Context { prec_gt_0_ : Prec_gt_0 prec }.
-(* floating-point format with abrupt underflow *)
-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.
+Inductive FTZ_format (x : R) : Prop :=
+ FTZ_spec (f : float beta) :
+ x = F2R f ->
+ (x <> 0%R -> Zpower beta (prec - 1) <= Z.abs (Fnum f) < Zpower beta prec)%Z ->
+ (emin <= Fexp f)%Z ->
+ FTZ_format x.
Definition FTZ_exp e := if Zlt_bool (e - prec) emin then (emin + prec - 1)%Z else (e - prec)%Z.
@@ -73,9 +69,10 @@ Qed.
Theorem FLXN_format_FTZ :
forall x, FTZ_format x -> FLXN_format beta prec x.
Proof.
-intros x ((xm, xe), (Hx1, (Hx2, Hx3))).
+intros x [[xm xe] Hx1 Hx2 Hx3].
eexists.
-apply (conj Hx1 Hx2).
+exact Hx1.
+exact Hx2.
Qed.
Theorem generic_format_FTZ :
@@ -83,9 +80,9 @@ Theorem generic_format_FTZ :
Proof.
intros x Hx.
cut (generic_format beta (FLX_exp prec) x).
-apply generic_inclusion_ln_beta.
+apply generic_inclusion_mag.
intros Zx.
-destruct Hx as ((xm, xe), (Hx1, (Hx2, Hx3))).
+destruct Hx as [[xm xe] Hx1 Hx2 Hx3].
simpl in Hx2, Hx3.
specialize (Hx2 Zx).
assert (Zxm: xm <> Z0).
@@ -94,11 +91,11 @@ rewrite Hx1, Zx.
apply F2R_0.
unfold FTZ_exp, FLX_exp.
rewrite Zlt_bool_false.
-apply Zle_refl.
-rewrite Hx1, ln_beta_F2R with (1 := Zxm).
-cut (prec - 1 < ln_beta beta (Z2R xm))%Z.
+apply Z.le_refl.
+rewrite Hx1, mag_F2R with (1 := Zxm).
+cut (prec - 1 < mag beta (IZR xm))%Z.
clear -Hx3 ; omega.
-apply ln_beta_gt_Zpower with (1 := Zxm).
+apply mag_gt_Zpower with (1 := Zxm).
apply Hx2.
apply generic_format_FLXN.
now apply FLXN_format_FTZ.
@@ -108,17 +105,14 @@ Theorem FTZ_format_generic :
forall x, generic_format beta FTZ_exp x -> FTZ_format x.
Proof.
intros x Hx.
-destruct (Req_dec x 0) as [Hx3|Hx3].
+destruct (Req_dec x 0) as [->|Hx3].
exists (Float beta 0 emin).
-split.
-unfold F2R. simpl.
-now rewrite Rmult_0_l.
-split.
+apply sym_eq, F2R_0.
intros H.
now elim H.
-apply Zle_refl.
-unfold generic_format, scaled_mantissa, canonic_exp, FTZ_exp in Hx.
-destruct (ln_beta beta x) as (ex, Hx4).
+apply Z.le_refl.
+unfold generic_format, scaled_mantissa, cexp, FTZ_exp in Hx.
+destruct (mag beta x) as (ex, Hx4).
simpl in Hx.
specialize (Hx4 Hx3).
generalize (Zlt_cases (ex - prec) emin) Hx. clear Hx.
@@ -129,43 +123,43 @@ rewrite Hx2, <- F2R_Zabs.
rewrite <- (Rmult_1_l (bpow ex)).
unfold F2R. simpl.
apply Rmult_le_compat.
-now apply (Z2R_le 0 1).
+now apply IZR_le.
apply bpow_ge_0.
-apply (Z2R_le 1).
+apply IZR_le.
apply (Zlt_le_succ 0).
-apply lt_Z2R.
+apply lt_IZR.
apply Rmult_lt_reg_r with (bpow (emin + prec - 1)).
apply bpow_gt_0.
rewrite Rmult_0_l.
-change (0 < F2R (Float beta (Zabs (Ztrunc (x * bpow (- (emin + prec - 1))))) (emin + prec - 1)))%R.
+change (0 < F2R (Float beta (Z.abs (Ztrunc (x * bpow (- (emin + prec - 1))))) (emin + prec - 1)))%R.
rewrite F2R_Zabs, <- Hx2.
now apply Rabs_pos_lt.
apply bpow_le.
omega.
rewrite Hx2.
eexists ; repeat split ; simpl.
-apply le_Z2R.
-rewrite Z2R_Zpower.
+apply le_IZR.
+rewrite IZR_Zpower.
apply Rmult_le_reg_r with (bpow (ex - prec)).
apply bpow_gt_0.
rewrite <- bpow_plus.
replace (prec - 1 + (ex - prec))%Z with (ex - 1)%Z by ring.
-change (bpow (ex - 1) <= F2R (Float beta (Zabs (Ztrunc (x * bpow (- (ex - prec))))) (ex - prec)))%R.
+change (bpow (ex - 1) <= F2R (Float beta (Z.abs (Ztrunc (x * bpow (- (ex - prec))))) (ex - prec)))%R.
rewrite F2R_Zabs, <- Hx2.
apply Hx4.
apply Zle_minus_le_0.
now apply (Zlt_le_succ 0).
-apply lt_Z2R.
-rewrite Z2R_Zpower.
+apply lt_IZR.
+rewrite IZR_Zpower.
apply Rmult_lt_reg_r with (bpow (ex - prec)).
apply bpow_gt_0.
rewrite <- bpow_plus.
replace (prec + (ex - prec))%Z with ex by ring.
-change (F2R (Float beta (Zabs (Ztrunc (x * bpow (- (ex - prec))))) (ex - prec)) < bpow ex)%R.
+change (F2R (Float beta (Z.abs (Ztrunc (x * bpow (- (ex - prec))))) (ex - prec)) < bpow ex)%R.
rewrite F2R_Zabs, <- Hx2.
apply Hx4.
now apply Zlt_le_weak.
-now apply Zge_le.
+now apply Z.ge_le.
Qed.
Theorem FTZ_format_satisfies_any :
@@ -191,11 +185,12 @@ apply generic_inclusion_ge.
intros e He.
unfold FTZ_exp.
rewrite Zlt_bool_false.
-apply Zle_refl.
+apply Z.le_refl.
omega.
Qed.
-Theorem ulp_FTZ_0: ulp beta FTZ_exp 0 = bpow (emin+prec-1).
+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).
@@ -230,9 +225,9 @@ case Rle_bool_spec ; intros Hx ;
4: easy.
(* 1 <= |x| *)
now apply Zrnd_le.
-rewrite <- (Zrnd_Z2R rnd 0).
+rewrite <- (Zrnd_IZR rnd 0).
apply Zrnd_le...
-apply Rle_trans with (Z2R (-1)). 2: now apply Z2R_le.
+apply Rle_trans with (-1)%R. 2: now apply IZR_le.
destruct (Rabs_ge_inv _ _ Hx) as [Hx1|Hx1].
exact Hx1.
elim Rle_not_lt with (1 := Hx1).
@@ -240,10 +235,10 @@ apply Rle_lt_trans with (2 := Hy).
apply Rle_trans with (1 := Hxy).
apply RRle_abs.
(* |x| < 1 *)
-rewrite <- (Zrnd_Z2R rnd 0).
+rewrite <- (Zrnd_IZR rnd 0).
apply Zrnd_le...
-apply Rle_trans with (Z2R 1).
-now apply Z2R_le.
+apply Rle_trans with 1%R.
+now apply IZR_le.
destruct (Rabs_ge_inv _ _ Hy) as [Hy1|Hy1].
elim Rle_not_lt with (1 := Hy1).
apply Rlt_le_trans with (2 := Hxy).
@@ -252,12 +247,12 @@ exact Hy1.
(* *)
intros n.
unfold Zrnd_FTZ.
-rewrite Zrnd_Z2R...
+rewrite Zrnd_IZR...
case Rle_bool_spec.
easy.
-rewrite <- Z2R_abs.
+rewrite <- abs_IZR.
intros H.
-generalize (lt_Z2R _ 1 H).
+generalize (lt_IZR _ 1 H).
clear.
now case n ; trivial ; simpl ; intros [p|p|].
Qed.
@@ -268,8 +263,8 @@ Theorem round_FTZ_FLX :
round beta FTZ_exp Zrnd_FTZ x = round beta (FLX_exp prec) rnd x.
Proof.
intros x Hx.
-unfold round, scaled_mantissa, canonic_exp.
-destruct (ln_beta beta x) as (ex, He). simpl.
+unfold round, scaled_mantissa, cexp.
+destruct (mag beta x) as (ex, He). simpl.
assert (Hx0: x <> 0%R).
intros Hx0.
apply Rle_not_lt with (1 := Hx).
@@ -306,14 +301,14 @@ Qed.
Theorem round_FTZ_small :
forall x : R,
(Rabs x < bpow (emin + prec - 1))%R ->
- round beta FTZ_exp Zrnd_FTZ x = R0.
+ round beta FTZ_exp Zrnd_FTZ x = 0%R.
Proof with auto with typeclass_instances.
intros x Hx.
destruct (Req_dec x 0) as [Hx0|Hx0].
rewrite Hx0.
apply round_0...
-unfold round, scaled_mantissa, canonic_exp.
-destruct (ln_beta beta x) as (ex, He). simpl.
+unfold round, scaled_mantissa, cexp.
+destruct (mag beta x) as (ex, He). simpl.
specialize (He Hx0).
unfold Zrnd_FTZ.
rewrite Rle_bool_false.
@@ -331,7 +326,7 @@ unfold FTZ_exp.
generalize (Zlt_cases (ex - prec) emin).
case Zlt_bool.
intros _.
-apply Zle_refl.
+apply Z.le_refl.
intros He'.
elim Rlt_not_le with (1 := Hx).
apply Rle_trans with (2 := proj1 He).
diff --git a/flocq/Core/Fcore_FLX.v b/flocq/Core/Fcore_FLX.v
deleted file mode 100644
index 55f6db61..00000000
--- a/flocq/Core/Fcore_FLX.v
+++ /dev/null
@@ -1,271 +0,0 @@
-(**
-This file is part of the Flocq formalization of floating-point
-arithmetic in Coq: http://flocq.gforge.inria.fr/
-
-Copyright (C) 2010-2013 Sylvie Boldo
-#<br />#
-Copyright (C) 2010-2013 Guillaume Melquiond
-
-This library is free software; you can redistribute it and/or
-modify it under the terms of the GNU Lesser General Public
-License as published by the Free Software Foundation; either
-version 3 of the License, or (at your option) any later version.
-
-This library is distributed in the hope that it will be useful,
-but WITHOUT ANY WARRANTY; without even the implied warranty of
-MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
-COPYING file for more details.
-*)
-
-(** * Floating-point format without 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_FIX.
-Require Import Fcore_ulp.
-Require Import Fcore_rnd_ne.
-
-Section RND_FLX.
-
-Variable beta : radix.
-
-Notation bpow e := (bpow beta e).
-
-Variable prec : Z.
-
-Class Prec_gt_0 :=
- prec_gt_0 : (0 < prec)%Z.
-
-Context { prec_gt_0_ : Prec_gt_0 }.
-
-(* unbounded floating-point format *)
-Definition FLX_format (x : R) :=
- exists f : float beta,
- x = F2R f /\ (Zabs (Fnum f) < Zpower beta prec)%Z.
-
-Definition FLX_exp (e : Z) := (e - prec)%Z.
-
-(** Properties of the FLX format *)
-
-Global Instance FLX_exp_valid : Valid_exp FLX_exp.
-Proof.
-intros k.
-unfold FLX_exp.
-generalize prec_gt_0.
-repeat split ; intros ; omega.
-Qed.
-
-Theorem FIX_format_FLX :
- forall x e,
- (bpow (e - 1) <= Rabs x <= bpow e)%R ->
- FLX_format x ->
- FIX_format beta (e - prec) x.
-Proof.
-clear prec_gt_0_.
-intros x e Hx ((xm, xe), (H1, H2)).
-rewrite H1, (F2R_prec_normalize beta xm xe e prec).
-now eexists.
-exact H2.
-now rewrite <- H1.
-Qed.
-
-Theorem FLX_format_generic :
- forall x, generic_format beta FLX_exp x -> FLX_format x.
-Proof.
-intros x H.
-rewrite H.
-unfold FLX_format.
-eexists ; repeat split.
-simpl.
-apply lt_Z2R.
-rewrite Z2R_abs.
-rewrite <- scaled_mantissa_generic with (1 := H).
-rewrite <- scaled_mantissa_abs.
-apply Rmult_lt_reg_r with (bpow (canonic_exp beta FLX_exp (Rabs x))).
-apply bpow_gt_0.
-rewrite scaled_mantissa_mult_bpow.
-rewrite Z2R_Zpower, <- bpow_plus.
-2: now apply Zlt_le_weak.
-unfold canonic_exp, FLX_exp.
-ring_simplify (prec + (ln_beta beta (Rabs x) - prec))%Z.
-rewrite ln_beta_abs.
-destruct (Req_dec x 0) as [Hx|Hx].
-rewrite Hx, Rabs_R0.
-apply bpow_gt_0.
-destruct (ln_beta beta x) as (ex, Ex).
-now apply Ex.
-Qed.
-
-Theorem generic_format_FLX :
- forall x, FLX_format x -> generic_format beta FLX_exp x.
-Proof.
-clear prec_gt_0_.
-intros x ((mx,ex),(H1,H2)).
-simpl in H2.
-rewrite H1.
-apply generic_format_F2R.
-intros Zmx.
-unfold canonic_exp, FLX_exp.
-rewrite ln_beta_F2R with (1 := Zmx).
-apply Zplus_le_reg_r with (prec - ex)%Z.
-ring_simplify.
-now apply ln_beta_le_Zpower.
-Qed.
-
-Theorem FLX_format_satisfies_any :
- satisfies_any FLX_format.
-Proof.
-refine (satisfies_any_eq _ _ _ (generic_format_satisfies_any beta FLX_exp)).
-intros x.
-split.
-apply FLX_format_generic.
-apply generic_format_FLX.
-Qed.
-
-Theorem FLX_format_FIX :
- forall x e,
- (bpow (e - 1) <= Rabs x <= bpow e)%R ->
- FIX_format beta (e - prec) x ->
- FLX_format x.
-Proof with auto with typeclass_instances.
-intros x e Hx Fx.
-apply FLX_format_generic.
-apply generic_format_FIX in Fx.
-revert Fx.
-apply generic_inclusion with (e := e)...
-apply Zle_refl.
-Qed.
-
-(** unbounded floating-point format with normal mantissas *)
-Definition FLXN_format (x : R) :=
- exists f : float beta,
- x = F2R f /\ (x <> R0 ->
- Zpower beta (prec - 1) <= Zabs (Fnum f) < Zpower beta prec)%Z.
-
-Theorem generic_format_FLXN :
- forall x, FLXN_format x -> generic_format beta FLX_exp x.
-Proof.
-intros x ((xm,ex),(H1,H2)).
-destruct (Req_dec x 0) as [Zx|Zx].
-rewrite Zx.
-apply generic_format_0.
-specialize (H2 Zx).
-apply generic_format_FLX.
-rewrite H1.
-eexists ; repeat split.
-apply H2.
-Qed.
-
-Theorem FLXN_format_generic :
- forall x, generic_format beta FLX_exp x -> FLXN_format x.
-Proof.
-intros x Hx.
-rewrite Hx.
-simpl.
-eexists ; split. split.
-simpl.
-rewrite <- Hx.
-intros Zx.
-split.
-(* *)
-apply le_Z2R.
-rewrite Z2R_Zpower.
-2: now apply Zlt_0_le_0_pred.
-rewrite Z2R_abs, <- scaled_mantissa_generic with (1 := Hx).
-apply Rmult_le_reg_r with (bpow (canonic_exp beta FLX_exp x)).
-apply bpow_gt_0.
-rewrite <- bpow_plus.
-rewrite <- scaled_mantissa_abs.
-rewrite <- canonic_exp_abs.
-rewrite scaled_mantissa_mult_bpow.
-unfold canonic_exp, FLX_exp.
-rewrite ln_beta_abs.
-ring_simplify (prec - 1 + (ln_beta beta x - prec))%Z.
-destruct (ln_beta beta x) as (ex,Ex).
-now apply Ex.
-(* *)
-apply lt_Z2R.
-rewrite Z2R_Zpower.
-2: now apply Zlt_le_weak.
-rewrite Z2R_abs, <- scaled_mantissa_generic with (1 := Hx).
-apply Rmult_lt_reg_r with (bpow (canonic_exp beta FLX_exp x)).
-apply bpow_gt_0.
-rewrite <- bpow_plus.
-rewrite <- scaled_mantissa_abs.
-rewrite <- canonic_exp_abs.
-rewrite scaled_mantissa_mult_bpow.
-unfold canonic_exp, FLX_exp.
-rewrite ln_beta_abs.
-ring_simplify (prec + (ln_beta beta x - prec))%Z.
-destruct (ln_beta beta x) as (ex,Ex).
-now apply Ex.
-Qed.
-
-Theorem FLXN_format_satisfies_any :
- satisfies_any FLXN_format.
-Proof.
-refine (satisfies_any_eq _ _ _ (generic_format_satisfies_any beta FLX_exp)).
-split ; intros H.
-now apply FLXN_format_generic.
-now apply generic_format_FLXN.
-Qed.
-
-Theorem ulp_FLX_0: (ulp beta FLX_exp 0 = 0)%R.
-Proof.
-unfold ulp; rewrite Req_bool_true; trivial.
-case (negligible_exp_spec FLX_exp).
-intros _; reflexivity.
-intros n H2; contradict H2.
-unfold FLX_exp; unfold Prec_gt_0 in prec_gt_0_; omega.
-Qed.
-
-Theorem ulp_FLX_le: forall x, (ulp beta FLX_exp x <= Rabs x * bpow (1-prec))%R.
-Proof.
-intros x; case (Req_dec x 0); intros Hx.
-rewrite Hx, ulp_FLX_0, Rabs_R0.
-right; ring.
-rewrite ulp_neq_0; try exact Hx.
-unfold canonic_exp, FLX_exp.
-replace (ln_beta beta x - prec)%Z with ((ln_beta beta x - 1) + (1-prec))%Z by ring.
-rewrite bpow_plus.
-apply Rmult_le_compat_r.
-apply bpow_ge_0.
-now apply bpow_ln_beta_le.
-Qed.
-
-
-Theorem ulp_FLX_ge: forall x, (Rabs x * bpow (-prec) <= ulp beta FLX_exp x)%R.
-Proof.
-intros x; case (Req_dec x 0); intros Hx.
-rewrite Hx, ulp_FLX_0, Rabs_R0.
-right; ring.
-rewrite ulp_neq_0; try exact Hx.
-unfold canonic_exp, FLX_exp.
-unfold Zminus; rewrite bpow_plus.
-apply Rmult_le_compat_r.
-apply bpow_ge_0.
-left; now apply bpow_ln_beta_gt.
-Qed.
-
-(** FLX is a nice format: it has a monotone exponent... *)
-Global Instance FLX_exp_monotone : Monotone_exp FLX_exp.
-Proof.
-intros ex ey Hxy.
-now apply Zplus_le_compat_r.
-Qed.
-
-(** and it allows a rounding to nearest, ties to even. *)
-Hypothesis NE_prop : Zeven beta = false \/ (1 < prec)%Z.
-
-Global Instance exists_NE_FLX : Exists_NE beta FLX_exp.
-Proof.
-destruct NE_prop as [H|H].
-now left.
-right.
-unfold FLX_exp.
-split ; omega.
-Qed.
-
-End RND_FLX.
diff --git a/flocq/Core/Fcore_float_prop.v b/flocq/Core/Float_prop.v
index a183bf0a..804dd397 100644
--- a/flocq/Core/Fcore_float_prop.v
+++ b/flocq/Core/Float_prop.v
@@ -2,9 +2,9 @@
This file is part of the Flocq formalization of floating-point
arithmetic in Coq: http://flocq.gforge.inria.fr/
-Copyright (C) 2010-2013 Sylvie Boldo
+Copyright (C) 2009-2018 Sylvie Boldo
#<br />#
-Copyright (C) 2010-2013 Guillaume Melquiond
+Copyright (C) 2009-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
@@ -18,40 +18,38 @@ COPYING file for more details.
*)
(** * Basic properties of floating-point formats: lemmas about mantissa, exponent... *)
-Require Import Fcore_Raux.
-Require Import Fcore_defs.
+Require Import Raux Defs Digits.
Section Float_prop.
Variable beta : radix.
-
Notation bpow e := (bpow beta e).
Theorem Rcompare_F2R :
forall e m1 m2 : Z,
- Rcompare (F2R (Float beta m1 e)) (F2R (Float beta m2 e)) = Zcompare m1 m2.
+ Rcompare (F2R (Float beta m1 e)) (F2R (Float beta m2 e)) = Z.compare m1 m2.
Proof.
intros e m1 m2.
unfold F2R. simpl.
rewrite Rcompare_mult_r.
-apply Rcompare_Z2R.
+apply Rcompare_IZR.
apply bpow_gt_0.
Qed.
(** Basic facts *)
-Theorem F2R_le_reg :
+Theorem le_F2R :
forall e m1 m2 : Z,
(F2R (Float beta m1 e) <= F2R (Float beta m2 e))%R ->
(m1 <= m2)%Z.
Proof.
intros e m1 m2 H.
-apply le_Z2R.
+apply le_IZR.
apply Rmult_le_reg_r with (bpow e).
apply bpow_gt_0.
exact H.
Qed.
-Theorem F2R_le_compat :
+Theorem F2R_le :
forall m1 m2 e : Z,
(m1 <= m2)%Z ->
(F2R (Float beta m1 e) <= F2R (Float beta m2 e))%R.
@@ -60,22 +58,22 @@ intros m1 m2 e H.
unfold F2R. simpl.
apply Rmult_le_compat_r.
apply bpow_ge_0.
-now apply Z2R_le.
+now apply IZR_le.
Qed.
-Theorem F2R_lt_reg :
+Theorem lt_F2R :
forall e m1 m2 : Z,
(F2R (Float beta m1 e) < F2R (Float beta m2 e))%R ->
(m1 < m2)%Z.
Proof.
intros e m1 m2 H.
-apply lt_Z2R.
+apply lt_IZR.
apply Rmult_lt_reg_r with (bpow e).
apply bpow_gt_0.
exact H.
Qed.
-Theorem F2R_lt_compat :
+Theorem F2R_lt :
forall e m1 m2 : Z,
(m1 < m2)%Z ->
(F2R (Float beta m1 e) < F2R (Float beta m2 e))%R.
@@ -84,10 +82,10 @@ intros e m1 m2 H.
unfold F2R. simpl.
apply Rmult_lt_compat_r.
apply bpow_gt_0.
-now apply Z2R_lt.
+now apply IZR_lt.
Qed.
-Theorem F2R_eq_compat :
+Theorem F2R_eq :
forall e m1 m2 : Z,
(m1 = m2)%Z ->
(F2R (Float beta m1 e) = F2R (Float beta m2 e))%R.
@@ -96,26 +94,26 @@ intros e m1 m2 H.
now apply (f_equal (fun m => F2R (Float beta m e))).
Qed.
-Theorem F2R_eq_reg :
+Theorem eq_F2R :
forall e m1 m2 : Z,
F2R (Float beta m1 e) = F2R (Float beta m2 e) ->
m1 = m2.
Proof.
intros e m1 m2 H.
apply Zle_antisym ;
- apply F2R_le_reg with e ;
+ apply le_F2R with e ;
rewrite H ;
apply Rle_refl.
Qed.
Theorem F2R_Zabs:
forall m e : Z,
- F2R (Float beta (Zabs m) e) = Rabs (F2R (Float beta m e)).
+ F2R (Float beta (Z.abs m) e) = Rabs (F2R (Float beta m e)).
Proof.
intros m e.
unfold F2R.
rewrite Rabs_mult.
-rewrite <- Z2R_abs.
+rewrite <- abs_IZR.
simpl.
apply f_equal.
apply sym_eq; apply Rabs_right.
@@ -125,12 +123,21 @@ Qed.
Theorem F2R_Zopp :
forall m e : Z,
- F2R (Float beta (Zopp m) e) = Ropp (F2R (Float beta m e)).
+ F2R (Float beta (Z.opp m) e) = Ropp (F2R (Float beta m e)).
Proof.
intros m e.
unfold F2R. simpl.
rewrite <- Ropp_mult_distr_l_reverse.
-now rewrite Z2R_opp.
+now rewrite opp_IZR.
+Qed.
+
+Theorem F2R_cond_Zopp :
+ forall b m e,
+ F2R (Float beta (cond_Zopp b m) e) = cond_Ropp b (F2R (Float beta m e)).
+Proof.
+intros [|] m e ; unfold F2R ; simpl.
+now rewrite opp_IZR, Ropp_mult_distr_l_reverse.
+apply refl_equal.
Qed.
(** Sign facts *)
@@ -143,125 +150,125 @@ unfold F2R. simpl.
apply Rmult_0_l.
Qed.
-Theorem F2R_eq_0_reg :
+Theorem eq_0_F2R :
forall m e : Z,
F2R (Float beta m e) = 0%R ->
m = Z0.
Proof.
intros m e H.
-apply F2R_eq_reg with e.
+apply eq_F2R with e.
now rewrite F2R_0.
Qed.
-Theorem F2R_ge_0_reg :
+Theorem ge_0_F2R :
forall m e : Z,
(0 <= F2R (Float beta m e))%R ->
(0 <= m)%Z.
Proof.
intros m e H.
-apply F2R_le_reg with e.
+apply le_F2R with e.
now rewrite F2R_0.
Qed.
-Theorem F2R_le_0_reg :
+Theorem le_0_F2R :
forall m e : Z,
(F2R (Float beta m e) <= 0)%R ->
(m <= 0)%Z.
Proof.
intros m e H.
-apply F2R_le_reg with e.
+apply le_F2R with e.
now rewrite F2R_0.
Qed.
-Theorem F2R_gt_0_reg :
+Theorem gt_0_F2R :
forall m e : Z,
(0 < F2R (Float beta m e))%R ->
(0 < m)%Z.
Proof.
intros m e H.
-apply F2R_lt_reg with e.
+apply lt_F2R with e.
now rewrite F2R_0.
Qed.
-Theorem F2R_lt_0_reg :
+Theorem lt_0_F2R :
forall m e : Z,
(F2R (Float beta m e) < 0)%R ->
(m < 0)%Z.
Proof.
intros m e H.
-apply F2R_lt_reg with e.
+apply lt_F2R with e.
now rewrite F2R_0.
Qed.
-Theorem F2R_ge_0_compat :
+Theorem F2R_ge_0 :
forall f : float beta,
(0 <= Fnum f)%Z ->
(0 <= F2R f)%R.
Proof.
intros f H.
rewrite <- F2R_0 with (Fexp f).
-now apply F2R_le_compat.
+now apply F2R_le.
Qed.
-Theorem F2R_le_0_compat :
+Theorem F2R_le_0 :
forall f : float beta,
(Fnum f <= 0)%Z ->
(F2R f <= 0)%R.
Proof.
intros f H.
rewrite <- F2R_0 with (Fexp f).
-now apply F2R_le_compat.
+now apply F2R_le.
Qed.
-Theorem F2R_gt_0_compat :
+Theorem F2R_gt_0 :
forall f : float beta,
(0 < Fnum f)%Z ->
(0 < F2R f)%R.
Proof.
intros f H.
rewrite <- F2R_0 with (Fexp f).
-now apply F2R_lt_compat.
+now apply F2R_lt.
Qed.
-Theorem F2R_lt_0_compat :
+Theorem F2R_lt_0 :
forall f : float beta,
(Fnum f < 0)%Z ->
(F2R f < 0)%R.
Proof.
intros f H.
rewrite <- F2R_0 with (Fexp f).
-now apply F2R_lt_compat.
+now apply F2R_lt.
Qed.
-Theorem F2R_neq_0_compat :
+Theorem F2R_neq_0 :
forall f : float beta,
(Fnum f <> 0)%Z ->
(F2R f <> 0)%R.
Proof.
intros f H H1.
apply H.
-now apply F2R_eq_0_reg with (Fexp f).
+now apply eq_0_F2R with (Fexp f).
Qed.
-Lemma Fnum_ge_0_compat: forall (f : float beta),
+Lemma Fnum_ge_0: forall (f : float beta),
(0 <= F2R f)%R -> (0 <= Fnum f)%Z.
Proof.
intros f H.
case (Zle_or_lt 0 (Fnum f)); trivial.
intros H1; contradict H.
apply Rlt_not_le.
-now apply F2R_lt_0_compat.
+now apply F2R_lt_0.
Qed.
-Lemma Fnum_le_0_compat: forall (f : float beta),
+Lemma Fnum_le_0: forall (f : float beta),
(F2R f <= 0)%R -> (Fnum f <= 0)%Z.
Proof.
intros f H.
case (Zle_or_lt (Fnum f) 0); trivial.
intros H1; contradict H.
apply Rlt_not_le.
-now apply F2R_gt_0_compat.
+now apply F2R_gt_0.
Qed.
(** Floats and bpow *)
@@ -281,7 +288,7 @@ Theorem bpow_le_F2R :
Proof.
intros m e H.
rewrite <- F2R_bpow.
-apply F2R_le_compat.
+apply F2R_le.
now apply (Zlt_le_succ 0).
Qed.
@@ -301,7 +308,7 @@ unfold F2R. simpl.
rewrite <- (Rmult_1_l (bpow e1)) at 1.
apply Rmult_le_compat_r.
apply bpow_ge_0.
-apply (Z2R_le 1).
+apply IZR_le.
now apply (Zlt_le_succ 0).
now apply Rlt_le.
(* . *)
@@ -309,14 +316,14 @@ revert H.
replace e2 with (e2 - e1 + e1)%Z by ring.
rewrite bpow_plus.
unfold F2R. simpl.
-rewrite <- (Z2R_Zpower beta (e2 - e1)).
+rewrite <- (IZR_Zpower beta (e2 - e1)).
intros H.
apply Rmult_le_compat_r.
apply bpow_ge_0.
apply Rmult_lt_reg_r in H.
-apply Z2R_le.
+apply IZR_le.
apply Zlt_le_succ.
-now apply lt_Z2R.
+now apply lt_IZR.
apply bpow_gt_0.
now apply Zle_minus_le_0.
Qed.
@@ -332,16 +339,16 @@ case (Zle_or_lt e1 e2); intros He.
replace e2 with (e2 - e1 + e1)%Z by ring.
rewrite bpow_plus.
unfold F2R. simpl.
-rewrite <- (Z2R_Zpower beta (e2 - e1)).
+rewrite <- (IZR_Zpower beta (e2 - e1)).
intros H.
apply Rmult_le_compat_r.
apply bpow_ge_0.
apply Rmult_lt_reg_r in H.
-apply Z2R_le.
+apply IZR_le.
rewrite (Zpred_succ (Zpower _ _)).
apply Zplus_le_compat_r.
apply Zlt_le_succ.
-now apply lt_Z2R.
+now apply lt_IZR.
apply bpow_gt_0.
now apply Zle_minus_le_0.
intros H.
@@ -352,14 +359,13 @@ now apply Zlt_le_weak.
unfold F2R. simpl.
apply Rmult_le_compat_r.
apply bpow_ge_0.
-replace 1%R with (Z2R 1) by reflexivity.
-apply Z2R_le.
+apply IZR_le.
omega.
Qed.
Theorem F2R_lt_bpow :
forall f : float beta, forall e',
- (Zabs (Fnum f) < Zpower beta (e' - Fexp f))%Z ->
+ (Z.abs (Fnum f) < Zpower beta (e' - Fexp f))%Z ->
(Rabs (F2R f) < bpow e')%R.
Proof.
intros (m, e) e' Hm.
@@ -369,8 +375,8 @@ unfold F2R. simpl.
apply Rmult_lt_reg_r with (bpow (-e)).
apply bpow_gt_0.
rewrite Rmult_assoc, <- 2!bpow_plus, Zplus_opp_r, Rmult_1_r.
-rewrite <-Z2R_Zpower. 2: now apply Zle_left.
-now apply Z2R_lt.
+rewrite <-IZR_Zpower. 2: now apply Zle_left.
+now apply IZR_lt.
elim Zlt_not_le with (1 := Hm).
simpl.
cut (e' - e < 0)%Z. 2: omega.
@@ -387,7 +393,7 @@ Theorem F2R_change_exp :
Proof.
intros e' m e He.
unfold F2R. simpl.
-rewrite Z2R_mult, Z2R_Zpower, Rmult_assoc.
+rewrite mult_IZR, IZR_Zpower, Rmult_assoc.
apply f_equal.
pattern e at 1 ; replace e with (e - e' + e')%Z by ring.
apply bpow_plus.
@@ -396,7 +402,7 @@ Qed.
Theorem F2R_prec_normalize :
forall m e e' p : Z,
- (Zabs m < Zpower beta p)%Z ->
+ (Z.abs m < Zpower beta p)%Z ->
(bpow (e' - 1)%Z <= Rabs (F2R (Float beta m e)))%R ->
F2R (Float beta m e) = F2R (Float beta (m * Zpower beta (e - e' + p)) (e' - p)).
Proof.
@@ -413,23 +419,23 @@ apply Rle_lt_trans with (1 := Hf).
rewrite <- F2R_Zabs, Zplus_comm, bpow_plus.
apply Rmult_lt_compat_r.
apply bpow_gt_0.
-rewrite <- Z2R_Zpower.
-now apply Z2R_lt.
+rewrite <- IZR_Zpower.
+now apply IZR_lt.
exact Hp.
Qed.
-(** Floats and ln_beta *)
-Theorem ln_beta_F2R_bounds :
+(** Floats and mag *)
+Theorem mag_F2R_bounds :
forall x m e, (0 < m)%Z ->
(F2R (Float beta m e) <= x < F2R (Float beta (m + 1) e))%R ->
- ln_beta beta x = ln_beta beta (F2R (Float beta m e)) :> Z.
+ mag beta x = mag beta (F2R (Float beta m e)) :> Z.
Proof.
intros x m e Hp (Hx,Hx2).
-destruct (ln_beta beta (F2R (Float beta m e))) as (ex, He).
+destruct (mag beta (F2R (Float beta m e))) as (ex, He).
simpl.
-apply ln_beta_unique.
+apply mag_unique.
assert (Hp1: (0 < F2R (Float beta m e))%R).
-now apply F2R_gt_0_compat.
+now apply F2R_gt_0.
specialize (He (Rgt_not_eq _ _ Hp1)).
rewrite Rabs_pos_eq in He. 2: now apply Rlt_le.
destruct He as (He1, He2).
@@ -442,22 +448,65 @@ apply Rlt_le_trans with (1 := Hx2).
now apply F2R_p1_le_bpow.
Qed.
-Theorem ln_beta_F2R :
+Theorem mag_F2R :
forall m e : Z,
m <> Z0 ->
- (ln_beta beta (F2R (Float beta m e)) = ln_beta beta (Z2R m) + e :> Z)%Z.
+ (mag beta (F2R (Float beta m e)) = mag beta (IZR m) + e :> Z)%Z.
Proof.
intros m e H.
unfold F2R ; simpl.
-apply ln_beta_mult_bpow.
-exact (Z2R_neq m 0 H).
+apply mag_mult_bpow.
+now apply IZR_neq.
+Qed.
+
+Theorem Zdigits_mag :
+ forall n,
+ n <> Z0 ->
+ Zdigits beta n = mag beta (IZR n).
+Proof.
+intros n Hn.
+destruct (mag beta (IZR n)) as (e, He) ; simpl.
+specialize (He (IZR_neq _ _ Hn)).
+rewrite <- abs_IZR in He.
+assert (Hd := Zdigits_correct beta n).
+assert (Hd' := Zdigits_gt_0 beta n).
+apply Zle_antisym ; apply (bpow_lt_bpow beta).
+apply Rle_lt_trans with (2 := proj2 He).
+rewrite <- IZR_Zpower by omega.
+now apply IZR_le.
+apply Rle_lt_trans with (1 := proj1 He).
+rewrite <- IZR_Zpower by omega.
+now apply IZR_lt.
+Qed.
+
+Theorem mag_F2R_Zdigits :
+ forall m e, m <> Z0 ->
+ (mag beta (F2R (Float beta m e)) = Zdigits beta m + e :> Z)%Z.
+Proof.
+intros m e Hm.
+rewrite mag_F2R with (1 := Hm).
+apply (f_equal (fun v => Zplus v e)).
+apply sym_eq.
+now apply Zdigits_mag.
+Qed.
+
+Theorem mag_F2R_bounds_Zdigits :
+ forall x m e, (0 < m)%Z ->
+ (F2R (Float beta m e) <= x < F2R (Float beta (m + 1) e))%R ->
+ mag beta x = (Zdigits beta m + e)%Z :> Z.
+Proof.
+intros x m e Hm Bx.
+apply mag_F2R_bounds with (1 := Hm) in Bx.
+rewrite Bx.
+apply mag_F2R_Zdigits.
+now apply Zgt_not_eq.
Qed.
Theorem float_distribution_pos :
forall m1 e1 m2 e2 : Z,
(0 < m1)%Z ->
(F2R (Float beta m1 e1) < F2R (Float beta m2 e2) < F2R (Float beta (m1 + 1) e1))%R ->
- (e2 < e1)%Z /\ (e1 + ln_beta beta (Z2R m1) = e2 + ln_beta beta (Z2R m2))%Z.
+ (e2 < e1)%Z /\ (e1 + mag beta (IZR m1) = e2 + mag beta (IZR m2))%Z.
Proof.
intros m1 e1 m2 e2 Hp1 (H12, H21).
assert (He: (e2 < e1)%Z).
@@ -465,35 +514,35 @@ assert (He: (e2 < e1)%Z).
apply Znot_ge_lt.
intros H0.
elim Rlt_not_le with (1 := H21).
-apply Zge_le in H0.
+apply Z.ge_le in H0.
apply (F2R_change_exp e1 m2 e2) in H0.
rewrite H0.
-apply F2R_le_compat.
+apply F2R_le.
apply Zlt_le_succ.
-apply (F2R_lt_reg e1).
+apply (lt_F2R e1).
now rewrite <- H0.
(* . *)
split.
exact He.
rewrite (Zplus_comm e1), (Zplus_comm e2).
assert (Hp2: (0 < m2)%Z).
-apply (F2R_gt_0_reg m2 e2).
+apply (gt_0_F2R m2 e2).
apply Rlt_trans with (2 := H12).
-now apply F2R_gt_0_compat.
-rewrite <- 2!ln_beta_F2R.
-destruct (ln_beta beta (F2R (Float beta m1 e1))) as (e1', H1).
+now apply F2R_gt_0.
+rewrite <- 2!mag_F2R.
+destruct (mag beta (F2R (Float beta m1 e1))) as (e1', H1).
simpl.
apply sym_eq.
-apply ln_beta_unique.
+apply mag_unique.
assert (H2 : (bpow (e1' - 1) <= F2R (Float beta m1 e1) < bpow e1')%R).
-rewrite <- (Zabs_eq m1), F2R_Zabs.
+rewrite <- (Z.abs_eq m1), F2R_Zabs.
apply H1.
apply Rgt_not_eq.
apply Rlt_gt.
-now apply F2R_gt_0_compat.
+now apply F2R_gt_0.
now apply Zlt_le_weak.
clear H1.
-rewrite <- F2R_Zabs, Zabs_eq.
+rewrite <- F2R_Zabs, Z.abs_eq.
split.
apply Rlt_le.
apply Rle_lt_trans with (2 := H12).
@@ -507,13 +556,4 @@ apply sym_not_eq.
now apply Zlt_not_eq.
Qed.
-Theorem F2R_cond_Zopp :
- forall b m e,
- F2R (Float beta (cond_Zopp b m) e) = cond_Ropp b (F2R (Float beta m e)).
-Proof.
-intros [|] m e ; unfold F2R ; simpl.
-now rewrite Z2R_opp, Ropp_mult_distr_l_reverse.
-apply refl_equal.
-Qed.
-
End Float_prop.
diff --git a/flocq/Core/Fcore_generic_fmt.v b/flocq/Core/Generic_fmt.v
index 668b4da2..cb37bd91 100644
--- a/flocq/Core/Fcore_generic_fmt.v
+++ b/flocq/Core/Generic_fmt.v
@@ -2,9 +2,9 @@
This file is part of the Flocq formalization of floating-point
arithmetic in Coq: http://flocq.gforge.inria.fr/
-Copyright (C) 2010-2013 Sylvie Boldo
+Copyright (C) 2009-2018 Sylvie Boldo
#<br />#
-Copyright (C) 2010-2013 Guillaume Melquiond
+Copyright (C) 2009-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
@@ -18,10 +18,7 @@ COPYING file for more details.
*)
(** * What is a real number belonging to a format, and many properties. *)
-Require Import Fcore_Raux.
-Require Import Fcore_defs.
-Require Import Fcore_rnd.
-Require Import Fcore_float_prop.
+Require Import Raux Defs Round_pred Float_prop.
Section Generic.
@@ -53,7 +50,7 @@ Proof.
intros k l Hk H.
apply Znot_ge_lt.
intros Hl.
-apply Zge_le in Hl.
+apply Z.ge_le in Hl.
assert (H' := proj2 (proj2 (valid_exp l) Hl) k).
omega.
Qed.
@@ -66,24 +63,24 @@ Proof.
intros k l Hk H.
apply Znot_ge_lt.
intros H'.
-apply Zge_le in H'.
-assert (Hl := Zle_trans _ _ _ H H').
+apply Z.ge_le in H'.
+assert (Hl := Z.le_trans _ _ _ H H').
apply valid_exp in Hl.
assert (H1 := proj2 Hl k H').
omega.
Qed.
-Definition canonic_exp x :=
- fexp (ln_beta beta x).
+Definition cexp x :=
+ fexp (mag beta x).
-Definition canonic (f : float beta) :=
- Fexp f = canonic_exp (F2R f).
+Definition canonical (f : float beta) :=
+ Fexp f = cexp (F2R f).
Definition scaled_mantissa x :=
- (x * bpow (- canonic_exp x))%R.
+ (x * bpow (- cexp x))%R.
Definition generic_format (x : R) :=
- x = F2R (Float beta (Ztrunc (scaled_mantissa x)) (canonic_exp x)).
+ x = F2R (Float beta (Ztrunc (scaled_mantissa x)) (cexp x)).
(** Basic facts *)
Theorem generic_format_0 :
@@ -91,26 +88,39 @@ Theorem generic_format_0 :
Proof.
unfold generic_format, scaled_mantissa.
rewrite Rmult_0_l.
-change (Ztrunc 0) with (Ztrunc (Z2R 0)).
-now rewrite Ztrunc_Z2R, F2R_0.
+now rewrite Ztrunc_IZR, F2R_0.
Qed.
-Theorem canonic_exp_opp :
+Theorem cexp_opp :
forall x,
- canonic_exp (-x) = canonic_exp x.
+ cexp (-x) = cexp x.
Proof.
intros x.
-unfold canonic_exp.
-now rewrite ln_beta_opp.
+unfold cexp.
+now rewrite mag_opp.
Qed.
-Theorem canonic_exp_abs :
+Theorem cexp_abs :
forall x,
- canonic_exp (Rabs x) = canonic_exp x.
+ cexp (Rabs x) = cexp x.
Proof.
intros x.
-unfold canonic_exp.
-now rewrite ln_beta_abs.
+unfold cexp.
+now rewrite mag_abs.
+Qed.
+
+Theorem canonical_generic_format :
+ forall x,
+ generic_format x ->
+ exists f : float beta,
+ x = F2R f /\ canonical f.
+Proof.
+intros x Hx.
+rewrite Hx.
+eexists.
+apply (conj eq_refl).
+unfold canonical.
+now rewrite <- Hx.
Qed.
Theorem generic_format_bpow :
@@ -118,11 +128,11 @@ Theorem generic_format_bpow :
generic_format (bpow e).
Proof.
intros e H.
-unfold generic_format, scaled_mantissa, canonic_exp.
-rewrite ln_beta_bpow.
+unfold generic_format, scaled_mantissa, cexp.
+rewrite mag_bpow.
rewrite <- bpow_plus.
-rewrite <- (Z2R_Zpower beta (e + - fexp (e + 1))).
-rewrite Ztrunc_Z2R.
+rewrite <- (IZR_Zpower beta (e + - fexp (e + 1))).
+rewrite Ztrunc_IZR.
rewrite <- F2R_bpow.
rewrite F2R_change_exp with (1 := H).
now rewrite Zmult_1_l.
@@ -140,110 +150,107 @@ now apply valid_exp_.
rewrite <- H.
apply valid_exp.
rewrite H.
-apply Zle_refl.
+apply Z.le_refl.
Qed.
Theorem generic_format_F2R :
forall m e,
- ( m <> 0 -> canonic_exp (F2R (Float beta m e)) <= e )%Z ->
+ ( m <> 0 -> cexp (F2R (Float beta m e)) <= e )%Z ->
generic_format (F2R (Float beta m e)).
Proof.
intros m e.
-destruct (Z_eq_dec m 0) as [Zm|Zm].
+destruct (Z.eq_dec m 0) as [Zm|Zm].
intros _.
rewrite Zm, F2R_0.
apply generic_format_0.
unfold generic_format, scaled_mantissa.
-set (e' := canonic_exp (F2R (Float beta m e))).
+set (e' := cexp (F2R (Float beta m e))).
intros He.
specialize (He Zm).
unfold F2R at 3. simpl.
rewrite F2R_change_exp with (1 := He).
-apply F2R_eq_compat.
-rewrite Rmult_assoc, <- bpow_plus, <- Z2R_Zpower, <- Z2R_mult.
-now rewrite Ztrunc_Z2R.
+apply F2R_eq.
+rewrite Rmult_assoc, <- bpow_plus, <- IZR_Zpower, <- mult_IZR.
+now rewrite Ztrunc_IZR.
now apply Zle_left.
Qed.
-Lemma generic_format_F2R': forall (x:R) (f:float beta),
- F2R f = x -> ((x <> 0)%R ->
- (canonic_exp x <= Fexp f)%Z) ->
- generic_format x.
+Lemma generic_format_F2R' :
+ forall (x : R) (f : float beta),
+ F2R f = x ->
+ (x <> 0%R -> (cexp x <= Fexp f)%Z) ->
+ generic_format x.
Proof.
intros x f H1 H2.
rewrite <- H1; destruct f as (m,e).
-apply generic_format_F2R.
+apply generic_format_F2R.
simpl in *; intros H3.
rewrite H1; apply H2.
intros Y; apply H3.
-apply F2R_eq_0_reg with beta e.
+apply eq_0_F2R with beta e.
now rewrite H1.
Qed.
-
-Theorem canonic_opp :
+Theorem canonical_opp :
forall m e,
- canonic (Float beta m e) ->
- canonic (Float beta (-m) e).
+ canonical (Float beta m e) ->
+ canonical (Float beta (-m) e).
Proof.
intros m e H.
-unfold canonic.
-now rewrite F2R_Zopp, canonic_exp_opp.
+unfold canonical.
+now rewrite F2R_Zopp, cexp_opp.
Qed.
-Theorem canonic_abs :
+Theorem canonical_abs :
forall m e,
- canonic (Float beta m e) ->
- canonic (Float beta (Zabs m) e).
+ canonical (Float beta m e) ->
+ canonical (Float beta (Z.abs m) e).
Proof.
intros m e H.
-unfold canonic.
-now rewrite F2R_Zabs, canonic_exp_abs.
+unfold canonical.
+now rewrite F2R_Zabs, cexp_abs.
Qed.
-Theorem canonic_0: canonic (Float beta 0 (fexp (ln_beta beta 0%R))).
+Theorem canonical_0 :
+ canonical (Float beta 0 (fexp (mag beta 0%R))).
Proof.
-unfold canonic; simpl; unfold canonic_exp.
-replace (F2R {| Fnum := 0; Fexp := fexp (ln_beta beta 0) |}) with 0%R.
-reflexivity.
-unfold F2R; simpl; ring.
+unfold canonical; simpl ; unfold cexp.
+now rewrite F2R_0.
Qed.
-
-
-Theorem canonic_unicity :
+Theorem canonical_unique :
forall f1 f2,
- canonic f1 ->
- canonic f2 ->
+ canonical f1 ->
+ canonical f2 ->
F2R f1 = F2R f2 ->
f1 = f2.
Proof.
intros (m1, e1) (m2, e2).
-unfold canonic. simpl.
+unfold canonical. simpl.
intros H1 H2 H.
rewrite H in H1.
rewrite <- H2 in H1. clear H2.
rewrite H1 in H |- *.
apply (f_equal (fun m => Float beta m e2)).
-apply F2R_eq_reg with (1 := H).
+apply eq_F2R with (1 := H).
Qed.
Theorem scaled_mantissa_generic :
forall x,
generic_format x ->
- scaled_mantissa x = Z2R (Ztrunc (scaled_mantissa x)).
+ scaled_mantissa x = IZR (Ztrunc (scaled_mantissa x)).
Proof.
intros x Hx.
unfold scaled_mantissa.
pattern x at 1 3 ; rewrite Hx.
unfold F2R. simpl.
rewrite Rmult_assoc, <- bpow_plus, Zplus_opp_r, Rmult_1_r.
-now rewrite Ztrunc_Z2R.
+now rewrite Ztrunc_IZR.
Qed.
Theorem scaled_mantissa_mult_bpow :
forall x,
- (scaled_mantissa x * bpow (canonic_exp x))%R = x.
+ (scaled_mantissa x * bpow (cexp x))%R = x.
Proof.
intros x.
unfold scaled_mantissa.
@@ -263,7 +270,7 @@ Theorem scaled_mantissa_opp :
Proof.
intros x.
unfold scaled_mantissa.
-rewrite canonic_exp_opp.
+rewrite cexp_opp.
now rewrite Ropp_mult_distr_l_reverse.
Qed.
@@ -273,7 +280,7 @@ Theorem scaled_mantissa_abs :
Proof.
intros x.
unfold scaled_mantissa.
-rewrite canonic_exp_abs, Rabs_mult.
+rewrite cexp_abs, Rabs_mult.
apply f_equal.
apply sym_eq.
apply Rabs_pos_eq.
@@ -285,7 +292,7 @@ Theorem generic_format_opp :
Proof.
intros x Hx.
unfold generic_format.
-rewrite scaled_mantissa_opp, canonic_exp_opp.
+rewrite scaled_mantissa_opp, cexp_opp.
rewrite Ztrunc_opp.
rewrite F2R_Zopp.
now apply f_equal.
@@ -296,7 +303,7 @@ Theorem generic_format_abs :
Proof.
intros x Hx.
unfold generic_format.
-rewrite scaled_mantissa_abs, canonic_exp_abs.
+rewrite scaled_mantissa_abs, cexp_abs.
rewrite Ztrunc_abs.
rewrite F2R_Zabs.
now apply f_equal.
@@ -308,7 +315,7 @@ Proof.
intros x.
unfold generic_format, Rabs.
case Rcase_abs ; intros _.
-rewrite scaled_mantissa_opp, canonic_exp_opp, Ztrunc_opp.
+rewrite scaled_mantissa_opp, cexp_opp, Ztrunc_opp.
intros H.
rewrite <- (Ropp_involutive x) at 1.
rewrite H, F2R_Zopp.
@@ -316,23 +323,23 @@ apply Ropp_involutive.
easy.
Qed.
-Theorem canonic_exp_fexp :
+Theorem cexp_fexp :
forall x ex,
(bpow (ex - 1) <= Rabs x < bpow ex)%R ->
- canonic_exp x = fexp ex.
+ cexp x = fexp ex.
Proof.
intros x ex Hx.
-unfold canonic_exp.
-now rewrite ln_beta_unique with (1 := Hx).
+unfold cexp.
+now rewrite mag_unique with (1 := Hx).
Qed.
-Theorem canonic_exp_fexp_pos :
+Theorem cexp_fexp_pos :
forall x ex,
(bpow (ex - 1) <= x < bpow ex)%R ->
- canonic_exp x = fexp ex.
+ cexp x = fexp ex.
Proof.
intros x ex Hx.
-apply canonic_exp_fexp.
+apply cexp_fexp.
rewrite Rabs_pos_eq.
exact Hx.
apply Rle_trans with (2 := proj1 Hx).
@@ -360,7 +367,7 @@ apply Rlt_le_trans with (1 := proj2 Hx).
now apply bpow_le.
Qed.
-Theorem scaled_mantissa_small :
+Theorem scaled_mantissa_lt_1 :
forall x ex,
(Rabs x < bpow ex)%R ->
(ex <= fexp ex)%Z ->
@@ -369,62 +376,62 @@ Proof.
intros x ex Ex He.
destruct (Req_dec x 0) as [Zx|Zx].
rewrite Zx, scaled_mantissa_0, Rabs_R0.
-now apply (Z2R_lt 0 1).
+now apply IZR_lt.
rewrite <- scaled_mantissa_abs.
unfold scaled_mantissa.
-rewrite canonic_exp_abs.
-unfold canonic_exp.
-destruct (ln_beta beta x) as (ex', Ex').
+rewrite cexp_abs.
+unfold cexp.
+destruct (mag beta x) as (ex', Ex').
simpl.
specialize (Ex' Zx).
apply (mantissa_small_pos _ _ Ex').
assert (ex' <= fexp ex)%Z.
-apply Zle_trans with (2 := He).
+apply Z.le_trans with (2 := He).
apply bpow_lt_bpow with beta.
now apply Rle_lt_trans with (2 := Ex).
now rewrite (proj2 (proj2 (valid_exp _) He)).
Qed.
-Theorem abs_scaled_mantissa_lt_bpow :
+Theorem scaled_mantissa_lt_bpow :
forall x,
- (Rabs (scaled_mantissa x) < bpow (ln_beta beta x - canonic_exp x))%R.
+ (Rabs (scaled_mantissa x) < bpow (mag beta x - cexp x))%R.
Proof.
intros x.
destruct (Req_dec x 0) as [Zx|Zx].
rewrite Zx, scaled_mantissa_0, Rabs_R0.
apply bpow_gt_0.
-apply Rlt_le_trans with (1 := bpow_ln_beta_gt beta _).
+apply Rlt_le_trans with (1 := bpow_mag_gt beta _).
apply bpow_le.
unfold scaled_mantissa.
-rewrite ln_beta_mult_bpow with (1 := Zx).
-apply Zle_refl.
+rewrite mag_mult_bpow with (1 := Zx).
+apply Z.le_refl.
Qed.
-Theorem ln_beta_generic_gt :
+Theorem mag_generic_gt :
forall x, (x <> 0)%R ->
generic_format x ->
- (canonic_exp x < ln_beta beta x)%Z.
+ (cexp x < mag beta x)%Z.
Proof.
intros x Zx Gx.
apply Znot_ge_lt.
-unfold canonic_exp.
-destruct (ln_beta beta x) as (ex,Ex) ; simpl.
+unfold cexp.
+destruct (mag beta x) as (ex,Ex) ; simpl.
specialize (Ex Zx).
intros H.
-apply Zge_le in H.
-generalize (scaled_mantissa_small x ex (proj2 Ex) H).
+apply Z.ge_le in H.
+generalize (scaled_mantissa_lt_1 x ex (proj2 Ex) H).
contradict Zx.
rewrite Gx.
replace (Ztrunc (scaled_mantissa x)) with Z0.
apply F2R_0.
-cut (Zabs (Ztrunc (scaled_mantissa x)) < 1)%Z.
+cut (Z.abs (Ztrunc (scaled_mantissa x)) < 1)%Z.
clear ; zify ; omega.
-apply lt_Z2R.
-rewrite Z2R_abs.
+apply lt_IZR.
+rewrite abs_IZR.
now rewrite <- scaled_mantissa_generic.
Qed.
-Theorem mantissa_DN_small_pos :
+Lemma mantissa_DN_small_pos :
forall x ex,
(bpow (ex - 1) <= x < bpow ex)%R ->
(ex <= fexp ex)%Z ->
@@ -436,7 +443,7 @@ assert (H := mantissa_small_pos x ex Hx He).
split ; try apply Rlt_le ; apply H.
Qed.
-Theorem mantissa_UP_small_pos :
+Lemma mantissa_UP_small_pos :
forall x ex,
(bpow (ex - 1) <= x < bpow ex)%R ->
(ex <= fexp ex)%Z ->
@@ -451,7 +458,7 @@ Qed.
(** Generic facts about any format *)
Theorem generic_format_discrete :
forall x m,
- let e := canonic_exp x in
+ let e := cexp x in
(F2R (Float beta m e) < x < F2R (Float beta (m + 1) e))%R ->
~ generic_format x.
Proof.
@@ -459,27 +466,27 @@ intros x m e (Hx,Hx2) Hf.
apply Rlt_not_le with (1 := Hx2). clear Hx2.
rewrite Hf.
fold e.
-apply F2R_le_compat.
+apply F2R_le.
apply Zlt_le_succ.
-apply lt_Z2R.
+apply lt_IZR.
rewrite <- scaled_mantissa_generic with (1 := Hf).
apply Rmult_lt_reg_r with (bpow e).
apply bpow_gt_0.
now rewrite scaled_mantissa_mult_bpow.
Qed.
-Theorem generic_format_canonic :
- forall f, canonic f ->
+Theorem generic_format_canonical :
+ forall f, canonical f ->
generic_format (F2R f).
Proof.
intros (m, e) Hf.
-unfold canonic in Hf. simpl in Hf.
+unfold canonical in Hf. simpl in Hf.
unfold generic_format, scaled_mantissa.
rewrite <- Hf.
-apply F2R_eq_compat.
+apply F2R_eq.
unfold F2R. simpl.
rewrite Rmult_assoc, <- bpow_plus, Zplus_opp_r, Rmult_1_r.
-now rewrite Ztrunc_Z2R.
+now rewrite Ztrunc_IZR.
Qed.
Theorem generic_format_ge_bpow :
@@ -492,10 +499,10 @@ Theorem generic_format_ge_bpow :
Proof.
intros emin Emin x Hx Fx.
rewrite Fx.
-apply Rle_trans with (bpow (fexp (ln_beta beta x))).
+apply Rle_trans with (bpow (fexp (mag beta x))).
now apply bpow_le.
apply bpow_le_F2R.
-apply F2R_gt_0_reg with beta (canonic_exp x).
+apply gt_0_F2R with beta (cexp x).
now rewrite <- Fx.
Qed.
@@ -504,13 +511,13 @@ Theorem abs_lt_bpow_prec:
(forall e, (e - prec <= fexp e)%Z) ->
(* OK with FLX, FLT and FTZ *)
forall x,
- (Rabs x < bpow (prec + canonic_exp x))%R.
+ (Rabs x < bpow (prec + cexp x))%R.
intros prec Hp x.
case (Req_dec x 0); intros Hxz.
rewrite Hxz, Rabs_R0.
apply bpow_gt_0.
-unfold canonic_exp.
-destruct (ln_beta beta x) as (ex,Ex) ; simpl.
+unfold cexp.
+destruct (mag beta x) as (ex,Ex) ; simpl.
specialize (Ex Hxz).
apply Rlt_le_trans with (1 := proj2 Ex).
apply bpow_le.
@@ -526,8 +533,8 @@ Proof.
intros e He.
apply Znot_gt_le.
contradict He.
-unfold generic_format, scaled_mantissa, canonic_exp, F2R. simpl.
-rewrite ln_beta_bpow, <- bpow_plus.
+unfold generic_format, scaled_mantissa, cexp, F2R. simpl.
+rewrite mag_bpow, <- bpow_plus.
apply Rgt_not_eq.
rewrite Ztrunc_floor.
2: apply bpow_ge_0.
@@ -559,7 +566,7 @@ Variable rnd : R -> Z.
Class Valid_rnd := {
Zrnd_le : forall x y, (x <= y)%R -> (rnd x <= rnd y)%Z ;
- Zrnd_Z2R : forall n, rnd (Z2R n) = n
+ Zrnd_IZR : forall n, rnd (IZR n) = n
}.
Context { valid_rnd : Valid_rnd }.
@@ -571,20 +578,20 @@ intros x.
destruct (Zle_or_lt (rnd x) (Zfloor x)) as [Hx|Hx].
left.
apply Zle_antisym with (1 := Hx).
-rewrite <- (Zrnd_Z2R (Zfloor x)).
+rewrite <- (Zrnd_IZR (Zfloor x)).
apply Zrnd_le.
apply Zfloor_lb.
right.
apply Zle_antisym.
-rewrite <- (Zrnd_Z2R (Zceil x)).
+rewrite <- (Zrnd_IZR (Zceil x)).
apply Zrnd_le.
apply Zceil_ub.
rewrite Zceil_floor_neq.
omega.
intros H.
rewrite <- H in Hx.
-rewrite Zfloor_Z2R, Zrnd_Z2R in Hx.
-apply Zlt_irrefl with (1 := Hx).
+rewrite Zfloor_IZR, Zrnd_IZR in Hx.
+apply Z.lt_irrefl with (1 := Hx).
Qed.
Theorem Zrnd_ZR_or_AW :
@@ -602,7 +609,7 @@ Qed.
(** the most useful one: R -> F *)
Definition round x :=
- F2R (Float beta (rnd (scaled_mantissa x)) (canonic_exp x)).
+ F2R (Float beta (rnd (scaled_mantissa x)) (cexp x)).
Theorem round_bounded_large_pos :
forall x ex,
@@ -612,7 +619,7 @@ Theorem round_bounded_large_pos :
Proof.
intros x ex He Hx.
unfold round, scaled_mantissa.
-rewrite (canonic_exp_fexp_pos _ _ Hx).
+rewrite (cexp_fexp_pos _ _ Hx).
unfold F2R. simpl.
destruct (Zrnd_DN_or_UP (x * bpow (- fexp ex))) as [Hr|Hr] ; rewrite Hr.
(* DN *)
@@ -621,11 +628,11 @@ replace (ex - 1)%Z with (ex - 1 + - fexp ex + fexp ex)%Z by ring.
rewrite bpow_plus.
apply Rmult_le_compat_r.
apply bpow_ge_0.
-assert (Hf: Z2R (Zpower beta (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)).
-apply Z2R_Zpower.
+assert (Hf: IZR (Zpower beta (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)).
+apply IZR_Zpower.
omega.
rewrite <- Hf.
-apply Z2R_le.
+apply IZR_le.
apply Zfloor_lub.
rewrite Hf.
rewrite bpow_plus.
@@ -648,11 +655,11 @@ pattern ex at 3 ; replace ex with (ex - fexp ex + fexp ex)%Z by ring.
rewrite bpow_plus.
apply Rmult_le_compat_r.
apply bpow_ge_0.
-assert (Hf: Z2R (Zpower beta (ex - fexp ex)) = bpow (ex - fexp ex)).
-apply Z2R_Zpower.
+assert (Hf: IZR (Zpower beta (ex - fexp ex)) = bpow (ex - fexp ex)).
+apply IZR_Zpower.
omega.
rewrite <- Hf.
-apply Z2R_le.
+apply IZR_le.
apply Zceil_glb.
rewrite Hf.
unfold Zminus.
@@ -671,13 +678,13 @@ Theorem round_bounded_small_pos :
Proof.
intros x ex He Hx.
unfold round, scaled_mantissa.
-rewrite (canonic_exp_fexp_pos _ _ Hx).
+rewrite (cexp_fexp_pos _ _ Hx).
unfold F2R. simpl.
destruct (Zrnd_DN_or_UP (x * bpow (-fexp ex))) as [Hr|Hr] ; rewrite Hr.
(* DN *)
left.
apply Rmult_eq_0_compat_r.
-apply (@f_equal _ _ Z2R _ Z0).
+apply IZR_eq.
apply Zfloor_imp.
refine (let H := _ in conj (Rlt_le _ _ (proj1 H)) (proj2 H)).
now apply mantissa_small_pos.
@@ -685,18 +692,18 @@ now apply mantissa_small_pos.
right.
pattern (bpow (fexp ex)) at 2 ; rewrite <- Rmult_1_l.
apply (f_equal (fun m => (m * bpow (fexp ex))%R)).
-apply (@f_equal _ _ Z2R _ 1%Z).
+apply IZR_eq.
apply Zceil_imp.
refine (let H := _ in conj (proj1 H) (Rlt_le _ _ (proj2 H))).
now apply mantissa_small_pos.
Qed.
-Theorem round_le_pos :
+Lemma round_le_pos :
forall x y, (0 < x)%R -> (x <= y)%R -> (round x <= round y)%R.
Proof.
intros x y Hx Hxy.
-destruct (ln_beta beta x) as [ex Hex].
-destruct (ln_beta beta y) as [ey Hey].
+destruct (mag beta x) as [ex Hex].
+destruct (mag beta y) as [ey Hey].
specialize (Hex (Rgt_not_eq _ _ Hx)).
specialize (Hey (Rgt_not_eq _ _ (Rlt_le_trans _ _ _ Hx Hxy))).
rewrite Rabs_pos_eq in Hex.
@@ -709,18 +716,18 @@ assert (He: (ex <= ey)%Z).
now apply Rle_lt_trans with y.
assert (Heq: fexp ex = fexp ey -> (round x <= round y)%R).
intros H.
- unfold round, scaled_mantissa, canonic_exp.
- rewrite ln_beta_unique_pos with (1 := Hex).
- rewrite ln_beta_unique_pos with (1 := Hey).
+ unfold round, scaled_mantissa, cexp.
+ rewrite mag_unique_pos with (1 := Hex).
+ rewrite mag_unique_pos with (1 := Hey).
rewrite H.
- apply F2R_le_compat.
+ apply F2R_le.
apply Zrnd_le.
apply Rmult_le_compat_r with (2 := Hxy).
apply bpow_ge_0.
destruct (Zle_or_lt ey (fexp ey)) as [Hy1|Hy1].
apply Heq.
apply valid_exp with (1 := Hy1).
- now apply Zle_trans with ey.
+ now apply Z.le_trans with ey.
destruct (Zle_lt_or_eq _ _ He) as [He'|He'].
2: now apply Heq, f_equal.
apply Rle_trans with (bpow (ey - 1)).
@@ -746,7 +753,7 @@ Proof.
intros x Hx.
unfold round.
rewrite scaled_mantissa_generic with (1 := Hx).
-rewrite Zrnd_Z2R.
+rewrite Zrnd_IZR.
now apply sym_eq.
Qed.
@@ -755,8 +762,7 @@ Theorem round_0 :
Proof.
unfold round, scaled_mantissa.
rewrite Rmult_0_l.
-change 0%R with (Z2R 0).
-rewrite Zrnd_Z2R.
+rewrite Zrnd_IZR.
apply F2R_0.
Qed.
@@ -774,13 +780,13 @@ apply bpow_gt_0.
apply (round_bounded_large_pos); assumption.
Qed.
-Theorem generic_format_round_pos :
+Lemma generic_format_round_pos :
forall x,
(0 < x)%R ->
generic_format (round x).
Proof.
intros x Hx0.
-destruct (ln_beta beta x) as (ex, Hex).
+destruct (mag beta x) as (ex, Hex).
specialize (Hex (Rgt_not_eq _ _ Hx0)).
rewrite Rabs_pos_eq in Hex. 2: now apply Rlt_le.
destruct (Zle_or_lt ex (fexp ex)) as [He|He].
@@ -798,8 +804,8 @@ apply generic_format_bpow.
now apply valid_exp.
apply generic_format_F2R.
intros _.
-rewrite (canonic_exp_fexp_pos (F2R _) _ (conj Hr1 Hr)).
-rewrite (canonic_exp_fexp_pos _ _ Hex).
+rewrite (cexp_fexp_pos (F2R _) _ (conj Hr1 Hr)).
+rewrite (cexp_fexp_pos _ _ Hex).
now apply Zeq_le.
Qed.
@@ -821,7 +827,7 @@ Section Zround_opp.
Variable rnd : R -> Z.
Context { valid_rnd : Valid_rnd rnd }.
-Definition Zrnd_opp x := Zopp (rnd (-x)).
+Definition Zrnd_opp x := Z.opp (rnd (-x)).
Global Instance valid_rnd_opp : Valid_rnd Zrnd_opp.
Proof with auto with typeclass_instances.
@@ -830,14 +836,14 @@ split.
intros x y Hxy.
unfold Zrnd_opp.
apply Zopp_le_cancel.
-rewrite 2!Zopp_involutive.
+rewrite 2!Z.opp_involutive.
apply Zrnd_le...
now apply Ropp_le_contravar.
(* *)
intros n.
unfold Zrnd_opp.
-rewrite <- Z2R_opp, Zrnd_Z2R...
-apply Zopp_involutive.
+rewrite <- opp_IZR, Zrnd_IZR...
+apply Z.opp_involutive.
Qed.
Theorem round_opp :
@@ -846,10 +852,10 @@ Theorem round_opp :
Proof.
intros x.
unfold round.
-rewrite <- F2R_Zopp, canonic_exp_opp, scaled_mantissa_opp.
-apply F2R_eq_compat.
+rewrite <- F2R_Zopp, cexp_opp, scaled_mantissa_opp.
+apply F2R_eq.
apply sym_eq.
-exact (Zopp_involutive _).
+exact (Z.opp_involutive _).
Qed.
End Zround_opp.
@@ -860,28 +866,28 @@ Global Instance valid_rnd_DN : Valid_rnd Zfloor.
Proof.
split.
apply Zfloor_le.
-apply Zfloor_Z2R.
+apply Zfloor_IZR.
Qed.
Global Instance valid_rnd_UP : Valid_rnd Zceil.
Proof.
split.
apply Zceil_le.
-apply Zceil_Z2R.
+apply Zceil_IZR.
Qed.
Global Instance valid_rnd_ZR : Valid_rnd Ztrunc.
Proof.
split.
apply Ztrunc_le.
-apply Ztrunc_Z2R.
+apply Ztrunc_IZR.
Qed.
Global Instance valid_rnd_AW : Valid_rnd Zaway.
Proof.
split.
apply Zaway_le.
-apply Zaway_Z2R.
+apply Zaway_IZR.
Qed.
Section monotone.
@@ -923,7 +929,7 @@ destruct (Rlt_or_le y 0) as [Hy|Hy].
(* . y < 0 *)
rewrite <- (Ropp_involutive x), <- (Ropp_involutive y).
rewrite (scaled_mantissa_opp (-x)), (scaled_mantissa_opp (-y)).
-rewrite (canonic_exp_opp (-x)), (canonic_exp_opp (-y)).
+rewrite (cexp_opp (-x)), (cexp_opp (-y)).
apply Ropp_le_cancel.
rewrite <- 2!F2R_Zopp.
apply (round_le_pos (Zrnd_opp rnd) (-y) (-x)).
@@ -932,16 +938,16 @@ now apply Ropp_lt_contravar.
now apply Ropp_le_contravar.
(* . 0 <= y *)
apply Rle_trans with 0%R.
-apply F2R_le_0_compat. simpl.
-rewrite <- (Zrnd_Z2R rnd 0).
+apply F2R_le_0. simpl.
+rewrite <- (Zrnd_IZR rnd 0).
apply Zrnd_le...
simpl.
-rewrite <- (Rmult_0_l (bpow (- fexp (ln_beta beta x)))).
+rewrite <- (Rmult_0_l (bpow (- fexp (mag beta x)))).
apply Rmult_le_compat_r.
apply bpow_ge_0.
now apply Rlt_le.
-apply F2R_ge_0_compat. simpl.
-rewrite <- (Zrnd_Z2R rnd 0).
+apply F2R_ge_0. simpl.
+rewrite <- (Zrnd_IZR rnd 0).
apply Zrnd_le...
apply Rmult_le_pos.
exact Hy.
@@ -949,9 +955,9 @@ apply bpow_ge_0.
(* x = 0 *)
rewrite Hx.
rewrite round_0...
-apply F2R_ge_0_compat.
+apply F2R_ge_0.
simpl.
-rewrite <- (Zrnd_Z2R rnd 0).
+rewrite <- (Zrnd_IZR rnd 0).
apply Zrnd_le...
apply Rmult_le_pos.
now rewrite <- Hx.
@@ -1071,8 +1077,8 @@ unfold round.
rewrite scaled_mantissa_opp.
rewrite <- F2R_Zopp.
unfold Zceil.
-rewrite Zopp_involutive.
-now rewrite canonic_exp_opp.
+rewrite Z.opp_involutive.
+now rewrite cexp_opp.
Qed.
Theorem round_UP_opp :
@@ -1085,7 +1091,7 @@ rewrite scaled_mantissa_opp.
rewrite <- F2R_Zopp.
unfold Zceil.
rewrite Ropp_involutive.
-now rewrite canonic_exp_opp.
+now rewrite cexp_opp.
Qed.
Theorem round_ZR_opp :
@@ -1094,7 +1100,7 @@ Theorem round_ZR_opp :
Proof.
intros x.
unfold round.
-rewrite scaled_mantissa_opp, canonic_exp_opp, Ztrunc_opp.
+rewrite scaled_mantissa_opp, cexp_opp, Ztrunc_opp.
apply F2R_Zopp.
Qed.
@@ -1123,7 +1129,7 @@ Theorem round_AW_opp :
Proof.
intros x.
unfold round.
-rewrite scaled_mantissa_opp, canonic_exp_opp, Zaway_opp.
+rewrite scaled_mantissa_opp, cexp_opp, Zaway_opp.
apply F2R_Zopp.
Qed.
@@ -1146,7 +1152,7 @@ apply round_le...
now apply Rge_le.
Qed.
-Theorem round_ZR_pos :
+Theorem round_ZR_DN :
forall x,
(0 <= x)%R ->
round Ztrunc x = round Zfloor x.
@@ -1156,13 +1162,13 @@ unfold round, Ztrunc.
case Rlt_bool_spec.
intros H.
elim Rlt_not_le with (1 := H).
-rewrite <- (Rmult_0_l (bpow (- canonic_exp x))).
+rewrite <- (Rmult_0_l (bpow (- cexp x))).
apply Rmult_le_compat_r with (2 := Hx).
apply bpow_ge_0.
easy.
Qed.
-Theorem round_ZR_neg :
+Theorem round_ZR_UP :
forall x,
(x <= 0)%R ->
round Ztrunc x = round Zceil x.
@@ -1173,15 +1179,14 @@ case Rlt_bool_spec.
easy.
intros [H|H].
elim Rlt_not_le with (1 := H).
-rewrite <- (Rmult_0_l (bpow (- canonic_exp x))).
+rewrite <- (Rmult_0_l (bpow (- cexp x))).
apply Rmult_le_compat_r with (2 := Hx).
apply bpow_ge_0.
rewrite <- H.
-change 0%R with (Z2R 0).
-now rewrite Zfloor_Z2R, Zceil_Z2R.
+now rewrite Zfloor_IZR, Zceil_IZR.
Qed.
-Theorem round_AW_pos :
+Theorem round_AW_UP :
forall x,
(0 <= x)%R ->
round Zaway x = round Zceil x.
@@ -1191,13 +1196,13 @@ unfold round, Zaway.
case Rlt_bool_spec.
intros H.
elim Rlt_not_le with (1 := H).
-rewrite <- (Rmult_0_l (bpow (- canonic_exp x))).
+rewrite <- (Rmult_0_l (bpow (- cexp x))).
apply Rmult_le_compat_r with (2 := Hx).
apply bpow_ge_0.
easy.
Qed.
-Theorem round_AW_neg :
+Theorem round_AW_DN :
forall x,
(x <= 0)%R ->
round Zaway x = round Zfloor x.
@@ -1208,12 +1213,11 @@ case Rlt_bool_spec.
easy.
intros [H|H].
elim Rlt_not_le with (1 := H).
-rewrite <- (Rmult_0_l (bpow (- canonic_exp x))).
+rewrite <- (Rmult_0_l (bpow (- cexp x))).
apply Rmult_le_compat_r with (2 := Hx).
apply bpow_ge_0.
rewrite <- H.
-change 0%R with (Z2R 0).
-now rewrite Zfloor_Z2R, Zceil_Z2R.
+now rewrite Zfloor_IZR, Zceil_IZR.
Qed.
Theorem generic_format_round :
@@ -1275,7 +1279,7 @@ Proof.
intros x.
rewrite <- (Ropp_involutive x).
rewrite round_UP_opp.
-apply Rnd_DN_UP_pt_sym.
+apply Rnd_UP_pt_opp.
apply generic_format_opp.
apply round_DN_pt.
Qed.
@@ -1286,22 +1290,22 @@ Theorem round_ZR_pt :
Proof.
intros x.
split ; intros Hx.
-rewrite round_ZR_pos with (1 := Hx).
+rewrite round_ZR_DN with (1 := Hx).
apply round_DN_pt.
-rewrite round_ZR_neg with (1 := Hx).
+rewrite round_ZR_UP with (1 := Hx).
apply round_UP_pt.
Qed.
-Theorem round_DN_small_pos :
+Lemma round_DN_small_pos :
forall x ex,
(bpow (ex - 1) <= x < bpow ex)%R ->
(ex <= fexp ex)%Z ->
round Zfloor x = 0%R.
Proof.
intros x ex Hx He.
-rewrite <- (F2R_0 beta (canonic_exp x)).
+rewrite <- (F2R_0 beta (cexp x)).
rewrite <- mantissa_DN_small_pos with (1 := Hx) (2 := He).
-now rewrite <- canonic_exp_fexp_pos with (1 := Hx).
+now rewrite <- cexp_fexp_pos with (1 := Hx).
Qed.
@@ -1329,7 +1333,7 @@ contradict Fx.
apply generic_format_round...
Qed.
-Theorem round_UP_small_pos :
+Lemma round_UP_small_pos :
forall x ex,
(bpow (ex - 1) <= x < bpow ex)%R ->
(ex <= fexp ex)%Z ->
@@ -1338,7 +1342,7 @@ Proof.
intros x ex Hx He.
rewrite <- F2R_bpow.
rewrite <- mantissa_UP_small_pos with (1 := Hx) (2 := He).
-now rewrite <- canonic_exp_fexp_pos with (1 := Hx).
+now rewrite <- cexp_fexp_pos with (1 := Hx).
Qed.
Theorem generic_format_EM :
@@ -1361,14 +1365,14 @@ Section round_large.
Variable rnd : R -> Z.
Context { valid_rnd : Valid_rnd rnd }.
-Theorem round_large_pos_ge_pow :
+Lemma round_large_pos_ge_bpow :
forall x e,
(0 < round rnd x)%R ->
(bpow e <= x)%R ->
(bpow e <= round rnd x)%R.
Proof.
intros x e Hd Hex.
-destruct (ln_beta beta x) as (ex, He).
+destruct (mag beta x) as (ex, He).
assert (Hx: (0 < x)%R).
apply Rlt_le_trans with (2 := Hex).
apply bpow_gt_0.
@@ -1391,95 +1395,95 @@ Qed.
End round_large.
-Theorem ln_beta_round_ZR :
+Theorem mag_round_ZR :
forall x,
(round Ztrunc x <> 0)%R ->
- (ln_beta beta (round Ztrunc x) = ln_beta beta x :> Z).
+ (mag beta (round Ztrunc x) = mag beta x :> Z).
Proof with auto with typeclass_instances.
intros x Zr.
destruct (Req_dec x 0) as [Zx|Zx].
rewrite Zx, round_0...
-apply ln_beta_unique.
-destruct (ln_beta beta x) as (ex, Ex) ; simpl.
+apply mag_unique.
+destruct (mag beta x) as (ex, Ex) ; simpl.
specialize (Ex Zx).
rewrite <- round_ZR_abs.
split.
-apply round_large_pos_ge_pow...
+apply round_large_pos_ge_bpow...
rewrite round_ZR_abs.
now apply Rabs_pos_lt.
apply Ex.
apply Rle_lt_trans with (2 := proj2 Ex).
-rewrite round_ZR_pos.
+rewrite round_ZR_DN.
apply round_DN_pt.
apply Rabs_pos.
Qed.
-Theorem ln_beta_round :
+Theorem mag_round :
forall rnd {Hrnd : Valid_rnd rnd} x,
(round rnd x <> 0)%R ->
- (ln_beta beta (round rnd x) = ln_beta beta x :> Z) \/
- Rabs (round rnd x) = bpow (Zmax (ln_beta beta x) (fexp (ln_beta beta x))).
+ (mag beta (round rnd x) = mag beta x :> Z) \/
+ Rabs (round rnd x) = bpow (Z.max (mag beta x) (fexp (mag beta x))).
Proof with auto with typeclass_instances.
intros rnd Hrnd x.
destruct (round_ZR_or_AW rnd x) as [Hr|Hr] ; rewrite Hr ; clear Hr rnd Hrnd.
left.
-now apply ln_beta_round_ZR.
+now apply mag_round_ZR.
intros Zr.
destruct (Req_dec x 0) as [Zx|Zx].
rewrite Zx, round_0...
-destruct (ln_beta beta x) as (ex, Ex) ; simpl.
+destruct (mag beta x) as (ex, Ex) ; simpl.
specialize (Ex Zx).
-rewrite <- ln_beta_abs.
+rewrite <- mag_abs.
rewrite <- round_AW_abs.
destruct (Zle_or_lt ex (fexp ex)) as [He|He].
right.
-rewrite Zmax_r with (1 := He).
-rewrite round_AW_pos with (1 := Rabs_pos _).
+rewrite Z.max_r with (1 := He).
+rewrite round_AW_UP with (1 := Rabs_pos _).
now apply round_UP_small_pos.
destruct (round_bounded_large_pos Zaway _ ex He Ex) as (H1,[H2|H2]).
left.
-apply ln_beta_unique.
+apply mag_unique.
rewrite <- round_AW_abs, Rabs_Rabsolu.
now split.
right.
-now rewrite Zmax_l with (1 := Zlt_le_weak _ _ He).
+now rewrite Z.max_l with (1 := Zlt_le_weak _ _ He).
Qed.
-Theorem ln_beta_DN :
+Theorem mag_DN :
forall x,
(0 < round Zfloor x)%R ->
- (ln_beta beta (round Zfloor x) = ln_beta beta x :> Z).
+ (mag beta (round Zfloor x) = mag beta x :> Z).
Proof.
intros x Hd.
assert (0 < x)%R.
apply Rlt_le_trans with (1 := Hd).
apply round_DN_pt.
revert Hd.
-rewrite <- round_ZR_pos.
+rewrite <- round_ZR_DN.
intros Hd.
-apply ln_beta_round_ZR.
+apply mag_round_ZR.
now apply Rgt_not_eq.
now apply Rlt_le.
Qed.
-Theorem canonic_exp_DN :
+Theorem cexp_DN :
forall x,
(0 < round Zfloor x)%R ->
- canonic_exp (round Zfloor x) = canonic_exp x.
+ cexp (round Zfloor x) = cexp x.
Proof.
intros x Hd.
apply (f_equal fexp).
-now apply ln_beta_DN.
+now apply mag_DN.
Qed.
Theorem scaled_mantissa_DN :
forall x,
(0 < round Zfloor x)%R ->
- scaled_mantissa (round Zfloor x) = Z2R (Zfloor (scaled_mantissa x)).
+ scaled_mantissa (round Zfloor x) = IZR (Zfloor (scaled_mantissa x)).
Proof.
intros x Hd.
unfold scaled_mantissa.
-rewrite canonic_exp_DN with (1 := Hd).
+rewrite cexp_DN with (1 := Hd).
unfold round, F2R. simpl.
now rewrite Rmult_assoc, <- bpow_plus, Zplus_opp_r, Rmult_1_r.
Qed.
@@ -1492,10 +1496,10 @@ Proof.
intros x f Hxf.
destruct (Rnd_N_pt_DN_or_UP _ _ _ Hxf).
left.
-apply Rnd_DN_pt_unicity with (1 := H).
+apply Rnd_DN_pt_unique with (1 := H).
apply round_DN_pt.
right.
-apply Rnd_UP_pt_unicity with (1 := H).
+apply Rnd_UP_pt_unique with (1 := H).
apply round_UP_pt.
Qed.
@@ -1516,20 +1520,20 @@ intros e x He Hx.
pattern x at 2 ; rewrite Hx.
unfold F2R at 2. simpl.
rewrite Rmult_assoc, <- bpow_plus.
-assert (H: Z2R (Zpower beta (canonic_exp x + - fexp e)) = bpow (canonic_exp x + - fexp e)).
-apply Z2R_Zpower.
-unfold canonic_exp.
-set (ex := ln_beta beta x).
+assert (H: IZR (Zpower beta (cexp x + - fexp e)) = bpow (cexp x + - fexp e)).
+apply IZR_Zpower.
+unfold cexp.
+set (ex := mag beta x).
generalize (exp_not_FTZ ex).
generalize (proj2 (proj2 (valid_exp _) He) (fexp ex + 1)%Z).
omega.
rewrite <- H.
-rewrite <- Z2R_mult, Ztrunc_Z2R.
+rewrite <- mult_IZR, Ztrunc_IZR.
unfold F2R. simpl.
-rewrite Z2R_mult.
+rewrite mult_IZR.
rewrite H.
rewrite Rmult_assoc, <- bpow_plus.
-now ring_simplify (canonic_exp x + - fexp e + fexp e)%Z.
+now ring_simplify (cexp x + - fexp e + fexp e)%Z.
Qed.
End not_FTZ.
@@ -1550,60 +1554,60 @@ now apply Zlt_le_succ.
now apply valid_exp.
Qed.
-Lemma canonic_exp_le_bpow :
+Lemma cexp_le_bpow :
forall (x : R) (e : Z),
x <> 0%R ->
(Rabs x < bpow e)%R ->
- (canonic_exp x <= fexp e)%Z.
+ (cexp x <= fexp e)%Z.
Proof.
intros x e Zx Hx.
apply monotone_exp.
-now apply ln_beta_le_bpow.
+now apply mag_le_bpow.
Qed.
-Lemma canonic_exp_ge_bpow :
+Lemma cexp_ge_bpow :
forall (x : R) (e : Z),
(bpow (e - 1) <= Rabs x)%R ->
- (fexp e <= canonic_exp x)%Z.
+ (fexp e <= cexp x)%Z.
Proof.
intros x e Hx.
apply monotone_exp.
rewrite (Zsucc_pred e).
apply Zlt_le_succ.
-now apply ln_beta_gt_bpow.
+now apply mag_gt_bpow.
Qed.
Variable rnd : R -> Z.
Context { valid_rnd : Valid_rnd rnd }.
-Theorem ln_beta_round_ge :
+Theorem mag_round_ge :
forall x,
round rnd x <> 0%R ->
- (ln_beta beta x <= ln_beta beta (round rnd x))%Z.
+ (mag beta x <= mag beta (round rnd x))%Z.
Proof with auto with typeclass_instances.
intros x.
destruct (round_ZR_or_AW rnd x) as [H|H] ; rewrite H ; clear H ; intros Zr.
-rewrite ln_beta_round_ZR with (1 := Zr).
-apply Zle_refl.
-apply ln_beta_le_abs.
+rewrite mag_round_ZR with (1 := Zr).
+apply Z.le_refl.
+apply mag_le_abs.
contradict Zr.
rewrite Zr.
apply round_0...
rewrite <- round_AW_abs.
-rewrite round_AW_pos.
+rewrite round_AW_UP.
apply round_UP_pt.
apply Rabs_pos.
Qed.
-Theorem canonic_exp_round_ge :
+Theorem cexp_round_ge :
forall x,
round rnd x <> 0%R ->
- (canonic_exp x <= canonic_exp (round rnd x))%Z.
+ (cexp x <= cexp (round rnd x))%Z.
Proof with auto with typeclass_instances.
intros x Zr.
-unfold canonic_exp.
+unfold cexp.
apply monotone_exp.
-now apply ln_beta_round_ge.
+now apply mag_round_ge.
Qed.
End monotone_exp.
@@ -1614,7 +1618,7 @@ Section Znearest.
Variable choice : Z -> bool.
Definition Znearest x :=
- match Rcompare (x - Z2R (Zfloor x)) (/2) with
+ match Rcompare (x - IZR (Zfloor x)) (/2) with
| Lt => Zfloor x
| Eq => if choice (Zfloor x) then Zceil x else Zfloor x
| Gt => Zceil x
@@ -1640,8 +1644,8 @@ Theorem Znearest_ge_floor :
Proof.
intros x.
destruct (Znearest_DN_or_UP x) as [Hx|Hx] ; rewrite Hx.
-apply Zle_refl.
-apply le_Z2R.
+apply Z.le_refl.
+apply le_IZR.
apply Rle_trans with x.
apply Zfloor_lb.
apply Zceil_ub.
@@ -1653,11 +1657,11 @@ Theorem Znearest_le_ceil :
Proof.
intros x.
destruct (Znearest_DN_or_UP x) as [Hx|Hx] ; rewrite Hx.
-apply le_Z2R.
+apply le_IZR.
apply Rle_trans with x.
apply Zfloor_lb.
apply Zceil_ub.
-apply Zle_refl.
+apply Z.le_refl.
Qed.
Global Instance valid_rnd_N : Valid_rnd Znearest.
@@ -1665,22 +1669,22 @@ Proof.
split.
(* *)
intros x y Hxy.
-destruct (Rle_or_lt (Z2R (Zceil x)) y) as [H|H].
-apply Zle_trans with (1 := Znearest_le_ceil x).
-apply Zle_trans with (2 := Znearest_ge_floor y).
+destruct (Rle_or_lt (IZR (Zceil x)) y) as [H|H].
+apply Z.le_trans with (1 := Znearest_le_ceil x).
+apply Z.le_trans with (2 := Znearest_ge_floor y).
now apply Zfloor_lub.
(* . *)
assert (Hf: Zfloor y = Zfloor x).
apply Zfloor_imp.
split.
apply Rle_trans with (2 := Zfloor_lb y).
-apply Z2R_le.
+apply IZR_le.
now apply Zfloor_le.
apply Rlt_le_trans with (1 := H).
-apply Z2R_le.
+apply IZR_le.
apply Zceil_glb.
apply Rlt_le.
-rewrite Z2R_plus.
+rewrite plus_IZR.
apply Zfloor_ub.
(* . *)
unfold Znearest at 1.
@@ -1696,15 +1700,15 @@ elim Rlt_not_le with (1 := Hy).
rewrite <- Hx.
now apply Rplus_le_compat_r.
replace y with x.
-apply Zle_refl.
-apply Rplus_eq_reg_l with (- Z2R (Zfloor x))%R.
-rewrite 2!(Rplus_comm (- (Z2R (Zfloor x)))).
-change (x - Z2R (Zfloor x) = y - Z2R (Zfloor x))%R.
+apply Z.le_refl.
+apply Rplus_eq_reg_l with (- IZR (Zfloor x))%R.
+rewrite 2!(Rplus_comm (- (IZR (Zfloor x)))).
+change (x - IZR (Zfloor x) = y - IZR (Zfloor x))%R.
now rewrite Hy.
-apply Zle_trans with (Zceil x).
+apply Z.le_trans with (Zceil x).
case choice.
-apply Zle_refl.
-apply le_Z2R.
+apply Z.le_refl.
+apply le_IZR.
apply Rle_trans with x.
apply Zfloor_lb.
apply Zceil_ub.
@@ -1719,79 +1723,19 @@ now apply Rplus_le_compat_r.
(* *)
intros n.
unfold Znearest.
-rewrite Zfloor_Z2R.
+rewrite Zfloor_IZR.
rewrite Rcompare_Lt.
easy.
unfold Rminus.
rewrite Rplus_opp_r.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
-Qed.
-
-Theorem Rcompare_floor_ceil_mid :
- forall x,
- Z2R (Zfloor x) <> x ->
- Rcompare (x - Z2R (Zfloor x)) (/ 2) = Rcompare (x - Z2R (Zfloor x)) (Z2R (Zceil x) - x).
-Proof.
-intros x Hx.
-rewrite Zceil_floor_neq with (1 := Hx).
-rewrite Z2R_plus. simpl.
-destruct (Rcompare_spec (x - Z2R (Zfloor x)) (/ 2)) as [H1|H1|H1] ; apply sym_eq.
-(* . *)
-apply Rcompare_Lt.
-apply Rplus_lt_reg_l with (x - Z2R (Zfloor x))%R.
-replace (x - Z2R (Zfloor x) + (x - Z2R (Zfloor x)))%R with ((x - Z2R (Zfloor x)) * 2)%R by ring.
-replace (x - Z2R (Zfloor x) + (Z2R (Zfloor x) + 1 - x))%R with (/2 * 2)%R by field.
-apply Rmult_lt_compat_r with (2 := H1).
-now apply (Z2R_lt 0 2).
-(* . *)
-apply Rcompare_Eq.
-replace (Z2R (Zfloor x) + 1 - x)%R with (1 - (x - Z2R (Zfloor x)))%R by ring.
-rewrite H1.
-field.
-(* . *)
-apply Rcompare_Gt.
-apply Rplus_lt_reg_l with (x - Z2R (Zfloor x))%R.
-replace (x - Z2R (Zfloor x) + (x - Z2R (Zfloor x)))%R with ((x - Z2R (Zfloor x)) * 2)%R by ring.
-replace (x - Z2R (Zfloor x) + (Z2R (Zfloor x) + 1 - x))%R with (/2 * 2)%R by field.
-apply Rmult_lt_compat_r with (2 := H1).
-now apply (Z2R_lt 0 2).
-Qed.
-
-Theorem Rcompare_ceil_floor_mid :
- forall x,
- Z2R (Zfloor x) <> x ->
- Rcompare (Z2R (Zceil x) - x) (/ 2) = Rcompare (Z2R (Zceil x) - x) (x - Z2R (Zfloor x)).
-Proof.
-intros x Hx.
-rewrite Zceil_floor_neq with (1 := Hx).
-rewrite Z2R_plus. simpl.
-destruct (Rcompare_spec (Z2R (Zfloor x) + 1 - x) (/ 2)) as [H1|H1|H1] ; apply sym_eq.
-(* . *)
-apply Rcompare_Lt.
-apply Rplus_lt_reg_l with (Z2R (Zfloor x) + 1 - x)%R.
-replace (Z2R (Zfloor x) + 1 - x + (Z2R (Zfloor x) + 1 - x))%R with ((Z2R (Zfloor x) + 1 - x) * 2)%R by ring.
-replace (Z2R (Zfloor x) + 1 - x + (x - Z2R (Zfloor x)))%R with (/2 * 2)%R by field.
-apply Rmult_lt_compat_r with (2 := H1).
-now apply (Z2R_lt 0 2).
-(* . *)
-apply Rcompare_Eq.
-replace (x - Z2R (Zfloor x))%R with (1 - (Z2R (Zfloor x) + 1 - x))%R by ring.
-rewrite H1.
-field.
-(* . *)
-apply Rcompare_Gt.
-apply Rplus_lt_reg_l with (Z2R (Zfloor x) + 1 - x)%R.
-replace (Z2R (Zfloor x) + 1 - x + (Z2R (Zfloor x) + 1 - x))%R with ((Z2R (Zfloor x) + 1 - x) * 2)%R by ring.
-replace (Z2R (Zfloor x) + 1 - x + (x - Z2R (Zfloor x)))%R with (/2 * 2)%R by field.
-apply Rmult_lt_compat_r with (2 := H1).
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
Qed.
Theorem Znearest_N_strict :
forall x,
- (x - Z2R (Zfloor x) <> /2)%R ->
- (Rabs (x - Z2R (Znearest x)) < /2)%R.
+ (x - IZR (Zfloor x) <> /2)%R ->
+ (Rabs (x - IZR (Znearest x)) < /2)%R.
Proof.
intros x Hx.
unfold Znearest.
@@ -1804,72 +1748,70 @@ now elim Hx.
rewrite Rabs_left1.
rewrite Ropp_minus_distr.
rewrite Zceil_floor_neq.
-rewrite Z2R_plus.
-simpl.
+rewrite plus_IZR.
apply Ropp_lt_cancel.
apply Rplus_lt_reg_l with 1%R.
replace (1 + -/2)%R with (/2)%R by field.
-now replace (1 + - (Z2R (Zfloor x) + 1 - x))%R with (x - Z2R (Zfloor x))%R by ring.
+now replace (1 + - (IZR (Zfloor x) + 1 - x))%R with (x - IZR (Zfloor x))%R by ring.
apply Rlt_not_eq.
-apply Rplus_lt_reg_l with (- Z2R (Zfloor x))%R.
+apply Rplus_lt_reg_l with (- IZR (Zfloor x))%R.
apply Rlt_trans with (/2)%R.
rewrite Rplus_opp_l.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
now rewrite <- (Rplus_comm x).
apply Rle_minus.
apply Zceil_ub.
Qed.
-Theorem Znearest_N :
+Theorem Znearest_half :
forall x,
- (Rabs (x - Z2R (Znearest x)) <= /2)%R.
+ (Rabs (x - IZR (Znearest x)) <= /2)%R.
Proof.
intros x.
-destruct (Req_dec (x - Z2R (Zfloor x)) (/2)) as [Hx|Hx].
+destruct (Req_dec (x - IZR (Zfloor x)) (/2)) as [Hx|Hx].
assert (K: (Rabs (/2) <= /2)%R).
apply Req_le.
apply Rabs_pos_eq.
apply Rlt_le.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
destruct (Znearest_DN_or_UP x) as [H|H] ; rewrite H ; clear H.
now rewrite Hx.
rewrite Zceil_floor_neq.
-rewrite Z2R_plus.
-simpl.
-replace (x - (Z2R (Zfloor x) + 1))%R with (x - Z2R (Zfloor x) - 1)%R by ring.
+rewrite plus_IZR.
+replace (x - (IZR (Zfloor x) + 1))%R with (x - IZR (Zfloor x) - 1)%R by ring.
rewrite Hx.
rewrite Rabs_minus_sym.
now replace (1 - /2)%R with (/2)%R by field.
apply Rlt_not_eq.
-apply Rplus_lt_reg_l with (- Z2R (Zfloor x))%R.
+apply Rplus_lt_reg_l with (- IZR (Zfloor x))%R.
rewrite Rplus_opp_l, Rplus_comm.
-fold (x - Z2R (Zfloor x))%R.
+fold (x - IZR (Zfloor x))%R.
rewrite Hx.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
apply Rlt_le.
now apply Znearest_N_strict.
Qed.
Theorem Znearest_imp :
forall x n,
- (Rabs (x - Z2R n) < /2)%R ->
+ (Rabs (x - IZR n) < /2)%R ->
Znearest x = n.
Proof.
intros x n Hd.
-cut (Zabs (Znearest x - n) < 1)%Z.
+cut (Z.abs (Znearest x - n) < 1)%Z.
clear ; zify ; omega.
-apply lt_Z2R.
-rewrite Z2R_abs, Z2R_minus.
-replace (Z2R (Znearest x) - Z2R n)%R with (- (x - Z2R (Znearest x)) + (x - Z2R n))%R by ring.
+apply lt_IZR.
+rewrite abs_IZR, minus_IZR.
+replace (IZR (Znearest x) - IZR n)%R with (- (x - IZR (Znearest x)) + (x - IZR n))%R by ring.
apply Rle_lt_trans with (1 := Rabs_triang _ _).
simpl.
replace 1%R with (/2 + /2)%R by field.
apply Rplus_le_lt_compat with (2 := Hd).
rewrite Rabs_Ropp.
-apply Znearest_N.
+apply Znearest_half.
Qed.
Theorem round_N_pt :
@@ -1880,7 +1822,7 @@ intros x.
set (d := round Zfloor x).
set (u := round Zceil x).
set (mx := scaled_mantissa x).
-set (bx := bpow (canonic_exp x)).
+set (bx := bpow (cexp x)).
(* . *)
assert (H: (Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R).
pattern x at -1 ; rewrite <- scaled_mantissa_mult_bpow.
@@ -1892,7 +1834,7 @@ rewrite <- Rmult_min_distr_r. 2: apply bpow_ge_0.
apply Rmult_le_compat_r.
apply bpow_ge_0.
unfold Znearest.
-destruct (Req_dec (Z2R (Zfloor mx)) mx) as [Hm|Hm].
+destruct (Req_dec (IZR (Zfloor mx)) mx) as [Hm|Hm].
(* .. *)
rewrite Hm.
unfold Rminus at 2.
@@ -1903,16 +1845,16 @@ unfold Rminus at -3.
rewrite Rplus_opp_r.
rewrite Rabs_R0.
unfold Rmin.
-destruct (Rle_dec 0 (Z2R (Zceil mx) - mx)) as [H|H].
+destruct (Rle_dec 0 (IZR (Zceil mx) - mx)) as [H|H].
apply Rle_refl.
apply Rle_0_minus.
apply Zceil_ub.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
(* .. *)
-rewrite Rcompare_floor_ceil_mid with (1 := Hm).
+rewrite Rcompare_floor_ceil_middle with (1 := Hm).
rewrite Rmin_compare.
-assert (H: (Rabs (mx - Z2R (Zfloor mx)) <= mx - Z2R (Zfloor mx))%R).
+assert (H: (Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R).
rewrite Rabs_pos_eq.
apply Rle_refl.
apply Rle_0_minus.
@@ -1928,7 +1870,7 @@ apply Rle_refl.
apply Rle_0_minus.
apply Zceil_ub.
(* . *)
-apply Rnd_DN_UP_pt_N with d u.
+apply Rnd_N_pt_DN_UP with d u.
apply generic_format_round.
auto with typeclass_instances.
now apply round_DN_pt.
@@ -1947,63 +1889,63 @@ Proof.
intros x.
pattern x at 1 4 ; rewrite <- scaled_mantissa_mult_bpow.
unfold round, Znearest, F2R. simpl.
-destruct (Req_dec (Z2R (Zfloor (scaled_mantissa x))) (scaled_mantissa x)) as [Fx|Fx].
+destruct (Req_dec (IZR (Zfloor (scaled_mantissa x))) (scaled_mantissa x)) as [Fx|Fx].
(* *)
intros _.
rewrite <- Fx.
-rewrite Zceil_Z2R, Zfloor_Z2R.
+rewrite Zceil_IZR, Zfloor_IZR.
set (m := Zfloor (scaled_mantissa x)).
-now case (Rcompare (Z2R m - Z2R m) (/ 2)) ; case (choice m).
+now case (Rcompare (IZR m - IZR m) (/ 2)) ; case (choice m).
(* *)
intros H.
-rewrite Rcompare_floor_ceil_mid with (1 := Fx).
+rewrite Rcompare_floor_ceil_middle with (1 := Fx).
rewrite Rcompare_Eq.
now case choice.
-apply Rmult_eq_reg_r with (bpow (canonic_exp x)).
+apply Rmult_eq_reg_r with (bpow (cexp x)).
now rewrite 2!Rmult_minus_distr_r.
apply Rgt_not_eq.
apply bpow_gt_0.
Qed.
-Lemma round_N_really_small_pos :
+Lemma round_N_small_pos :
forall x,
forall ex,
- (Fcore_Raux.bpow beta (ex - 1) <= x < Fcore_Raux.bpow beta ex)%R ->
+ (Raux.bpow beta (ex - 1) <= x < Raux.bpow beta ex)%R ->
(ex < fexp ex)%Z ->
(round Znearest x = 0)%R.
Proof.
intros x ex Hex Hf.
-unfold round, F2R, scaled_mantissa, canonic_exp; simpl.
-apply (Rmult_eq_reg_r (bpow (- fexp (ln_beta beta x))));
+unfold round, F2R, scaled_mantissa, cexp; simpl.
+apply (Rmult_eq_reg_r (bpow (- fexp (mag beta x))));
[|now apply Rgt_not_eq; apply bpow_gt_0].
rewrite Rmult_0_l, Rmult_assoc, <- bpow_plus.
replace (_ + - _)%Z with 0%Z by ring; simpl; rewrite Rmult_1_r.
-change 0%R with (Z2R 0); apply f_equal.
+apply IZR_eq.
apply Znearest_imp.
-simpl; unfold Rminus; rewrite Ropp_0; rewrite Rplus_0_r.
+unfold Rminus; rewrite Ropp_0; rewrite Rplus_0_r.
assert (H : (x >= 0)%R).
{ apply Rle_ge; apply Rle_trans with (bpow (ex - 1)); [|exact (proj1 Hex)].
now apply bpow_ge_0. }
-assert (H' : (x * bpow (- fexp (ln_beta beta x)) >= 0)%R).
+assert (H' : (x * bpow (- fexp (mag beta x)) >= 0)%R).
{ apply Rle_ge; apply Rmult_le_pos.
- now apply Rge_le.
- now apply bpow_ge_0. }
rewrite Rabs_right; [|exact H'].
-apply (Rmult_lt_reg_r (bpow (fexp (ln_beta beta x)))); [now apply bpow_gt_0|].
+apply (Rmult_lt_reg_r (bpow (fexp (mag beta x)))); [now apply bpow_gt_0|].
rewrite Rmult_assoc, <- bpow_plus.
replace (- _ + _)%Z with 0%Z by ring; simpl; rewrite Rmult_1_r.
apply (Rlt_le_trans _ _ _ (proj2 Hex)).
-apply Rle_trans with (bpow (fexp (ln_beta beta x) - 1)).
+apply Rle_trans with (bpow (fexp (mag beta x) - 1)).
- apply bpow_le.
- rewrite (ln_beta_unique beta x ex); [omega|].
+ rewrite (mag_unique beta x ex); [omega|].
now rewrite Rabs_right.
- unfold Zminus; rewrite bpow_plus.
rewrite Rmult_comm.
apply Rmult_le_compat_r; [now apply bpow_ge_0|].
- unfold Fcore_Raux.bpow, Z.pow_pos; simpl.
+ unfold Raux.bpow, Z.pow_pos; simpl.
rewrite Zmult_1_r.
apply Rinv_le; [exact Rlt_0_2|].
- change 2%R with (Z2R 2); apply Z2R_le.
+ apply IZR_le.
destruct beta as (beta_val,beta_prop).
now apply Zle_bool_imp_le.
Qed.
@@ -2024,7 +1966,7 @@ set (f := round (Znearest (Zle_bool 0)) x).
intros Rxf.
destruct (Req_dec (x - round Zfloor x) (round Zceil x - x)) as [Hm|Hm].
(* *)
-apply Rnd_NA_N_pt.
+apply Rnd_NA_pt_N.
exact generic_format_0.
exact Rxf.
destruct (Rle_or_lt 0 x) as [Hx|Hx].
@@ -2038,7 +1980,7 @@ apply (round_UP_pt x).
apply Zfloor_lub.
apply Rmult_le_pos with (1 := Hx).
apply bpow_ge_0.
-apply Rnd_N_pt_pos with (2 := Hx) (3 := Rxf).
+apply Rnd_N_pt_ge_0 with (2 := Hx) (3 := Rxf).
exact generic_format_0.
(* . *)
rewrite Rabs_left with (1 := Hx).
@@ -2048,21 +1990,21 @@ unfold f.
rewrite round_N_middle with (1 := Hm).
rewrite Zle_bool_false.
apply (round_DN_pt x).
-apply lt_Z2R.
+apply lt_IZR.
apply Rle_lt_trans with (scaled_mantissa x).
apply Zfloor_lb.
simpl.
-rewrite <- (Rmult_0_l (bpow (- canonic_exp x))).
+rewrite <- (Rmult_0_l (bpow (- cexp x))).
apply Rmult_lt_compat_r with (2 := Hx).
apply bpow_gt_0.
-apply Rnd_N_pt_neg with (3 := Rxf).
+apply Rnd_N_pt_le_0 with (3 := Rxf).
exact generic_format_0.
now apply Rlt_le.
(* *)
split.
apply Rxf.
intros g Rxg.
-rewrite Rnd_N_pt_unicity with (3 := Hm) (4 := Rxf) (5 := Rxg).
+rewrite Rnd_N_pt_unique with (3 := Hm) (4 := Rxf) (5 := Rxg).
apply Rle_refl.
apply round_DN_pt.
apply round_UP_pt.
@@ -2077,25 +2019,25 @@ Theorem Znearest_opp :
Znearest choice (- x) = (- Znearest (fun t => negb (choice (- (t + 1))%Z)) x)%Z.
Proof with auto with typeclass_instances.
intros choice x.
-destruct (Req_dec (Z2R (Zfloor x)) x) as [Hx|Hx].
+destruct (Req_dec (IZR (Zfloor x)) x) as [Hx|Hx].
rewrite <- Hx.
-rewrite <- Z2R_opp.
-rewrite 2!Zrnd_Z2R...
+rewrite <- opp_IZR.
+rewrite 2!Zrnd_IZR...
unfold Znearest.
-replace (- x - Z2R (Zfloor (-x)))%R with (Z2R (Zceil x) - x)%R.
-rewrite Rcompare_ceil_floor_mid with (1 := Hx).
-rewrite Rcompare_floor_ceil_mid with (1 := Hx).
+replace (- x - IZR (Zfloor (-x)))%R with (IZR (Zceil x) - x)%R.
+rewrite Rcompare_ceil_floor_middle with (1 := Hx).
+rewrite Rcompare_floor_ceil_middle with (1 := Hx).
rewrite Rcompare_sym.
rewrite <- Zceil_floor_neq with (1 := Hx).
unfold Zceil.
rewrite Ropp_involutive.
case Rcompare ; simpl ; trivial.
-rewrite Zopp_involutive.
+rewrite Z.opp_involutive.
case (choice (Zfloor (- x))) ; simpl ; trivial.
-now rewrite Zopp_involutive.
-now rewrite Zopp_involutive.
+now rewrite Z.opp_involutive.
+now rewrite Z.opp_involutive.
unfold Zceil.
-rewrite Z2R_opp.
+rewrite opp_IZR.
apply Rplus_comm.
Qed.
@@ -2106,15 +2048,30 @@ Theorem round_N_opp :
Proof.
intros choice x.
unfold round, F2R. simpl.
-rewrite canonic_exp_opp.
+rewrite cexp_opp.
rewrite scaled_mantissa_opp.
rewrite Znearest_opp.
-rewrite Z2R_opp.
+rewrite opp_IZR.
now rewrite Ropp_mult_distr_l_reverse.
Qed.
End rndN_opp.
+Lemma round_N_small :
+ forall choice,
+ forall x,
+ forall ex,
+ (Raux.bpow beta (ex - 1) <= Rabs x < Raux.bpow beta ex)%R ->
+ (ex < fexp ex)%Z ->
+ (round (Znearest choice) x = 0)%R.
+Proof.
+intros choice x ex Hx Hex.
+destruct (Rle_or_lt 0 x) as [Px|Nx].
+{ now revert Hex; apply round_N_small_pos; revert Hx; rewrite Rabs_pos_eq. }
+rewrite <-(Ropp_involutive x), round_N_opp, <-Ropp_0; f_equal.
+now revert Hex; apply round_N_small_pos; revert Hx; rewrite Rabs_left.
+Qed.
+
End Format.
(** Inclusion of a format into an extended format *)
@@ -2125,9 +2082,9 @@ Variables fexp1 fexp2 : Z -> Z.
Context { valid_exp1 : Valid_exp fexp1 }.
Context { valid_exp2 : Valid_exp fexp2 }.
-Theorem generic_inclusion_ln_beta :
+Theorem generic_inclusion_mag :
forall x,
- ( x <> R0 -> (fexp2 (ln_beta beta x) <= fexp1 (ln_beta beta x))%Z ) ->
+ ( x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z ) ->
generic_format fexp1 x ->
generic_format fexp2 x.
Proof.
@@ -2139,7 +2096,7 @@ rewrite <- Fx.
apply He.
contradict Zx.
rewrite Zx, scaled_mantissa_0.
-apply (Ztrunc_Z2R 0).
+apply Ztrunc_IZR.
Qed.
Theorem generic_inclusion_lt_ge :
@@ -2151,12 +2108,12 @@ Theorem generic_inclusion_lt_ge :
generic_format fexp2 x.
Proof.
intros e1 e2 He x (Hx1,Hx2).
-apply generic_inclusion_ln_beta.
+apply generic_inclusion_mag.
intros Zx.
apply He.
split.
-now apply ln_beta_gt_bpow.
-now apply ln_beta_le_bpow.
+now apply mag_gt_bpow.
+now apply mag_le_bpow.
Qed.
Theorem generic_inclusion :
@@ -2168,13 +2125,13 @@ Theorem generic_inclusion :
generic_format fexp2 x.
Proof with auto with typeclass_instances.
intros e He x (Hx1,[Hx2|Hx2]).
-apply generic_inclusion_ln_beta.
-now rewrite ln_beta_unique with (1 := conj Hx1 Hx2).
+apply generic_inclusion_mag.
+now rewrite mag_unique with (1 := conj Hx1 Hx2).
intros Fx.
apply generic_format_abs_inv.
rewrite Hx2.
apply generic_format_bpow'...
-apply Zle_trans with (1 := He).
+apply Z.le_trans with (1 := He).
apply generic_format_bpow_inv...
rewrite <- Hx2.
now apply generic_format_abs.
@@ -2191,18 +2148,18 @@ Theorem generic_inclusion_le_ge :
Proof.
intros e1 e2 He' He x (Hx1,[Hx2|Hx2]).
(* *)
-apply generic_inclusion_ln_beta.
+apply generic_inclusion_mag.
intros Zx.
apply He.
split.
-now apply ln_beta_gt_bpow.
-now apply ln_beta_le_bpow.
+now apply mag_gt_bpow.
+now apply mag_le_bpow.
(* *)
apply generic_inclusion with (e := e2).
apply He.
split.
apply He'.
-apply Zle_refl.
+apply Z.le_refl.
rewrite Hx2.
split.
apply bpow_le.
@@ -2219,13 +2176,13 @@ Theorem generic_inclusion_le :
generic_format fexp2 x.
Proof.
intros e2 He x [Hx|Hx].
-apply generic_inclusion_ln_beta.
+apply generic_inclusion_mag.
intros Zx.
apply He.
-now apply ln_beta_le_bpow.
+now apply mag_le_bpow.
apply generic_inclusion with (e := e2).
apply He.
-apply Zle_refl.
+apply Z.le_refl.
rewrite Hx.
split.
apply bpow_le.
@@ -2242,10 +2199,10 @@ Theorem generic_inclusion_ge :
generic_format fexp2 x.
Proof.
intros e1 He x Hx.
-apply generic_inclusion_ln_beta.
+apply generic_inclusion_mag.
intros Zx.
apply He.
-now apply ln_beta_gt_bpow.
+now apply mag_gt_bpow.
Qed.
Variable rnd : R -> Z.
@@ -2263,9 +2220,9 @@ revert rnd valid_rnd x Gx.
refine (round_abs_abs _ (fun x y => generic_format fexp1 x -> generic_format fexp1 y) _).
intros rnd valid_rnd x [Hx|Hx] Gx.
(* x > 0 *)
-generalize (ln_beta_generic_gt _ x (Rgt_not_eq _ _ Hx) Gx).
-unfold canonic_exp.
-destruct (ln_beta beta x) as (ex,Ex) ; simpl.
+generalize (mag_generic_gt _ x (Rgt_not_eq _ _ Hx) Gx).
+unfold cexp.
+destruct (mag beta x) as (ex,Ex) ; simpl.
specialize (Ex (Rgt_not_eq _ _ Hx)).
intros He'.
rewrite Rabs_pos_eq in Ex by now apply Rlt_le.
@@ -2279,25 +2236,25 @@ apply generic_format_bpow'...
apply Zlt_le_weak.
apply valid_exp_large with ex...
(* - x large for fexp2 *)
-destruct (Zle_or_lt (canonic_exp fexp2 x) (canonic_exp fexp1 x)) as [He''|He''].
+destruct (Zle_or_lt (cexp fexp2 x) (cexp fexp1 x)) as [He''|He''].
(* - - round fexp2 x is representable for fexp1 *)
rewrite round_generic...
rewrite Gx.
apply generic_format_F2R.
fold (round fexp1 Ztrunc x).
intros Zx.
-unfold canonic_exp at 1.
-rewrite ln_beta_round_ZR...
+unfold cexp at 1.
+rewrite mag_round_ZR...
contradict Zx.
-apply F2R_eq_0_reg with (1 := Zx).
+apply eq_0_F2R with (1 := Zx).
(* - - round fexp2 x has too many digits for fexp1 *)
destruct (round_bounded_large_pos fexp2 rnd x ex He Ex) as (Hr1,[Hr2|Hr2]).
apply generic_format_F2R.
intros Zx.
fold (round fexp2 rnd x).
-unfold canonic_exp at 1.
-rewrite ln_beta_unique_pos with (1 := conj Hr1 Hr2).
-rewrite <- ln_beta_unique_pos with (1 := Ex).
+unfold cexp at 1.
+rewrite mag_unique_pos with (1 := conj Hr1 Hr2).
+rewrite <- mag_unique_pos with (1 := Ex).
now apply Zlt_le_weak.
rewrite Hr2.
apply generic_format_bpow'...
@@ -2327,7 +2284,7 @@ apply Ropp_eq_compat.
apply round_ext.
clear x; intro x.
unfold Znearest.
-case_eq (Rcompare (x - Z2R (Zfloor x)) (/ 2)); intro C;
+case_eq (Rcompare (x - IZR (Zfloor x)) (/ 2)); intro C;
[|reflexivity|reflexivity].
apply Rcompare_Eq_inv in C.
assert (H : negb (0 <=? - (Zfloor x + 1))%Z = (0 <=? Zfloor x)%Z);
diff --git a/flocq/Core/Fcore_Raux.v b/flocq/Core/Raux.v
index 77235e63..8273a55b 100644
--- a/flocq/Core/Fcore_Raux.v
+++ b/flocq/Core/Raux.v
@@ -2,9 +2,9 @@
This file is part of the Flocq formalization of floating-point
arithmetic in Coq: http://flocq.gforge.inria.fr/
-Copyright (C) 2010-2013 Sylvie Boldo
+Copyright (C) 2009-2018 Sylvie Boldo
#<br />#
-Copyright (C) 2010-2013 Guillaume Melquiond
+Copyright (C) 2009-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
@@ -18,9 +18,9 @@ COPYING file for more details.
*)
(** * Missing definitions/lemmas *)
-Require Export Reals.
-Require Export ZArith.
-Require Export Fcore_Zaux.
+Require Import Psatz.
+Require Export Reals ZArith.
+Require Export Zaux.
Section Rmissing.
@@ -58,12 +58,13 @@ Theorem Rabs_minus_le:
(Rabs (x-y) <= x)%R.
Proof.
intros x y Hx Hy.
-unfold Rabs; case (Rcase_abs (x - y)); intros H.
-apply Rplus_le_reg_l with x; ring_simplify; assumption.
-apply Rplus_le_reg_l with (-x)%R; ring_simplify.
-auto with real.
+apply Rabs_le.
+lra.
Qed.
+Theorem Rabs_eq_R0 x : (Rabs x = 0 -> x = 0)%R.
+Proof. split_Rabs; lra. Qed.
+
Theorem Rplus_eq_reg_r :
forall r r1 r2 : R,
(r1 + r = r2 + r)%R -> (r1 = r2)%R.
@@ -73,53 +74,6 @@ apply Rplus_eq_reg_l with r.
now rewrite 2!(Rplus_comm r).
Qed.
-Theorem Rplus_lt_reg_l :
- forall r r1 r2 : R,
- (r + r1 < r + r2)%R -> (r1 < r2)%R.
-Proof.
-intros.
-solve [ apply Rplus_lt_reg_l with (1 := H) |
- apply Rplus_lt_reg_r with (1 := H) ].
-Qed.
-
-Theorem Rplus_lt_reg_r :
- forall r r1 r2 : R,
- (r1 + r < r2 + r)%R -> (r1 < r2)%R.
-Proof.
-intros.
-apply Rplus_lt_reg_l with r.
-now rewrite 2!(Rplus_comm r).
-Qed.
-
-Theorem Rplus_le_reg_r :
- forall r r1 r2 : R,
- (r1 + r <= r2 + r)%R -> (r1 <= r2)%R.
-Proof.
-intros.
-apply Rplus_le_reg_l with r.
-now rewrite 2!(Rplus_comm r).
-Qed.
-
-Theorem Rmult_lt_reg_r :
- forall r r1 r2 : R, (0 < r)%R ->
- (r1 * r < r2 * r)%R -> (r1 < r2)%R.
-Proof.
-intros.
-apply Rmult_lt_reg_l with r.
-exact H.
-now rewrite 2!(Rmult_comm r).
-Qed.
-
-Theorem Rmult_le_reg_r :
- forall r r1 r2 : R, (0 < r)%R ->
- (r1 * r <= r2 * r)%R -> (r1 <= r2)%R.
-Proof.
-intros.
-apply Rmult_le_reg_l with r.
-exact H.
-now rewrite 2!(Rmult_comm r).
-Qed.
-
Theorem Rmult_lt_compat :
forall r1 r2 r3 r4,
(0 <= r1)%R -> (0 <= r3)%R -> (r1 < r2)%R -> (r3 < r4)%R ->
@@ -135,16 +89,6 @@ apply Rle_lt_trans with (r1 * r4)%R.
+ exact H12.
Qed.
-Theorem Rmult_eq_reg_r :
- forall r r1 r2 : R, (r1 * r)%R = (r2 * r)%R ->
- r <> 0%R -> r1 = r2.
-Proof.
-intros r r1 r2 H1 H2.
-apply Rmult_eq_reg_l with r.
-now rewrite 2!(Rmult_comm r).
-exact H2.
-Qed.
-
Theorem Rmult_minus_distr_r :
forall r r1 r2 : R,
((r1 - r2) * r = r1 * r - r2 * r)%R.
@@ -154,13 +98,18 @@ rewrite <- 3!(Rmult_comm r).
apply Rmult_minus_distr_l.
Qed.
-Lemma Rmult_neq_reg_r: forall r1 r2 r3:R, (r2 * r1 <> r3 * r1)%R -> r2 <> r3.
+Lemma Rmult_neq_reg_r :
+ forall r1 r2 r3 : R, (r2 * r1 <> r3 * r1)%R -> r2 <> r3.
+Proof.
intros r1 r2 r3 H1 H2.
apply H1; rewrite H2; ring.
Qed.
-Lemma Rmult_neq_compat_r: forall r1 r2 r3:R, (r1 <> 0)%R -> (r2 <> r3)%R
- -> (r2 *r1 <> r3*r1)%R.
+Lemma Rmult_neq_compat_r :
+ forall r1 r2 r3 : R,
+ (r1 <> 0)%R -> (r2 <> r3)%R ->
+ (r2 * r1 <> r3 * r1)%R.
+Proof.
intros r1 r2 r3 H H1 H2.
now apply H1, Rmult_eq_reg_r with r1.
Qed.
@@ -227,7 +176,6 @@ rewrite Rmax_right; trivial.
now apply Ropp_le_contravar.
Qed.
-
Theorem exp_le :
forall x y : R,
(x <= y)%R -> (exp x <= exp y)%R.
@@ -288,6 +236,14 @@ destruct (Req_dec x 0) as [Zx|Nzx].
now apply Nzx, Rle_antisym; [|apply Rge_le].
Qed.
+Lemma Rsqr_le_abs_0_alt :
+ forall x y,
+ (x² <= y² -> x <= Rabs y)%R.
+Proof.
+intros x y H.
+apply (Rle_trans _ (Rabs x)); [apply Rle_abs|apply (Rsqr_le_abs_0 _ _ H)].
+Qed.
+
Theorem Rabs_le :
forall x y,
(-y <= x <= y)%R -> (Rabs x <= y)%R.
@@ -387,187 +343,35 @@ Qed.
End Rmissing.
-Section Z2R.
+Section IZR.
-(** Z2R function (Z -> R) *)
-Fixpoint P2R (p : positive) :=
- match p with
- | xH => 1%R
- | xO xH => 2%R
- | xO t => (2 * P2R t)%R
- | xI xH => 3%R
- | xI t => (1 + 2 * P2R t)%R
- end.
-
-Definition Z2R n :=
- match n with
- | Zpos p => P2R p
- | Zneg p => Ropp (P2R p)
- | Z0 => 0%R
- end.
-
-Theorem P2R_INR :
- forall n, P2R n = INR (nat_of_P n).
-Proof.
-induction n ; simpl ; try (
- rewrite IHn ;
- rewrite <- (mult_INR 2) ;
- rewrite <- (nat_of_P_mult_morphism 2) ;
- change (2 * n)%positive with (xO n)).
-(* xI *)
-rewrite (Rplus_comm 1).
-change (nat_of_P (xO n)) with (Pmult_nat n 2).
-case n ; intros ; simpl ; try apply refl_equal.
-case (Pmult_nat p 4) ; intros ; try apply refl_equal.
-rewrite Rplus_0_l.
-apply refl_equal.
-apply Rplus_comm.
-(* xO *)
-case n ; intros ; apply refl_equal.
-(* xH *)
-apply refl_equal.
-Qed.
-
-Theorem Z2R_IZR :
- forall n, Z2R n = IZR n.
-Proof.
-intro.
-case n ; intros ; unfold Z2R.
-apply refl_equal.
-rewrite <- positive_nat_Z, <- INR_IZR_INZ.
-apply P2R_INR.
-change (IZR (Zneg p)) with (Ropp (IZR (Zpos p))).
-apply Ropp_eq_compat.
-rewrite <- positive_nat_Z, <- INR_IZR_INZ.
-apply P2R_INR.
-Qed.
-
-Theorem Z2R_opp :
- forall n, Z2R (-n) = (- Z2R n)%R.
-Proof.
-intros.
-repeat rewrite Z2R_IZR.
-apply Ropp_Ropp_IZR.
-Qed.
-
-Theorem Z2R_plus :
- forall m n, (Z2R (m + n) = Z2R m + Z2R n)%R.
-Proof.
-intros.
-repeat rewrite Z2R_IZR.
-apply plus_IZR.
-Qed.
-
-Theorem minus_IZR :
- forall n m : Z,
- IZR (n - m) = (IZR n - IZR m)%R.
-Proof.
-intros.
-unfold Zminus.
-rewrite plus_IZR.
-rewrite Ropp_Ropp_IZR.
-apply refl_equal.
-Qed.
-
-Theorem Z2R_minus :
- forall m n, (Z2R (m - n) = Z2R m - Z2R n)%R.
-Proof.
-intros.
-repeat rewrite Z2R_IZR.
-apply minus_IZR.
-Qed.
-
-Theorem Z2R_mult :
- forall m n, (Z2R (m * n) = Z2R m * Z2R n)%R.
-Proof.
-intros.
-repeat rewrite Z2R_IZR.
-apply mult_IZR.
-Qed.
-
-Theorem le_Z2R :
- forall m n, (Z2R m <= Z2R n)%R -> (m <= n)%Z.
-Proof.
-intros m n.
-repeat rewrite Z2R_IZR.
-apply le_IZR.
-Qed.
-
-Theorem Z2R_le :
- forall m n, (m <= n)%Z -> (Z2R m <= Z2R n)%R.
-Proof.
-intros m n.
-repeat rewrite Z2R_IZR.
-apply IZR_le.
-Qed.
-
-Theorem lt_Z2R :
- forall m n, (Z2R m < Z2R n)%R -> (m < n)%Z.
-Proof.
-intros m n.
-repeat rewrite Z2R_IZR.
-apply lt_IZR.
-Qed.
-
-Theorem Z2R_lt :
- forall m n, (m < n)%Z -> (Z2R m < Z2R n)%R.
-Proof.
-intros m n.
-repeat rewrite Z2R_IZR.
-apply IZR_lt.
-Qed.
-
-Theorem Z2R_le_lt :
- forall m n p, (m <= n < p)%Z -> (Z2R m <= Z2R n < Z2R p)%R.
+Theorem IZR_le_lt :
+ forall m n p, (m <= n < p)%Z -> (IZR m <= IZR n < IZR p)%R.
Proof.
intros m n p (H1, H2).
split.
-now apply Z2R_le.
-now apply Z2R_lt.
+now apply IZR_le.
+now apply IZR_lt.
Qed.
-Theorem le_lt_Z2R :
- forall m n p, (Z2R m <= Z2R n < Z2R p)%R -> (m <= n < p)%Z.
+Theorem le_lt_IZR :
+ forall m n p, (IZR m <= IZR n < IZR p)%R -> (m <= n < p)%Z.
Proof.
intros m n p (H1, H2).
split.
-now apply le_Z2R.
-now apply lt_Z2R.
-Qed.
-
-Theorem eq_Z2R :
- forall m n, (Z2R m = Z2R n)%R -> (m = n)%Z.
-Proof.
-intros m n H.
-apply eq_IZR.
-now rewrite <- 2!Z2R_IZR.
+now apply le_IZR.
+now apply lt_IZR.
Qed.
-Theorem neq_Z2R :
- forall m n, (Z2R m <> Z2R n)%R -> (m <> n)%Z.
+Theorem neq_IZR :
+ forall m n, (IZR m <> IZR n)%R -> (m <> n)%Z.
Proof.
intros m n H H'.
apply H.
now apply f_equal.
Qed.
-Theorem Z2R_neq :
- forall m n, (m <> n)%Z -> (Z2R m <> Z2R n)%R.
-Proof.
-intros m n.
-repeat rewrite Z2R_IZR.
-apply IZR_neq.
-Qed.
-
-Theorem Z2R_abs :
- forall z, Z2R (Zabs z) = Rabs (Z2R z).
-Proof.
-intros.
-repeat rewrite Z2R_IZR.
-now rewrite Rabs_Zabs.
-Qed.
-
-End Z2R.
+End IZR.
(** Decidable comparison on reals *)
Section Rcompare.
@@ -691,17 +495,17 @@ contradict H.
now apply Rcompare_Gt.
Qed.
-Theorem Rcompare_Z2R :
- forall x y, Rcompare (Z2R x) (Z2R y) = Zcompare x y.
+Theorem Rcompare_IZR :
+ forall x y, Rcompare (IZR x) (IZR y) = Z.compare x y.
Proof.
intros x y.
case Rcompare_spec ; intros H ; apply sym_eq.
apply Zcompare_Lt.
-now apply lt_Z2R.
+now apply lt_IZR.
apply Zcompare_Eq.
-now apply eq_Z2R.
+now apply eq_IZR.
apply Zcompare_Gt.
-now apply lt_Z2R.
+now apply lt_IZR.
Qed.
Theorem Rcompare_sym :
@@ -715,6 +519,16 @@ now apply Rcompare_Eq.
now apply Rcompare_Lt.
Qed.
+Lemma Rcompare_opp :
+ forall x y,
+ Rcompare (- x) (- y) = Rcompare y x.
+Proof.
+intros x y.
+destruct (Rcompare_spec y x);
+ destruct (Rcompare_spec (- x) (- y));
+ try reflexivity; exfalso; lra.
+Qed.
+
Theorem Rcompare_plus_r :
forall z x y,
Rcompare (x + z) (y + z) = Rcompare x y.
@@ -773,7 +587,7 @@ rewrite <- (Rcompare_mult_r (/2) (x - d)).
field_simplify (x + (- x / 2 - d / 2))%R.
now field_simplify ((d + u) / 2 + (- x / 2 - d / 2))%R.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
Qed.
Theorem Rcompare_half_l :
@@ -784,8 +598,8 @@ rewrite <- (Rcompare_mult_r 2%R).
unfold Rdiv.
rewrite Rmult_assoc, Rinv_l, Rmult_1_r.
now rewrite Rmult_comm.
-now apply (Z2R_neq 2 0).
-now apply (Z2R_lt 0 2).
+now apply IZR_neq.
+now apply IZR_lt.
Qed.
Theorem Rcompare_half_r :
@@ -796,23 +610,23 @@ rewrite <- (Rcompare_mult_r 2%R).
unfold Rdiv.
rewrite Rmult_assoc, Rinv_l, Rmult_1_r.
now rewrite Rmult_comm.
-now apply (Z2R_neq 2 0).
-now apply (Z2R_lt 0 2).
+now apply IZR_neq.
+now apply IZR_lt.
Qed.
Theorem Rcompare_sqr :
forall x y,
- (0 <= x)%R -> (0 <= y)%R ->
- Rcompare (x * x) (y * y) = Rcompare x y.
+ Rcompare (x * x) (y * y) = Rcompare (Rabs x) (Rabs y).
Proof.
-intros x y Hx Hy.
-destruct (Rcompare_spec x y) as [H|H|H].
+intros x y.
+destruct (Rcompare_spec (Rabs x) (Rabs y)) as [H|H|H].
apply Rcompare_Lt.
-now apply Rsqr_incrst_1.
-rewrite H.
+now apply Rsqr_lt_abs_1.
+change (Rcompare (Rsqr x) (Rsqr y) = Eq).
+rewrite Rsqr_abs, H, (Rsqr_abs y).
now apply Rcompare_Eq.
apply Rcompare_Gt.
-now apply Rsqr_incrst_1.
+now apply Rsqr_lt_abs_1.
Qed.
Theorem Rmin_compare :
@@ -941,6 +755,14 @@ rewrite <- negb_Rlt_bool.
now rewrite Rle_bool_true.
Qed.
+Lemma Rlt_bool_opp :
+ forall x y,
+ Rlt_bool (- x) (- y) = Rlt_bool y x.
+Proof.
+intros x y.
+now unfold Rlt_bool; rewrite Rcompare_opp.
+Qed.
+
End Rlt_bool.
Section Req_bool.
@@ -997,13 +819,12 @@ Definition Zfloor (x : R) := (up x - 1)%Z.
Theorem Zfloor_lb :
forall x : R,
- (Z2R (Zfloor x) <= x)%R.
+ (IZR (Zfloor x) <= x)%R.
Proof.
intros x.
unfold Zfloor.
-rewrite Z2R_minus.
+rewrite minus_IZR.
simpl.
-rewrite Z2R_IZR.
apply Rplus_le_reg_r with (1 - x)%R.
ring_simplify.
exact (proj2 (archimed x)).
@@ -1011,55 +832,54 @@ Qed.
Theorem Zfloor_ub :
forall x : R,
- (x < Z2R (Zfloor x) + 1)%R.
+ (x < IZR (Zfloor x) + 1)%R.
Proof.
intros x.
unfold Zfloor.
-rewrite Z2R_minus.
+rewrite minus_IZR.
unfold Rminus.
rewrite Rplus_assoc.
rewrite Rplus_opp_l, Rplus_0_r.
-rewrite Z2R_IZR.
exact (proj1 (archimed x)).
Qed.
Theorem Zfloor_lub :
forall n x,
- (Z2R n <= x)%R ->
+ (IZR n <= x)%R ->
(n <= Zfloor x)%Z.
Proof.
intros n x Hnx.
apply Zlt_succ_le.
-apply lt_Z2R.
+apply lt_IZR.
apply Rle_lt_trans with (1 := Hnx).
-unfold Zsucc.
-rewrite Z2R_plus.
+unfold Z.succ.
+rewrite plus_IZR.
apply Zfloor_ub.
Qed.
Theorem Zfloor_imp :
forall n x,
- (Z2R n <= x < Z2R (n + 1))%R ->
+ (IZR n <= x < IZR (n + 1))%R ->
Zfloor x = n.
Proof.
intros n x Hnx.
apply Zle_antisym.
apply Zlt_succ_le.
-apply lt_Z2R.
+apply lt_IZR.
apply Rle_lt_trans with (2 := proj2 Hnx).
apply Zfloor_lb.
now apply Zfloor_lub.
Qed.
-Theorem Zfloor_Z2R :
+Theorem Zfloor_IZR :
forall n,
- Zfloor (Z2R n) = n.
+ Zfloor (IZR n) = n.
Proof.
intros n.
apply Zfloor_imp.
split.
apply Rle_refl.
-apply Z2R_lt.
+apply IZR_lt.
apply Zlt_succ.
Qed.
@@ -1077,11 +897,11 @@ Definition Zceil (x : R) := (- Zfloor (- x))%Z.
Theorem Zceil_ub :
forall x : R,
- (x <= Z2R (Zceil x))%R.
+ (x <= IZR (Zceil x))%R.
Proof.
intros x.
unfold Zceil.
-rewrite Z2R_opp.
+rewrite opp_IZR.
apply Ropp_le_cancel.
rewrite Ropp_involutive.
apply Zfloor_lb.
@@ -1089,45 +909,45 @@ Qed.
Theorem Zceil_glb :
forall n x,
- (x <= Z2R n)%R ->
+ (x <= IZR n)%R ->
(Zceil x <= n)%Z.
Proof.
intros n x Hnx.
unfold Zceil.
apply Zopp_le_cancel.
-rewrite Zopp_involutive.
+rewrite Z.opp_involutive.
apply Zfloor_lub.
-rewrite Z2R_opp.
+rewrite opp_IZR.
now apply Ropp_le_contravar.
Qed.
Theorem Zceil_imp :
forall n x,
- (Z2R (n - 1) < x <= Z2R n)%R ->
+ (IZR (n - 1) < x <= IZR n)%R ->
Zceil x = n.
Proof.
intros n x Hnx.
unfold Zceil.
-rewrite <- (Zopp_involutive n).
+rewrite <- (Z.opp_involutive n).
apply f_equal.
apply Zfloor_imp.
split.
-rewrite Z2R_opp.
+rewrite opp_IZR.
now apply Ropp_le_contravar.
-rewrite <- (Zopp_involutive 1).
+rewrite <- (Z.opp_involutive 1).
rewrite <- Zopp_plus_distr.
-rewrite Z2R_opp.
+rewrite opp_IZR.
now apply Ropp_lt_contravar.
Qed.
-Theorem Zceil_Z2R :
+Theorem Zceil_IZR :
forall n,
- Zceil (Z2R n) = n.
+ Zceil (IZR n) = n.
Proof.
intros n.
unfold Zceil.
-rewrite <- Z2R_opp, Zfloor_Z2R.
-apply Zopp_involutive.
+rewrite <- opp_IZR, Zfloor_IZR.
+apply Z.opp_involutive.
Qed.
Theorem Zceil_le :
@@ -1142,7 +962,7 @@ Qed.
Theorem Zceil_floor_neq :
forall x : R,
- (Z2R (Zfloor x) <> x)%R ->
+ (IZR (Zfloor x) <> x)%R ->
(Zceil x = Zfloor x + 1)%Z.
Proof.
intros x Hx.
@@ -1156,21 +976,21 @@ apply Rle_antisym.
apply Zfloor_lb.
exact H.
apply Rlt_le.
-rewrite Z2R_plus.
+rewrite plus_IZR.
apply Zfloor_ub.
Qed.
Definition Ztrunc x := if Rlt_bool x 0 then Zceil x else Zfloor x.
-Theorem Ztrunc_Z2R :
+Theorem Ztrunc_IZR :
forall n,
- Ztrunc (Z2R n) = n.
+ Ztrunc (IZR n) = n.
Proof.
intros n.
unfold Ztrunc.
case Rlt_bool_spec ; intro H.
-apply Zceil_Z2R.
-apply Zfloor_Z2R.
+apply Zceil_IZR.
+apply Zfloor_IZR.
Qed.
Theorem Ztrunc_floor :
@@ -1196,9 +1016,8 @@ unfold Ztrunc.
case Rlt_bool_spec ; intro H.
apply refl_equal.
rewrite (Rle_antisym _ _ Hx H).
-change 0%R with (Z2R 0).
-rewrite Zceil_Z2R.
-apply Zfloor_Z2R.
+rewrite Zceil_IZR.
+apply Zfloor_IZR.
Qed.
Theorem Ztrunc_le :
@@ -1211,7 +1030,7 @@ case Rlt_bool_spec ; intro Hx.
unfold Ztrunc.
case Rlt_bool_spec ; intro Hy.
now apply Zceil_le.
-apply Zle_trans with 0%Z.
+apply Z.le_trans with 0%Z.
apply Zceil_glb.
now apply Rlt_le.
now apply Zfloor_lub.
@@ -1222,14 +1041,14 @@ Qed.
Theorem Ztrunc_opp :
forall x,
- Ztrunc (- x) = Zopp (Ztrunc x).
+ Ztrunc (- x) = Z.opp (Ztrunc x).
Proof.
intros x.
unfold Ztrunc at 2.
case Rlt_bool_spec ; intros Hx.
rewrite Ztrunc_floor.
apply sym_eq.
-apply Zopp_involutive.
+apply Z.opp_involutive.
rewrite <- Ropp_0.
apply Ropp_le_contravar.
now apply Rlt_le.
@@ -1242,7 +1061,7 @@ Qed.
Theorem Ztrunc_abs :
forall x,
- Ztrunc (Rabs x) = Zabs (Ztrunc x).
+ Ztrunc (Rabs x) = Z.abs (Ztrunc x).
Proof.
intros x.
rewrite Ztrunc_floor. 2: apply Rabs_pos.
@@ -1251,19 +1070,19 @@ case Rlt_bool_spec ; intro H.
rewrite Rabs_left with (1 := H).
rewrite Zabs_non_eq.
apply sym_eq.
-apply Zopp_involutive.
+apply Z.opp_involutive.
apply Zceil_glb.
now apply Rlt_le.
rewrite Rabs_pos_eq with (1 := H).
apply sym_eq.
-apply Zabs_eq.
+apply Z.abs_eq.
now apply Zfloor_lub.
Qed.
Theorem Ztrunc_lub :
forall n x,
- (Z2R n <= Rabs x)%R ->
- (n <= Zabs (Ztrunc x))%Z.
+ (IZR n <= Rabs x)%R ->
+ (n <= Z.abs (Ztrunc x))%Z.
Proof.
intros n x H.
rewrite <- Ztrunc_abs.
@@ -1273,15 +1092,15 @@ Qed.
Definition Zaway x := if Rlt_bool x 0 then Zfloor x else Zceil x.
-Theorem Zaway_Z2R :
+Theorem Zaway_IZR :
forall n,
- Zaway (Z2R n) = n.
+ Zaway (IZR n) = n.
Proof.
intros n.
unfold Zaway.
case Rlt_bool_spec ; intro H.
-apply Zfloor_Z2R.
-apply Zceil_Z2R.
+apply Zfloor_IZR.
+apply Zceil_IZR.
Qed.
Theorem Zaway_ceil :
@@ -1307,9 +1126,8 @@ unfold Zaway.
case Rlt_bool_spec ; intro H.
apply refl_equal.
rewrite (Rle_antisym _ _ Hx H).
-change 0%R with (Z2R 0).
-rewrite Zfloor_Z2R.
-apply Zceil_Z2R.
+rewrite Zfloor_IZR.
+apply Zceil_IZR.
Qed.
Theorem Zaway_le :
@@ -1322,7 +1140,7 @@ case Rlt_bool_spec ; intro Hx.
unfold Zaway.
case Rlt_bool_spec ; intro Hy.
now apply Zfloor_le.
-apply le_Z2R.
+apply le_IZR.
apply Rle_trans with 0%R.
apply Rlt_le.
apply Rle_lt_trans with (2 := Hx).
@@ -1336,7 +1154,7 @@ Qed.
Theorem Zaway_opp :
forall x,
- Zaway (- x) = Zopp (Zaway x).
+ Zaway (- x) = Z.opp (Zaway x).
Proof.
intros x.
unfold Zaway at 2.
@@ -1348,14 +1166,14 @@ apply Rlt_le.
now apply Ropp_0_gt_lt_contravar.
rewrite Zaway_floor.
apply sym_eq.
-apply Zopp_involutive.
+apply Z.opp_involutive.
rewrite <- Ropp_0.
now apply Ropp_le_contravar.
Qed.
Theorem Zaway_abs :
forall x,
- Zaway (Rabs x) = Zabs (Zaway x).
+ Zaway (Rabs x) = Z.abs (Zaway x).
Proof.
intros x.
rewrite Zaway_ceil. 2: apply Rabs_pos.
@@ -1365,66 +1183,126 @@ rewrite Rabs_left with (1 := H).
rewrite Zabs_non_eq.
apply (f_equal (fun v => - Zfloor v)%Z).
apply Ropp_involutive.
-apply le_Z2R.
+apply le_IZR.
apply Rlt_le.
apply Rle_lt_trans with (2 := H).
apply Zfloor_lb.
rewrite Rabs_pos_eq with (1 := H).
apply sym_eq.
-apply Zabs_eq.
-apply le_Z2R.
+apply Z.abs_eq.
+apply le_IZR.
apply Rle_trans with (1 := H).
apply Zceil_ub.
Qed.
End Floor_Ceil.
+Theorem Rcompare_floor_ceil_middle :
+ forall x,
+ IZR (Zfloor x) <> x ->
+ Rcompare (x - IZR (Zfloor x)) (/ 2) = Rcompare (x - IZR (Zfloor x)) (IZR (Zceil x) - x).
+Proof.
+intros x Hx.
+rewrite Zceil_floor_neq with (1 := Hx).
+rewrite plus_IZR.
+destruct (Rcompare_spec (x - IZR (Zfloor x)) (/ 2)) as [H1|H1|H1] ; apply sym_eq.
+(* . *)
+apply Rcompare_Lt.
+apply Rplus_lt_reg_l with (x - IZR (Zfloor x))%R.
+replace (x - IZR (Zfloor x) + (x - IZR (Zfloor x)))%R with ((x - IZR (Zfloor x)) * 2)%R by ring.
+replace (x - IZR (Zfloor x) + (IZR (Zfloor x) + 1 - x))%R with (/2 * 2)%R by field.
+apply Rmult_lt_compat_r with (2 := H1).
+now apply IZR_lt.
+(* . *)
+apply Rcompare_Eq.
+replace (IZR (Zfloor x) + 1 - x)%R with (1 - (x - IZR (Zfloor x)))%R by ring.
+rewrite H1.
+field.
+(* . *)
+apply Rcompare_Gt.
+apply Rplus_lt_reg_l with (x - IZR (Zfloor x))%R.
+replace (x - IZR (Zfloor x) + (x - IZR (Zfloor x)))%R with ((x - IZR (Zfloor x)) * 2)%R by ring.
+replace (x - IZR (Zfloor x) + (IZR (Zfloor x) + 1 - x))%R with (/2 * 2)%R by field.
+apply Rmult_lt_compat_r with (2 := H1).
+now apply IZR_lt.
+Qed.
+
+Theorem Rcompare_ceil_floor_middle :
+ forall x,
+ IZR (Zfloor x) <> x ->
+ Rcompare (IZR (Zceil x) - x) (/ 2) = Rcompare (IZR (Zceil x) - x) (x - IZR (Zfloor x)).
+Proof.
+intros x Hx.
+rewrite Zceil_floor_neq with (1 := Hx).
+rewrite plus_IZR.
+destruct (Rcompare_spec (IZR (Zfloor x) + 1 - x) (/ 2)) as [H1|H1|H1] ; apply sym_eq.
+(* . *)
+apply Rcompare_Lt.
+apply Rplus_lt_reg_l with (IZR (Zfloor x) + 1 - x)%R.
+replace (IZR (Zfloor x) + 1 - x + (IZR (Zfloor x) + 1 - x))%R with ((IZR (Zfloor x) + 1 - x) * 2)%R by ring.
+replace (IZR (Zfloor x) + 1 - x + (x - IZR (Zfloor x)))%R with (/2 * 2)%R by field.
+apply Rmult_lt_compat_r with (2 := H1).
+now apply IZR_lt.
+(* . *)
+apply Rcompare_Eq.
+replace (x - IZR (Zfloor x))%R with (1 - (IZR (Zfloor x) + 1 - x))%R by ring.
+rewrite H1.
+field.
+(* . *)
+apply Rcompare_Gt.
+apply Rplus_lt_reg_l with (IZR (Zfloor x) + 1 - x)%R.
+replace (IZR (Zfloor x) + 1 - x + (IZR (Zfloor x) + 1 - x))%R with ((IZR (Zfloor x) + 1 - x) * 2)%R by ring.
+replace (IZR (Zfloor x) + 1 - x + (x - IZR (Zfloor x)))%R with (/2 * 2)%R by field.
+apply Rmult_lt_compat_r with (2 := H1).
+now apply IZR_lt.
+Qed.
+
Section Zdiv_Rdiv.
Theorem Zfloor_div :
forall x y,
y <> Z0 ->
- Zfloor (Z2R x / Z2R y) = (x / y)%Z.
+ Zfloor (IZR x / IZR y) = (x / y)%Z.
Proof.
intros x y Zy.
generalize (Z_div_mod_eq_full x y Zy).
intros Hx.
rewrite Hx at 1.
-assert (Zy': Z2R y <> R0).
+assert (Zy': IZR y <> 0%R).
contradict Zy.
-now apply eq_Z2R.
+now apply eq_IZR.
unfold Rdiv.
-rewrite Z2R_plus, Rmult_plus_distr_r, Z2R_mult.
-replace (Z2R y * Z2R (x / y) * / Z2R y)%R with (Z2R (x / y)) by now field.
+rewrite plus_IZR, Rmult_plus_distr_r, mult_IZR.
+replace (IZR y * IZR (x / y) * / IZR y)%R with (IZR (x / y)) by now field.
apply Zfloor_imp.
-rewrite Z2R_plus.
-assert (0 <= Z2R (x mod y) * / Z2R y < 1)%R.
+rewrite plus_IZR.
+assert (0 <= IZR (x mod y) * / IZR y < 1)%R.
(* *)
-assert (forall x' y', (0 < y')%Z -> 0 <= Z2R (x' mod y') * / Z2R y' < 1)%R.
+assert (forall x' y', (0 < y')%Z -> 0 <= IZR (x' mod y') * / IZR y' < 1)%R.
(* . *)
clear.
intros x y Hy.
split.
apply Rmult_le_pos.
-apply (Z2R_le 0).
+apply IZR_le.
refine (proj1 (Z_mod_lt _ _ _)).
-now apply Zlt_gt.
+now apply Z.lt_gt.
apply Rlt_le.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0).
-apply Rmult_lt_reg_r with (Z2R y).
-now apply (Z2R_lt 0).
+now apply IZR_lt.
+apply Rmult_lt_reg_r with (IZR y).
+now apply IZR_lt.
rewrite Rmult_1_l, Rmult_assoc, Rinv_l, Rmult_1_r.
-apply Z2R_lt.
+apply IZR_lt.
eapply Z_mod_lt.
-now apply Zlt_gt.
+now apply Z.lt_gt.
apply Rgt_not_eq.
-now apply (Z2R_lt 0).
+now apply IZR_lt.
(* . *)
destruct (Z_lt_le_dec y 0) as [Hy|Hy].
rewrite <- Rmult_opp_opp.
rewrite Ropp_inv_permute with (1 := Zy').
-rewrite <- 2!Z2R_opp.
+rewrite <- 2!opp_IZR.
rewrite <- Zmod_opp_opp.
apply H.
clear -Hy. omega.
@@ -1432,7 +1310,7 @@ apply H.
clear -Zy Hy. omega.
(* *)
split.
-pattern (Z2R (x / y)) at 1 ; rewrite <- Rplus_0_r.
+pattern (IZR (x / y)) at 1 ; rewrite <- Rplus_0_r.
apply Rplus_le_compat_l.
apply H.
apply Rplus_lt_compat_l.
@@ -1445,11 +1323,11 @@ Section pow.
Variable r : radix.
-Theorem radix_pos : (0 < Z2R r)%R.
+Theorem radix_pos : (0 < IZR r)%R.
Proof.
destruct r as (v, Hr). simpl.
-apply (Z2R_lt 0).
-apply Zlt_le_trans with 2%Z.
+apply IZR_lt.
+apply Z.lt_le_trans with 2%Z.
easy.
now apply Zle_bool_imp_le.
Qed.
@@ -1457,14 +1335,14 @@ Qed.
(** Well-used function called bpow for computing the power function #&beta;#^e *)
Definition bpow e :=
match e with
- | Zpos p => Z2R (Zpower_pos r p)
- | Zneg p => Rinv (Z2R (Zpower_pos r p))
+ | Zpos p => IZR (Zpower_pos r p)
+ | Zneg p => Rinv (IZR (Zpower_pos r p))
| Z0 => 1%R
end.
-Theorem Z2R_Zpower_pos :
+Theorem IZR_Zpower_pos :
forall n m,
- Z2R (Zpower_pos n m) = powerRZ (Z2R n) (Zpos m).
+ IZR (Zpower_pos n m) = powerRZ (IZR n) (Zpos m).
Proof.
intros.
rewrite Zpower_pos_nat.
@@ -1473,19 +1351,19 @@ induction (nat_of_P m).
apply refl_equal.
unfold Zpower_nat.
simpl.
-rewrite Z2R_mult.
+rewrite mult_IZR.
apply Rmult_eq_compat_l.
exact IHn0.
Qed.
Theorem bpow_powerRZ :
forall e,
- bpow e = powerRZ (Z2R r) e.
+ bpow e = powerRZ (IZR r) e.
Proof.
destruct e ; unfold bpow.
reflexivity.
-now rewrite Z2R_Zpower_pos.
-now rewrite Z2R_Zpower_pos.
+now rewrite IZR_Zpower_pos.
+now rewrite IZR_Zpower_pos.
Qed.
Theorem bpow_ge_0 :
@@ -1517,14 +1395,14 @@ apply radix_pos.
Qed.
Theorem bpow_1 :
- bpow 1 = Z2R r.
+ bpow 1 = IZR r.
Proof.
unfold bpow, Zpower_pos. simpl.
now rewrite Zmult_1_r.
Qed.
-Theorem bpow_plus1 :
- forall e : Z, (bpow (e + 1) = Z2R r * bpow e)%R.
+Theorem bpow_plus_1 :
+ forall e : Z, (bpow (e + 1) = IZR r * bpow e)%R.
Proof.
intros.
rewrite <- bpow_1.
@@ -1544,9 +1422,9 @@ apply Rgt_not_eq.
apply (bpow_gt_0 (Zpos p)).
Qed.
-Theorem Z2R_Zpower_nat :
+Theorem IZR_Zpower_nat :
forall e : nat,
- Z2R (Zpower_nat r e) = bpow (Z_of_nat e).
+ IZR (Zpower_nat r e) = bpow (Z_of_nat e).
Proof.
intros [|e].
split.
@@ -1555,10 +1433,10 @@ rewrite <- Zpower_pos_nat.
now rewrite <- Zpos_eq_Z_of_nat_o_nat_of_P.
Qed.
-Theorem Z2R_Zpower :
+Theorem IZR_Zpower :
forall e : Z,
(0 <= e)%Z ->
- Z2R (Zpower r e) = bpow e.
+ IZR (Zpower r e) = bpow e.
Proof.
intros [|e|e] H.
split.
@@ -1579,8 +1457,8 @@ apply bpow_gt_0.
assert (0 < e2 - e1)%Z by omega.
destruct (e2 - e1)%Z ; try discriminate H0.
clear.
-rewrite <- Z2R_Zpower by easy.
-apply (Z2R_lt 1).
+rewrite <- IZR_Zpower by easy.
+apply IZR_lt.
now apply Zpower_gt_1.
Qed.
@@ -1589,7 +1467,7 @@ Theorem lt_bpow :
(bpow e1 < bpow e2)%R -> (e1 < e2)%Z.
Proof.
intros e1 e2 H.
-apply Zgt_lt.
+apply Z.gt_lt.
apply Znot_le_gt.
intros H'.
apply Rlt_not_le with (1 := H).
@@ -1608,7 +1486,7 @@ intros e1 e2 H.
apply Rnot_lt_le.
intros H'.
apply Zle_not_gt with (1 := H).
-apply Zlt_gt.
+apply Z.lt_gt.
now apply lt_bpow.
Qed.
@@ -1621,7 +1499,7 @@ apply Znot_gt_le.
intros H'.
apply Rle_not_lt with (1 := H).
apply bpow_lt.
-now apply Zgt_lt.
+now apply Z.gt_lt.
Qed.
Theorem bpow_inj :
@@ -1638,15 +1516,15 @@ Qed.
Theorem bpow_exp :
forall e : Z,
- bpow e = exp (Z2R e * ln (Z2R r)).
+ bpow e = exp (IZR e * ln (IZR r)).
Proof.
(* positive case *)
-assert (forall e, bpow (Zpos e) = exp (Z2R (Zpos e) * ln (Z2R r))).
+assert (forall e, bpow (Zpos e) = exp (IZR (Zpos e) * ln (IZR r))).
intros e.
unfold bpow.
rewrite Zpower_pos_nat.
-unfold Z2R at 2.
-rewrite P2R_INR.
+rewrite <- positive_nat_Z.
+rewrite <- INR_IZR_INZ.
induction (nat_of_P e).
rewrite Rmult_0_l.
now rewrite exp_0.
@@ -1657,7 +1535,7 @@ rewrite exp_plus.
rewrite Rmult_1_l.
rewrite exp_ln.
rewrite <- IHn.
-rewrite <- Z2R_mult.
+rewrite <- mult_IZR.
now rewrite Zmult_comm.
apply radix_pos.
(* general case *)
@@ -1666,31 +1544,50 @@ rewrite Rmult_0_l.
now rewrite exp_0.
apply H.
unfold bpow.
-change (Z2R (Zpower_pos r e)) with (bpow (Zpos e)).
+change (IZR (Zpower_pos r e)) with (bpow (Zpos e)).
rewrite H.
rewrite <- exp_Ropp.
rewrite <- Ropp_mult_distr_l_reverse.
-now rewrite <- Z2R_opp.
+now rewrite <- opp_IZR.
+Qed.
+
+Lemma sqrt_bpow :
+ forall e,
+ sqrt (bpow (2 * e)) = bpow e.
+Proof.
+intro e.
+change 2%Z with (1 + 1)%Z; rewrite Z.mul_add_distr_r, Z.mul_1_l, bpow_plus.
+apply sqrt_square, bpow_ge_0.
Qed.
-(** Another well-used function for having the logarithm of a real number x to the base #&beta;# *)
-Record ln_beta_prop x := {
- ln_beta_val :> Z ;
- _ : (x <> 0)%R -> (bpow (ln_beta_val - 1)%Z <= Rabs x < bpow ln_beta_val)%R
+Lemma sqrt_bpow_ge :
+ forall e,
+ (bpow (e / 2) <= sqrt (bpow e))%R.
+Proof.
+intro e.
+rewrite <- (sqrt_square (bpow _)); [|now apply bpow_ge_0].
+apply sqrt_le_1_alt; rewrite <- bpow_plus; apply bpow_le.
+now replace (_ + _)%Z with (2 * (e / 2))%Z by ring; apply Z_mult_div_ge.
+Qed.
+
+(** Another well-used function for having the magnitude of a real number x to the base #&beta;# *)
+Record mag_prop x := {
+ mag_val :> Z ;
+ _ : (x <> 0)%R -> (bpow (mag_val - 1)%Z <= Rabs x < bpow mag_val)%R
}.
-Definition ln_beta :
- forall x : R, ln_beta_prop x.
+Definition mag :
+ forall x : R, mag_prop x.
Proof.
intros x.
-set (fact := ln (Z2R r)).
+set (fact := ln (IZR r)).
(* . *)
assert (0 < fact)%R.
apply exp_lt_inv.
rewrite exp_0.
unfold fact.
rewrite exp_ln.
-apply (Z2R_lt 1).
+apply IZR_lt.
apply radix_gt_1.
apply radix_pos.
(* . *)
@@ -1703,19 +1600,19 @@ rewrite 2!bpow_exp.
fold fact.
pattern x at 2 3 ; replace x with (exp (ln x * / fact * fact)).
split.
-rewrite Z2R_minus.
+rewrite minus_IZR.
apply exp_le.
apply Rmult_le_compat_r.
now apply Rlt_le.
unfold Rminus.
-rewrite Z2R_plus.
+rewrite plus_IZR.
rewrite Rplus_assoc.
rewrite Rplus_opp_r, Rplus_0_r.
apply Zfloor_lb.
apply exp_increasing.
apply Rmult_lt_compat_r.
exact H.
-rewrite Z2R_plus.
+rewrite plus_IZR.
apply Zfloor_ub.
rewrite Rmult_assoc.
rewrite Rinv_l.
@@ -1748,55 +1645,55 @@ apply Zle_antisym ;
assumption.
Qed.
-Theorem ln_beta_unique :
+Theorem mag_unique :
forall (x : R) (e : Z),
(bpow (e - 1) <= Rabs x < bpow e)%R ->
- ln_beta x = e :> Z.
+ mag x = e :> Z.
Proof.
intros x e1 He.
destruct (Req_dec x 0) as [Hx|Hx].
elim Rle_not_lt with (1 := proj1 He).
rewrite Hx, Rabs_R0.
apply bpow_gt_0.
-destruct (ln_beta x) as (e2, Hx2).
+destruct (mag x) as (e2, Hx2).
simpl.
apply bpow_unique with (2 := He).
now apply Hx2.
Qed.
-Theorem ln_beta_opp :
+Theorem mag_opp :
forall x,
- ln_beta (-x) = ln_beta x :> Z.
+ mag (-x) = mag x :> Z.
Proof.
intros x.
destruct (Req_dec x 0) as [Hx|Hx].
now rewrite Hx, Ropp_0.
-destruct (ln_beta x) as (e, He).
+destruct (mag x) as (e, He).
simpl.
specialize (He Hx).
-apply ln_beta_unique.
+apply mag_unique.
now rewrite Rabs_Ropp.
Qed.
-Theorem ln_beta_abs :
+Theorem mag_abs :
forall x,
- ln_beta (Rabs x) = ln_beta x :> Z.
+ mag (Rabs x) = mag x :> Z.
Proof.
intros x.
unfold Rabs.
case Rcase_abs ; intros _.
-apply ln_beta_opp.
+apply mag_opp.
apply refl_equal.
Qed.
-Theorem ln_beta_unique_pos :
+Theorem mag_unique_pos :
forall (x : R) (e : Z),
(bpow (e - 1) <= x < bpow e)%R ->
- ln_beta x = e :> Z.
+ mag x = e :> Z.
Proof.
intros x e1 He1.
-rewrite <- ln_beta_abs.
-apply ln_beta_unique.
+rewrite <- mag_abs.
+apply mag_unique.
rewrite 2!Rabs_pos_eq.
exact He1.
apply Rle_trans with (2 := proj1 He1).
@@ -1804,14 +1701,14 @@ apply bpow_ge_0.
apply Rabs_pos.
Qed.
-Theorem ln_beta_le_abs :
+Theorem mag_le_abs :
forall x y,
(x <> 0)%R -> (Rabs x <= Rabs y)%R ->
- (ln_beta x <= ln_beta y)%Z.
+ (mag x <= mag y)%Z.
Proof.
intros x y H0x Hxy.
-destruct (ln_beta x) as (ex, Hx).
-destruct (ln_beta y) as (ey, Hy).
+destruct (mag x) as (ex, Hx).
+destruct (mag y) as (ey, Hy).
simpl.
apply bpow_lt_bpow.
specialize (Hx H0x).
@@ -1825,13 +1722,13 @@ rewrite Hy', Rabs_R0.
apply Rle_refl.
Qed.
-Theorem ln_beta_le :
+Theorem mag_le :
forall x y,
(0 < x)%R -> (x <= y)%R ->
- (ln_beta x <= ln_beta y)%Z.
+ (mag x <= mag y)%Z.
Proof.
intros x y H0x Hxy.
-apply ln_beta_le_abs.
+apply mag_le_abs.
now apply Rgt_not_eq.
rewrite 2!Rabs_pos_eq.
exact Hxy.
@@ -1840,17 +1737,17 @@ now apply Rlt_le.
now apply Rlt_le.
Qed.
-Lemma ln_beta_lt_pos :
+Lemma lt_mag :
forall x y,
(0 < y)%R ->
- (ln_beta x < ln_beta y)%Z -> (x < y)%R.
+ (mag x < mag y)%Z -> (x < y)%R.
Proof.
intros x y Py.
case (Rle_or_lt x 0); intros Px.
intros H.
now apply Rle_lt_trans with 0%R.
-destruct (ln_beta x) as (ex, Hex).
-destruct (ln_beta y) as (ey, Hey).
+destruct (mag x) as (ex, Hex).
+destruct (mag y) as (ey, Hey).
simpl.
intro H.
destruct Hex as (_,Hex); [now apply Rgt_not_eq|].
@@ -1862,11 +1759,11 @@ apply Rle_trans with (bpow (ey - 1)); [|exact Hey].
now apply bpow_le; omega.
Qed.
-Theorem ln_beta_bpow :
- forall e, (ln_beta (bpow e) = e + 1 :> Z)%Z.
+Theorem mag_bpow :
+ forall e, (mag (bpow e) = e + 1 :> Z)%Z.
Proof.
intros e.
-apply ln_beta_unique.
+apply mag_unique.
rewrite Rabs_right.
replace (e + 1 - 1)%Z with e by ring.
split.
@@ -1877,14 +1774,14 @@ apply Rle_ge.
apply bpow_ge_0.
Qed.
-Theorem ln_beta_mult_bpow :
+Theorem mag_mult_bpow :
forall x e, x <> 0%R ->
- (ln_beta (x * bpow e) = ln_beta x + e :>Z)%Z.
+ (mag (x * bpow e) = mag x + e :>Z)%Z.
Proof.
intros x e Zx.
-destruct (ln_beta x) as (ex, Ex) ; simpl.
+destruct (mag x) as (ex, Ex) ; simpl.
specialize (Ex Zx).
-apply ln_beta_unique.
+apply mag_unique.
rewrite Rabs_mult.
rewrite (Rabs_pos_eq (bpow e)) by apply bpow_ge_0.
split.
@@ -1899,26 +1796,26 @@ apply bpow_gt_0.
apply Ex.
Qed.
-Theorem ln_beta_le_bpow :
+Theorem mag_le_bpow :
forall x e,
x <> 0%R ->
(Rabs x < bpow e)%R ->
- (ln_beta x <= e)%Z.
+ (mag x <= e)%Z.
Proof.
intros x e Zx Hx.
-destruct (ln_beta x) as (ex, Ex) ; simpl.
+destruct (mag x) as (ex, Ex) ; simpl.
specialize (Ex Zx).
apply bpow_lt_bpow.
now apply Rle_lt_trans with (Rabs x).
Qed.
-Theorem ln_beta_gt_bpow :
+Theorem mag_gt_bpow :
forall x e,
(bpow e <= Rabs x)%R ->
- (e < ln_beta x)%Z.
+ (e < mag x)%Z.
Proof.
intros x e Hx.
-destruct (ln_beta x) as (ex, Ex) ; simpl.
+destruct (mag x) as (ex, Ex) ; simpl.
apply lt_bpow.
apply Rle_lt_trans with (1 := Hx).
apply Ex.
@@ -1928,92 +1825,92 @@ rewrite Zx, Rabs_R0.
apply bpow_gt_0.
Qed.
-Theorem ln_beta_ge_bpow :
+Theorem mag_ge_bpow :
forall x e,
(bpow (e - 1) <= Rabs x)%R ->
- (e <= ln_beta x)%Z.
+ (e <= mag x)%Z.
Proof.
intros x e H.
destruct (Rlt_or_le (Rabs x) (bpow e)) as [Hxe|Hxe].
- (* Rabs x w bpow e *)
- assert (ln_beta x = e :> Z) as Hln.
- now apply ln_beta_unique; split.
+ assert (mag x = e :> Z) as Hln.
+ now apply mag_unique; split.
rewrite Hln.
now apply Z.le_refl.
- (* bpow e <= Rabs x *)
apply Zlt_le_weak.
- now apply ln_beta_gt_bpow.
+ now apply mag_gt_bpow.
Qed.
-Theorem bpow_ln_beta_gt :
+Theorem bpow_mag_gt :
forall x,
- (Rabs x < bpow (ln_beta x))%R.
+ (Rabs x < bpow (mag x))%R.
Proof.
intros x.
destruct (Req_dec x 0) as [Zx|Zx].
rewrite Zx, Rabs_R0.
apply bpow_gt_0.
-destruct (ln_beta x) as (ex, Ex) ; simpl.
+destruct (mag x) as (ex, Ex) ; simpl.
now apply Ex.
Qed.
-Theorem bpow_ln_beta_le :
+Theorem bpow_mag_le :
forall x, (x <> 0)%R ->
- (bpow (ln_beta x-1) <= Rabs x)%R.
+ (bpow (mag x-1) <= Rabs x)%R.
Proof.
intros x Hx.
-destruct (ln_beta x) as (ex, Ex) ; simpl.
+destruct (mag x) as (ex, Ex) ; simpl.
now apply Ex.
Qed.
-Theorem ln_beta_le_Zpower :
+Theorem mag_le_Zpower :
forall m e,
m <> Z0 ->
- (Zabs m < Zpower r e)%Z->
- (ln_beta (Z2R m) <= e)%Z.
+ (Z.abs m < Zpower r e)%Z->
+ (mag (IZR m) <= e)%Z.
Proof.
intros m e Zm Hm.
-apply ln_beta_le_bpow.
-exact (Z2R_neq m 0 Zm).
+apply mag_le_bpow.
+now apply IZR_neq.
destruct (Zle_or_lt 0 e).
-rewrite <- Z2R_abs, <- Z2R_Zpower with (1 := H).
-now apply Z2R_lt.
+rewrite <- abs_IZR, <- IZR_Zpower with (1 := H).
+now apply IZR_lt.
elim Zm.
-cut (Zabs m < 0)%Z.
+cut (Z.abs m < 0)%Z.
now case m.
clear -Hm H.
now destruct e.
Qed.
-Theorem ln_beta_gt_Zpower :
+Theorem mag_gt_Zpower :
forall m e,
m <> Z0 ->
- (Zpower r e <= Zabs m)%Z ->
- (e < ln_beta (Z2R m))%Z.
+ (Zpower r e <= Z.abs m)%Z ->
+ (e < mag (IZR m))%Z.
Proof.
intros m e Zm Hm.
-apply ln_beta_gt_bpow.
-rewrite <- Z2R_abs.
+apply mag_gt_bpow.
+rewrite <- abs_IZR.
destruct (Zle_or_lt 0 e).
-rewrite <- Z2R_Zpower with (1 := H).
-now apply Z2R_le.
+rewrite <- IZR_Zpower with (1 := H).
+now apply IZR_le.
apply Rle_trans with (bpow 0).
apply bpow_le.
now apply Zlt_le_weak.
-apply (Z2R_le 1).
+apply IZR_le.
clear -Zm.
zify ; omega.
Qed.
-Lemma ln_beta_mult :
+Lemma mag_mult :
forall x y,
(x <> 0)%R -> (y <> 0)%R ->
- (ln_beta x + ln_beta y - 1 <= ln_beta (x * y) <= ln_beta x + ln_beta y)%Z.
+ (mag x + mag y - 1 <= mag (x * y) <= mag x + mag y)%Z.
Proof.
intros x y Hx Hy.
-destruct (ln_beta x) as (ex, Hx2).
-destruct (ln_beta y) as (ey, Hy2).
+destruct (mag x) as (ex, Hx2).
+destruct (mag y) as (ey, Hy2).
simpl.
destruct (Hx2 Hx) as (Hx21,Hx22); clear Hx2.
destruct (Hy2 Hy) as (Hy21,Hy22); clear Hy2.
@@ -2029,26 +1926,26 @@ assert (Hxy2 : (Rabs (x * y) < bpow (ex + ey))%R).
now apply Rle_trans with (bpow (ex - 1)); try apply bpow_ge_0.
now apply Rle_trans with (bpow (ey - 1)); try apply bpow_ge_0. }
split.
-- now apply ln_beta_ge_bpow.
-- apply ln_beta_le_bpow.
+- now apply mag_ge_bpow.
+- apply mag_le_bpow.
+ now apply Rmult_integral_contrapositive_currified.
+ assumption.
Qed.
-Lemma ln_beta_plus :
+Lemma mag_plus :
forall x y,
(0 < y)%R -> (y <= x)%R ->
- (ln_beta x <= ln_beta (x + y) <= ln_beta x + 1)%Z.
+ (mag x <= mag (x + y) <= mag x + 1)%Z.
Proof.
assert (Hr : (2 <= r)%Z).
{ destruct r as (beta_val,beta_prop).
now apply Zle_bool_imp_le. }
intros x y Hy Hxy.
assert (Hx : (0 < x)%R) by apply (Rlt_le_trans _ _ _ Hy Hxy).
-destruct (ln_beta x) as (ex,Hex); simpl.
+destruct (mag x) as (ex,Hex); simpl.
destruct Hex as (Hex0,Hex1); [now apply Rgt_not_eq|].
assert (Haxy : (Rabs (x + y) < bpow (ex + 1))%R).
-{ rewrite bpow_plus1.
+{ rewrite bpow_plus_1.
apply Rlt_le_trans with (2 * bpow ex)%R.
- rewrite Rabs_pos_eq.
apply Rle_lt_trans with (2 * Rabs x)%R.
@@ -2062,7 +1959,7 @@ assert (Haxy : (Rabs (x + y) < bpow (ex + 1))%R).
now apply Rlt_le, Rplus_lt_compat.
- apply Rmult_le_compat_r.
now apply bpow_ge_0.
- now apply (Z2R_le 2). }
+ now apply IZR_le. }
assert (Haxy2 : (bpow (ex - 1) <= Rabs (x + y))%R).
{ apply (Rle_trans _ _ _ Hex0).
rewrite Rabs_right; [|now apply Rgt_ge].
@@ -2071,20 +1968,20 @@ assert (Haxy2 : (bpow (ex - 1) <= Rabs (x + y))%R).
apply Rplus_le_compat_l.
now apply Rlt_le. }
split.
-- now apply ln_beta_ge_bpow.
-- apply ln_beta_le_bpow.
+- now apply mag_ge_bpow.
+- apply mag_le_bpow.
+ now apply tech_Rplus; [apply Rlt_le|].
+ exact Haxy.
Qed.
-Lemma ln_beta_minus :
+Lemma mag_minus :
forall x y,
(0 < y)%R -> (y < x)%R ->
- (ln_beta (x - y) <= ln_beta x)%Z.
+ (mag (x - y) <= mag x)%Z.
Proof.
intros x y Py Hxy.
assert (Px : (0 < x)%R) by apply (Rlt_trans _ _ _ Py Hxy).
-apply ln_beta_le.
+apply mag_le.
- now apply Rlt_Rminus.
- rewrite <- (Rplus_0_r x) at 2.
apply Rplus_le_compat_l.
@@ -2092,19 +1989,19 @@ apply ln_beta_le.
now apply Ropp_le_contravar; apply Rlt_le.
Qed.
-Lemma ln_beta_minus_lb :
+Lemma mag_minus_lb :
forall x y,
(0 < x)%R -> (0 < y)%R ->
- (ln_beta y <= ln_beta x - 2)%Z ->
- (ln_beta x - 1 <= ln_beta (x - y))%Z.
+ (mag y <= mag x - 2)%Z ->
+ (mag x - 1 <= mag (x - y))%Z.
Proof.
assert (Hbeta : (2 <= r)%Z).
{ destruct r as (beta_val,beta_prop).
now apply Zle_bool_imp_le. }
intros x y Px Py Hln.
-assert (Oxy : (y < x)%R); [apply ln_beta_lt_pos;[assumption|omega]|].
-destruct (ln_beta x) as (ex,Hex).
-destruct (ln_beta y) as (ey,Hey).
+assert (Oxy : (y < x)%R); [apply lt_mag;[assumption|omega]|].
+destruct (mag x) as (ex,Hex).
+destruct (mag y) as (ey,Hey).
simpl in Hln |- *.
destruct Hex as (Hex,_); [now apply Rgt_not_eq|].
destruct Hey as (_,Hey); [now apply Rgt_not_eq|].
@@ -2112,9 +2009,9 @@ assert (Hbx : (bpow (ex - 2) + bpow (ex - 2) <= x)%R).
{ ring_simplify.
apply Rle_trans with (bpow (ex - 1)).
- replace (ex - 1)%Z with (ex - 2 + 1)%Z by ring.
- rewrite bpow_plus1.
+ rewrite bpow_plus_1.
apply Rmult_le_compat_r; [now apply bpow_ge_0|].
- now change 2%R with (Z2R 2); apply Z2R_le.
+ now apply IZR_le.
- now rewrite Rabs_right in Hex; [|apply Rle_ge; apply Rlt_le]. }
assert (Hby : (y < bpow (ex - 2))%R).
{ apply Rlt_le_trans with (bpow ey).
@@ -2126,98 +2023,95 @@ assert (Hbxy : (bpow (ex - 2) <= x - y)%R).
replace (bpow (ex - 2))%R with (bpow (ex - 2) + bpow (ex - 2)
- bpow (ex - 2))%R by ring.
now apply Rplus_le_compat. }
-apply ln_beta_ge_bpow.
+apply mag_ge_bpow.
replace (ex - 1 - 1)%Z with (ex - 2)%Z by ring.
now apply Rabs_ge; right.
Qed.
-Lemma ln_beta_div :
+Lemma mag_div :
forall x y : R,
- (0 < x)%R -> (0 < y)%R ->
- (ln_beta x - ln_beta y <= ln_beta (x / y) <= ln_beta x - ln_beta y + 1)%Z.
+ x <> 0%R -> y <> 0%R ->
+ (mag x - mag y <= mag (x / y) <= mag x - mag y + 1)%Z.
Proof.
intros x y Px Py.
-destruct (ln_beta x) as (ex,Hex).
-destruct (ln_beta y) as (ey,Hey).
+destruct (mag x) as (ex,Hex).
+destruct (mag y) as (ey,Hey).
simpl.
unfold Rdiv.
-rewrite Rabs_right in Hex; [|now apply Rle_ge; apply Rlt_le].
-rewrite Rabs_right in Hey; [|now apply Rle_ge; apply Rlt_le].
-assert (Heiy : (bpow (- ey) < / y <= bpow (- ey + 1))%R).
-{ split.
+assert (Heiy : (bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R).
+{ rewrite Rabs_Rinv by easy.
+ split.
- rewrite bpow_opp.
apply Rinv_lt_contravar.
- + apply Rmult_lt_0_compat; [exact Py|].
+ + apply Rmult_lt_0_compat.
+ now apply Rabs_pos_lt.
now apply bpow_gt_0.
- + apply Hey.
- now apply Rgt_not_eq.
+ + now apply Hey.
- replace (_ + _)%Z with (- (ey - 1))%Z by ring.
rewrite bpow_opp.
apply Rinv_le; [now apply bpow_gt_0|].
- apply Hey.
- now apply Rgt_not_eq. }
+ now apply Hey. }
split.
-- apply ln_beta_ge_bpow.
- apply Rabs_ge; right.
+- apply mag_ge_bpow.
replace (_ - _)%Z with (ex - 1 - ey)%Z by ring.
unfold Zminus at 1; rewrite bpow_plus.
+ rewrite Rabs_mult.
apply Rmult_le_compat.
+ now apply bpow_ge_0.
+ now apply bpow_ge_0.
- + apply Hex.
- now apply Rgt_not_eq.
- + apply Rlt_le; apply Heiy.
-- assert (Pxy : (0 < x * / y)%R).
- { apply Rmult_lt_0_compat; [exact Px|].
- now apply Rinv_0_lt_compat. }
- apply ln_beta_le_bpow.
- + now apply Rgt_not_eq.
- + rewrite Rabs_right; [|now apply Rle_ge; apply Rlt_le].
- replace (_ + 1)%Z with (ex + (- ey + 1))%Z by ring.
+ + now apply Hex.
+ + now apply Rlt_le; apply Heiy.
+- apply mag_le_bpow.
+ + apply Rmult_integral_contrapositive_currified.
+ exact Px.
+ now apply Rinv_neq_0_compat.
+ + replace (_ + 1)%Z with (ex + (- ey + 1))%Z by ring.
rewrite bpow_plus.
- apply Rlt_le_trans with (bpow ex * / y)%R.
- * apply Rmult_lt_compat_r; [now apply Rinv_0_lt_compat|].
- apply Hex.
- now apply Rgt_not_eq.
+ apply Rlt_le_trans with (bpow ex * Rabs (/ y))%R.
+ * rewrite Rabs_mult.
+ apply Rmult_lt_compat_r.
+ apply Rabs_pos_lt.
+ now apply Rinv_neq_0_compat.
+ now apply Hex.
* apply Rmult_le_compat_l; [now apply bpow_ge_0|].
apply Heiy.
Qed.
-Lemma ln_beta_sqrt :
+Lemma mag_sqrt :
forall x,
(0 < x)%R ->
- (2 * ln_beta (sqrt x) - 1 <= ln_beta x <= 2 * ln_beta (sqrt x))%Z.
+ mag (sqrt x) = Z.div2 (mag x + 1) :> Z.
Proof.
intros x Px.
-assert (H : (bpow (2 * ln_beta (sqrt x) - 1 - 1) <= Rabs x
- < bpow (2 * ln_beta (sqrt x)))%R).
-{ split.
- - apply Rge_le; rewrite <- (sqrt_def x) at 1;
- [|now apply Rlt_le]; apply Rle_ge.
- rewrite Rabs_mult.
- replace (_ - _)%Z with (ln_beta (sqrt x) - 1
- + (ln_beta (sqrt x) - 1))%Z by ring.
- rewrite bpow_plus.
- assert (H : (bpow (ln_beta (sqrt x) - 1) <= Rabs (sqrt x))%R).
- { destruct (ln_beta (sqrt x)) as (esx,Hesx); simpl.
- apply Hesx.
- apply Rgt_not_eq; apply Rlt_gt.
- now apply sqrt_lt_R0. }
- now apply Rmult_le_compat; [now apply bpow_ge_0|now apply bpow_ge_0| |].
- - rewrite <- (sqrt_def x) at 1; [|now apply Rlt_le].
- rewrite Rabs_mult.
- change 2%Z with (1 + 1)%Z; rewrite Zmult_plus_distr_l;
- rewrite Zmult_1_l.
- rewrite bpow_plus.
- assert (H : (Rabs (sqrt x) < bpow (ln_beta (sqrt x)))%R).
- { destruct (ln_beta (sqrt x)) as (esx,Hesx); simpl.
- apply Hesx.
- apply Rgt_not_eq; apply Rlt_gt.
- now apply sqrt_lt_R0. }
- now apply Rmult_lt_compat; [now apply Rabs_pos|now apply Rabs_pos| |]. }
+apply mag_unique.
+destruct mag as [e He].
+simpl.
+specialize (He (Rgt_not_eq _ _ Px)).
+rewrite Rabs_pos_eq in He by now apply Rlt_le.
split.
-- now apply ln_beta_ge_bpow.
-- now apply ln_beta_le_bpow; [now apply Rgt_not_eq|].
+- rewrite <- (Rabs_pos_eq (bpow _)) by apply bpow_ge_0.
+ apply Rsqr_le_abs_0.
+ rewrite Rsqr_sqrt by now apply Rlt_le.
+ apply Rle_trans with (2 := proj1 He).
+ unfold Rsqr ; rewrite <- bpow_plus.
+ apply bpow_le.
+ generalize (Zdiv2_odd_eqn (e + 1)).
+ destruct Z.odd ; intros ; omega.
+- rewrite <- (Rabs_pos_eq (bpow _)) by apply bpow_ge_0.
+ apply Rsqr_lt_abs_0.
+ rewrite Rsqr_sqrt by now apply Rlt_le.
+ apply Rlt_le_trans with (1 := proj2 He).
+ unfold Rsqr ; rewrite <- bpow_plus.
+ apply bpow_le.
+ generalize (Zdiv2_odd_eqn (e + 1)).
+ destruct Z.odd ; intros ; omega.
+Qed.
+
+Lemma mag_1 : mag 1 = 1%Z :> Z.
+Proof.
+apply mag_unique_pos; rewrite bpow_1; simpl; split; [now right|apply IZR_lt].
+assert (H := Zle_bool_imp_le _ _ (radix_prop r)); revert H.
+now apply Z.lt_le_trans.
Qed.
End pow.
@@ -2248,12 +2142,12 @@ Section cond_Ropp.
Definition cond_Ropp (b : bool) m := if b then Ropp m else m.
-Theorem Z2R_cond_Zopp :
+Theorem IZR_cond_Zopp :
forall b m,
- Z2R (cond_Zopp b m) = cond_Ropp b (Z2R m).
+ IZR (cond_Zopp b m) = cond_Ropp b (IZR m).
Proof.
intros [|] m.
-apply Z2R_opp.
+apply opp_IZR.
apply refl_equal.
Qed.
@@ -2286,22 +2180,6 @@ apply Ropp_involutive.
apply refl_equal.
Qed.
-Theorem cond_Ropp_even_function :
- forall {A : Type} (f : R -> A),
- (forall x, f (Ropp x) = f x) ->
- forall b x, f (cond_Ropp b x) = f x.
-Proof.
-now intros A f Hf [|] x ; simpl.
-Qed.
-
-Theorem cond_Ropp_odd_function :
- forall (f : R -> R),
- (forall x, f (Ropp x) = Ropp (f x)) ->
- forall b x, f (cond_Ropp b x) = cond_Ropp b (f x).
-Proof.
-now intros f Hf [|] x ; simpl.
-Qed.
-
Theorem cond_Ropp_inj :
forall b x y,
cond_Ropp b x = cond_Ropp b y -> x = y.
@@ -2391,7 +2269,7 @@ destruct (Rle_lt_dec l 0) as [Hl|Hl].
apply ub.
now apply HE.
left.
-set (N := Zabs_nat (up (/l) - 2)).
+set (N := Z.abs_nat (up (/l) - 2)).
exists N.
assert (HN: (INR N + 1 = IZR (up (/ l)) - 1)%R).
unfold N.
@@ -2399,7 +2277,7 @@ assert (HN: (INR N + 1 = IZR (up (/ l)) - 1)%R).
rewrite inj_Zabs_nat.
replace (IZR (up (/ l)) - 1)%R with (IZR (up (/ l) - 2) + 1)%R.
apply (f_equal (fun v => IZR v + 1)%R).
- apply Zabs_eq.
+ apply Z.abs_eq.
apply Zle_minus_le_0.
apply (Zlt_le_succ 1).
apply lt_IZR.
@@ -2484,10 +2362,10 @@ intros n; apply H.
destruct K as (n, Hn).
left; now exists (-Z.of_nat n)%Z.
right; intros n; case (Zle_or_lt 0 n); intros M.
-rewrite <- (Zabs_eq n); trivial.
+rewrite <- (Z.abs_eq n); trivial.
rewrite <- Zabs2Nat.id_abs.
apply J.
-rewrite <- (Zopp_involutive n).
+rewrite <- (Z.opp_involutive n).
rewrite <- (Z.abs_neq n).
rewrite <- Zabs2Nat.id_abs.
apply K.
diff --git a/flocq/Core/Fcore_rnd_ne.v b/flocq/Core/Round_NE.v
index 2d67e709..20b60ef5 100644
--- a/flocq/Core/Fcore_rnd_ne.v
+++ b/flocq/Core/Round_NE.v
@@ -2,9 +2,9 @@
This file is part of the Flocq formalization of floating-point
arithmetic in Coq: http://flocq.gforge.inria.fr/
-Copyright (C) 2010-2013 Sylvie Boldo
+Copyright (C) 2009-2018 Sylvie Boldo
#<br />#
-Copyright (C) 2010-2013 Guillaume Melquiond
+Copyright (C) 2009-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
@@ -18,14 +18,9 @@ COPYING file for more details.
*)
(** * Rounding to nearest, ties to even: existence, unicity... *)
-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 Raux Defs Round_pred Generic_fmt Float_prop Ulp.
-Notation ZnearestE := (Znearest (fun x => negb (Zeven x))).
+Notation ZnearestE := (Znearest (fun x => negb (Z.even x))).
Section Fcore_rnd_NE.
@@ -38,10 +33,10 @@ Variable fexp : Z -> Z.
Context { valid_exp : Valid_exp fexp }.
Notation format := (generic_format beta fexp).
-Notation canonic := (canonic beta fexp).
+Notation canonical := (canonical beta fexp).
Definition NE_prop (_ : R) f :=
- exists g : float beta, f = F2R g /\ canonic g /\ Zeven (Fnum g) = true.
+ exists g : float beta, f = F2R g /\ canonical g /\ Z.even (Fnum g) = true.
Definition Rnd_NE_pt :=
Rnd_NG_pt format NE_prop.
@@ -50,20 +45,20 @@ Definition DN_UP_parity_pos_prop :=
forall x xd xu,
(0 < x)%R ->
~ format x ->
- canonic xd ->
- canonic xu ->
+ canonical xd ->
+ canonical xu ->
F2R xd = round beta fexp Zfloor x ->
F2R xu = round beta fexp Zceil x ->
- Zeven (Fnum xu) = negb (Zeven (Fnum xd)).
+ Z.even (Fnum xu) = negb (Z.even (Fnum xd)).
Definition DN_UP_parity_prop :=
forall x xd xu,
~ format x ->
- canonic xd ->
- canonic xu ->
+ canonical xd ->
+ canonical xu ->
F2R xd = round beta fexp Zfloor x ->
F2R xu = round beta fexp Zceil x ->
- Zeven (Fnum xu) = negb (Zeven (Fnum xd)).
+ Z.even (Fnum xu) = negb (Z.even (Fnum xd)).
Lemma DN_UP_parity_aux :
DN_UP_parity_pos_prop ->
@@ -83,18 +78,18 @@ now rewrite Ropp_involutive, Ropp_0.
destruct xd as (md, ed).
destruct xu as (mu, eu).
simpl.
-rewrite <- (Bool.negb_involutive (Zeven mu)).
+rewrite <- (Bool.negb_involutive (Z.even mu)).
apply f_equal.
apply sym_eq.
-rewrite <- (Zeven_opp mu), <- (Zeven_opp md).
-change (Zeven (Fnum (Float beta (-md) ed)) = negb (Zeven (Fnum (Float beta (-mu) eu)))).
+rewrite <- (Z.even_opp mu), <- (Z.even_opp md).
+change (Z.even (Fnum (Float beta (-md) ed)) = negb (Z.even (Fnum (Float beta (-mu) eu)))).
apply (Hpos (-x)%R _ _ Hx').
intros H.
apply Hfx.
rewrite <- Ropp_involutive.
now apply generic_format_opp.
-now apply canonic_opp.
-now apply canonic_opp.
+now apply canonical_opp.
+now apply canonical_opp.
rewrite round_DN_opp, F2R_Zopp.
now apply f_equal.
rewrite round_UP_opp, F2R_Zopp.
@@ -102,7 +97,7 @@ now apply f_equal.
Qed.
Class Exists_NE :=
- exists_NE : Zeven beta = false \/ forall e,
+ exists_NE : Z.even beta = false \/ forall e,
((fexp e < e)%Z -> (fexp (e + 1) < e)%Z) /\ ((e <= fexp e)%Z -> fexp (fexp e + 1) = fexp e).
Context { exists_NE_ : Exists_NE }.
@@ -111,22 +106,22 @@ Theorem DN_UP_parity_generic_pos :
DN_UP_parity_pos_prop.
Proof with auto with typeclass_instances.
intros x xd xu H0x Hfx Hd Hu Hxd Hxu.
-destruct (ln_beta beta x) as (ex, Hexa).
+destruct (mag beta x) as (ex, Hexa).
specialize (Hexa (Rgt_not_eq _ _ H0x)).
generalize Hexa. intros Hex.
rewrite (Rabs_pos_eq _ (Rlt_le _ _ H0x)) in Hex.
destruct (Zle_or_lt ex (fexp ex)) as [Hxe|Hxe].
(* small x *)
assert (Hd3 : Fnum xd = Z0).
-apply F2R_eq_0_reg with beta (Fexp xd).
+apply eq_0_F2R with beta (Fexp xd).
change (F2R xd = R0).
rewrite Hxd.
apply round_DN_small_pos with (1 := Hex) (2 := Hxe).
assert (Hu3 : xu = Float beta (1 * Zpower beta (fexp ex - fexp (fexp ex + 1))) (fexp (fexp ex + 1))).
-apply canonic_unicity with (1 := Hu).
+apply canonical_unique with (1 := Hu).
apply (f_equal fexp).
rewrite <- F2R_change_exp.
-now rewrite F2R_bpow, ln_beta_bpow.
+now rewrite F2R_bpow, mag_bpow.
now apply valid_exp.
rewrite <- F2R_change_exp.
rewrite F2R_bpow.
@@ -172,10 +167,10 @@ rewrite Hxu.
apply round_bounded_large_pos...
(* - xu = bpow ex *)
assert (Hu3: xu = Float beta (1 * Zpower beta (ex - fexp (ex + 1))) (fexp (ex + 1))).
-apply canonic_unicity with (1 := Hu).
+apply canonical_unique with (1 := Hu).
apply (f_equal fexp).
rewrite <- F2R_change_exp.
-now rewrite F2R_bpow, ln_beta_bpow.
+now rewrite F2R_bpow, mag_bpow.
now apply valid_exp.
rewrite <- Hu2.
apply sym_eq.
@@ -185,15 +180,15 @@ exact Hxe2.
assert (Hd3: xd = Float beta (Zpower beta (ex - fexp ex) - 1) (fexp ex)).
assert (H: F2R xd = F2R (Float beta (Zpower beta (ex - fexp ex) - 1) (fexp ex))).
unfold F2R. simpl.
-rewrite Z2R_minus.
+rewrite minus_IZR.
unfold Rminus.
rewrite Rmult_plus_distr_r.
-rewrite Z2R_Zpower, <- bpow_plus.
+rewrite IZR_Zpower, <- bpow_plus.
ring_simplify (ex - fexp ex + fexp ex)%Z.
rewrite Hu2, Hud.
rewrite ulp_neq_0;[idtac|now apply Rgt_not_eq].
-unfold canonic_exp.
-rewrite ln_beta_unique with beta x ex.
+unfold cexp.
+rewrite mag_unique with beta x ex.
unfold F2R.
simpl. ring.
rewrite Rabs_pos_eq.
@@ -201,25 +196,25 @@ exact Hex.
now apply Rlt_le.
apply Zle_minus_le_0.
now apply Zlt_le_weak.
-apply canonic_unicity with (1 := Hd) (3 := H).
+apply canonical_unique with (1 := Hd) (3 := H).
apply (f_equal fexp).
rewrite <- H.
apply sym_eq.
-now apply ln_beta_unique.
+now apply mag_unique.
rewrite Hd3, Hu3.
unfold Fnum.
-rewrite Zeven_mult. simpl.
+rewrite Z.even_mul. simpl.
unfold Zminus at 2.
-rewrite Zeven_plus.
+rewrite Z.even_add.
rewrite eqb_sym. simpl.
-fold (negb (Zeven (beta ^ (ex - fexp ex)))).
+fold (negb (Z.even (beta ^ (ex - fexp ex)))).
rewrite Bool.negb_involutive.
-rewrite (Zeven_Zpower beta (ex - fexp ex)). 2: omega.
+rewrite (Z.even_pow beta (ex - fexp ex)). 2: omega.
destruct exists_NE_.
rewrite H.
apply Zeven_Zpower_odd with (2 := H).
now apply Zle_minus_le_0.
-apply Zeven_Zpower.
+apply Z.even_pow.
specialize (H ex).
omega.
(* - xu < bpow ex *)
@@ -227,17 +222,17 @@ revert Hud.
rewrite ulp_neq_0;[idtac|now apply Rgt_not_eq].
unfold F2R.
rewrite Hd, Hu.
-unfold canonic_exp.
-rewrite ln_beta_unique with beta (F2R xu) ex.
-rewrite ln_beta_unique with (1 := Hd4).
-rewrite ln_beta_unique with (1 := Hexa).
+unfold cexp.
+rewrite mag_unique with beta (F2R xu) ex.
+rewrite mag_unique with (1 := Hd4).
+rewrite mag_unique with (1 := Hexa).
intros H.
replace (Fnum xu) with (Fnum xd + 1)%Z.
-rewrite Zeven_plus.
+rewrite Z.even_add.
now apply eqb_sym.
apply sym_eq.
-apply eq_Z2R.
-rewrite Z2R_plus.
+apply eq_IZR.
+rewrite plus_IZR.
apply Rmult_eq_reg_r with (bpow (fexp ex)).
rewrite H.
simpl. ring.
@@ -270,38 +265,38 @@ now apply generic_format_satisfies_any.
intros x d u Hf Hd Hu.
generalize (proj1 Hd).
unfold generic_format.
-set (ed := canonic_exp beta fexp d).
+set (ed := cexp beta fexp d).
set (md := Ztrunc (scaled_mantissa beta fexp d)).
intros Hd1.
-case_eq (Zeven md) ; [ intros He | intros Ho ].
+case_eq (Z.even md) ; [ intros He | intros Ho ].
right.
exists (Float beta md ed).
-unfold Fcore_generic_fmt.canonic.
+unfold Generic_fmt.canonical.
rewrite <- Hd1.
now repeat split.
left.
generalize (proj1 Hu).
unfold generic_format.
-set (eu := canonic_exp beta fexp u).
+set (eu := cexp beta fexp u).
set (mu := Ztrunc (scaled_mantissa beta fexp u)).
intros Hu1.
rewrite Hu1.
eexists ; repeat split.
-unfold Fcore_generic_fmt.canonic.
+unfold Generic_fmt.canonical.
now rewrite <- Hu1.
rewrite (DN_UP_parity_generic x (Float beta md ed) (Float beta mu eu)).
simpl.
now rewrite Ho.
exact Hf.
-unfold Fcore_generic_fmt.canonic.
+unfold Generic_fmt.canonical.
now rewrite <- Hd1.
-unfold Fcore_generic_fmt.canonic.
+unfold Generic_fmt.canonical.
now rewrite <- Hu1.
rewrite <- Hd1.
-apply Rnd_DN_pt_unicity with (1 := Hd).
+apply Rnd_DN_pt_unique with (1 := Hd).
now apply round_DN_pt.
rewrite <- Hu1.
-apply Rnd_UP_pt_unicity with (1 := Hu).
+apply Rnd_UP_pt_unique with (1 := Hu).
now apply round_UP_pt.
Qed.
@@ -323,15 +318,16 @@ apply Hx.
apply sym_eq.
now apply Rnd_DN_pt_idempotent with (1 := Hd).
rewrite <- Hd1.
-apply Rnd_DN_pt_unicity with (1 := Hd).
+apply Rnd_DN_pt_unique with (1 := Hd).
now apply round_DN_pt.
rewrite <- Hu1.
-apply Rnd_UP_pt_unicity with (1 := Hu).
+apply Rnd_UP_pt_unique with (1 := Hu).
now apply round_UP_pt.
Qed.
Theorem Rnd_NE_pt_round :
round_pred Rnd_NE_pt.
+Proof.
split.
apply Rnd_NE_pt_total.
apply Rnd_NE_pt_monotone.
@@ -348,14 +344,14 @@ now apply round_N_pt.
unfold NE_prop.
set (mx := scaled_mantissa beta fexp x).
set (xr := round beta fexp ZnearestE x).
-destruct (Req_dec (mx - Z2R (Zfloor mx)) (/2)) as [Hm|Hm].
+destruct (Req_dec (mx - IZR (Zfloor mx)) (/2)) as [Hm|Hm].
(* midpoint *)
left.
-exists (Float beta (Ztrunc (scaled_mantissa beta fexp xr)) (canonic_exp beta fexp xr)).
+exists (Float beta (Ztrunc (scaled_mantissa beta fexp xr)) (cexp beta fexp xr)).
split.
apply round_N_pt...
split.
-unfold Fcore_generic_fmt.canonic. simpl.
+unfold Generic_fmt.canonical. simpl.
apply f_equal.
apply round_N_pt...
simpl.
@@ -363,23 +359,22 @@ unfold xr, round, Znearest.
fold mx.
rewrite Hm.
rewrite Rcompare_Eq. 2: apply refl_equal.
-case_eq (Zeven (Zfloor mx)) ; intros Hmx.
+case_eq (Z.even (Zfloor mx)) ; intros Hmx.
(* . even floor *)
-change (Zeven (Ztrunc (scaled_mantissa beta fexp (round beta fexp Zfloor x))) = true).
+change (Z.even (Ztrunc (scaled_mantissa beta fexp (round beta fexp Zfloor x))) = true).
destruct (Rle_or_lt (round beta fexp Zfloor x) 0) as [Hr|Hr].
rewrite (Rle_antisym _ _ Hr).
unfold scaled_mantissa.
rewrite Rmult_0_l.
-change 0%R with (Z2R 0).
-now rewrite (Ztrunc_Z2R 0).
+now rewrite Ztrunc_IZR.
rewrite <- (round_0 beta fexp Zfloor).
apply round_le...
now apply Rlt_le.
rewrite scaled_mantissa_DN...
-now rewrite Ztrunc_Z2R.
+now rewrite Ztrunc_IZR.
(* . odd floor *)
-change (Zeven (Ztrunc (scaled_mantissa beta fexp (round beta fexp Zceil x))) = true).
-destruct (ln_beta beta x) as (ex, Hex).
+change (Z.even (Ztrunc (scaled_mantissa beta fexp (round beta fexp Zceil x))) = true).
+destruct (mag beta x) as (ex, Hex).
specialize (Hex (Rgt_not_eq _ _ Hx)).
rewrite (Rabs_pos_eq _ (Rlt_le _ _ Hx)) in Hex.
destruct (Z_lt_le_dec (fexp ex) ex) as [He|He].
@@ -394,56 +389,56 @@ rewrite Rplus_opp_r in Hm.
elim (Rlt_irrefl 0).
rewrite Hm at 2.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
destruct (proj2 Hu) as [Hu'|Hu'].
(* ... u <> bpow *)
unfold scaled_mantissa.
-rewrite canonic_exp_fexp_pos with (1 := conj (proj1 Hu) Hu').
+rewrite cexp_fexp_pos with (1 := conj (proj1 Hu) Hu').
unfold round, F2R. simpl.
-rewrite canonic_exp_fexp_pos with (1 := Hex).
+rewrite cexp_fexp_pos with (1 := Hex).
rewrite Rmult_assoc, <- bpow_plus, Zplus_opp_r, Rmult_1_r.
-rewrite Ztrunc_Z2R.
+rewrite Ztrunc_IZR.
fold mx.
rewrite Hfc.
-now rewrite Zeven_plus, Hmx.
+now rewrite Z.even_add, Hmx.
(* ... u = bpow *)
rewrite Hu'.
-unfold scaled_mantissa, canonic_exp.
-rewrite ln_beta_bpow.
-rewrite <- bpow_plus, <- Z2R_Zpower.
-rewrite Ztrunc_Z2R.
-case_eq (Zeven beta) ; intros Hr.
+unfold scaled_mantissa, cexp.
+rewrite mag_bpow.
+rewrite <- bpow_plus, <- IZR_Zpower.
+rewrite Ztrunc_IZR.
+case_eq (Z.even beta) ; intros Hr.
destruct exists_NE_ as [Hs|Hs].
now rewrite Hs in Hr.
destruct (Hs ex) as (H,_).
-rewrite Zeven_Zpower.
+rewrite Z.even_pow.
exact Hr.
omega.
-assert (Zeven (Zfloor mx) = true). 2: now rewrite H in Hmx.
+assert (Z.even (Zfloor mx) = true). 2: now rewrite H in Hmx.
replace (Zfloor mx) with (Zceil mx + -1)%Z by omega.
-rewrite Zeven_plus.
+rewrite Z.even_add.
apply eqb_true.
unfold mx.
replace (Zceil (scaled_mantissa beta fexp x)) with (Zpower beta (ex - fexp ex)).
rewrite Zeven_Zpower_odd with (2 := Hr).
easy.
omega.
-apply eq_Z2R.
-rewrite Z2R_Zpower. 2: omega.
+apply eq_IZR.
+rewrite IZR_Zpower. 2: omega.
apply Rmult_eq_reg_r with (bpow (fexp ex)).
unfold Zminus.
rewrite bpow_plus.
rewrite Rmult_assoc, <- bpow_plus, Zplus_opp_l, Rmult_1_r.
-pattern (fexp ex) ; rewrite <- canonic_exp_fexp_pos with (1 := Hex).
+pattern (fexp ex) ; rewrite <- cexp_fexp_pos with (1 := Hex).
now apply sym_eq.
apply Rgt_not_eq.
apply bpow_gt_0.
generalize (proj1 (valid_exp ex) He).
omega.
(* .. small pos *)
-assert (Zeven (Zfloor mx) = true). 2: now rewrite H in Hmx.
+assert (Z.even (Zfloor mx) = true). 2: now rewrite H in Hmx.
unfold mx, scaled_mantissa.
-rewrite canonic_exp_fexp_pos with (1 := Hex).
+rewrite cexp_fexp_pos with (1 := Hex).
now rewrite mantissa_DN_small_pos.
(* not midpoint *)
right.
@@ -456,7 +451,7 @@ rewrite Hxg.
apply Hg.
set (d := round beta fexp Zfloor x).
set (u := round beta fexp Zceil x).
-apply Rnd_N_pt_unicity with (d := d) (u := u) (4 := Hg).
+apply Rnd_N_pt_unique with (d := d) (u := u) (4 := Hg).
now apply round_DN_pt.
now apply round_UP_pt.
2: now apply round_N_pt.
@@ -467,7 +462,7 @@ intros H.
apply Rmult_eq_reg_r in H.
apply Hm.
apply Rcompare_Eq_inv.
-rewrite Rcompare_floor_ceil_mid.
+rewrite Rcompare_floor_ceil_middle.
now apply Rcompare_Eq.
contradict Hxg.
apply sym_eq.
@@ -475,7 +470,7 @@ apply Rnd_N_pt_idempotent with (1 := Hg).
rewrite <- (scaled_mantissa_mult_bpow beta fexp x).
fold mx.
rewrite <- Hxg.
-change (Z2R (Zfloor mx) * bpow (canonic_exp beta fexp x))%R with d.
+change (IZR (Zfloor mx) * bpow (cexp beta fexp x))%R with d.
now eapply round_DN_pt.
apply Rgt_not_eq.
apply bpow_gt_0.
@@ -487,7 +482,7 @@ Theorem round_NE_opp :
Proof.
intros x.
unfold round. simpl.
-rewrite scaled_mantissa_opp, canonic_exp_opp.
+rewrite scaled_mantissa_opp, cexp_opp.
rewrite Znearest_opp.
rewrite <- F2R_Zopp.
apply (f_equal (fun v => F2R (Float beta (-v) _))).
@@ -496,8 +491,8 @@ unfold Znearest.
case Rcompare ; trivial.
apply (f_equal (fun (b : bool) => if b then Zceil m else Zfloor m)).
rewrite Bool.negb_involutive.
-rewrite Zeven_opp.
-rewrite Zeven_plus.
+rewrite Z.even_opp.
+rewrite Z.even_add.
now rewrite eqb_sym.
Qed.
@@ -526,7 +521,7 @@ Theorem round_NE_pt :
Proof with auto with typeclass_instances.
intros x.
destruct (total_order_T x 0) as [[Hx|Hx]|Hx].
-apply Rnd_NG_pt_sym.
+apply Rnd_NG_pt_opp_inv.
apply generic_format_opp.
unfold NE_prop.
intros _ f ((mg,eg),(H1,(H2,H3))).
@@ -534,9 +529,9 @@ exists (Float beta (- mg) eg).
repeat split.
rewrite H1.
now rewrite F2R_Zopp.
-now apply canonic_opp.
+now apply canonical_opp.
simpl.
-now rewrite Zeven_opp.
+now rewrite Z.even_opp.
rewrite <- round_NE_opp.
apply round_NE_pt_pos.
now apply Ropp_0_gt_lt_contravar.
diff --git a/flocq/Core/Fcore_rnd.v b/flocq/Core/Round_pred.v
index e5091684..428a4bac 100644
--- a/flocq/Core/Fcore_rnd.v
+++ b/flocq/Core/Round_pred.v
@@ -2,9 +2,9 @@
This file is part of the Flocq formalization of floating-point
arithmetic in Coq: http://flocq.gforge.inria.fr/
-Copyright (C) 2010-2013 Sylvie Boldo
+Copyright (C) 2009-2018 Sylvie Boldo
#<br />#
-Copyright (C) 2010-2013 Guillaume Melquiond
+Copyright (C) 2009-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
@@ -18,13 +18,30 @@ COPYING file for more details.
*)
(** * Roundings: properties and/or functions *)
-Require Import Fcore_Raux.
-Require Import Fcore_defs.
+Require Import Raux Defs.
Section RND_prop.
Open Scope R_scope.
+Definition Rnd_DN (F : R -> Prop) (rnd : R -> R) :=
+ forall x : R, Rnd_DN_pt F x (rnd x).
+
+Definition Rnd_UP (F : R -> Prop) (rnd : R -> R) :=
+ forall x : R, Rnd_UP_pt F x (rnd x).
+
+Definition Rnd_ZR (F : R -> Prop) (rnd : R -> R) :=
+ forall x : R, Rnd_ZR_pt F x (rnd x).
+
+Definition Rnd_N (F : R -> Prop) (rnd : R -> R) :=
+ forall x : R, Rnd_N_pt F x (rnd x).
+
+Definition Rnd_NG (F : R -> Prop) (P : R -> R -> Prop) (rnd : R -> R) :=
+ forall x : R, Rnd_NG_pt F P x (rnd x).
+
+Definition Rnd_NA (F : R -> Prop) (rnd : R -> R) :=
+ forall x : R, Rnd_NA_pt F x (rnd x).
+
Theorem round_val_of_pred :
forall rnd : R -> R -> Prop,
round_pred rnd ->
@@ -63,7 +80,7 @@ intros x.
now destruct round_val_of_pred as (f, H1).
Qed.
-Theorem round_unicity :
+Theorem round_unique :
forall rnd : R -> R -> Prop,
round_pred_monotone rnd ->
forall x f1 f2,
@@ -87,25 +104,25 @@ apply Hx1.
now apply Rle_trans with (2 := Hxy).
Qed.
-Theorem Rnd_DN_pt_unicity :
+Theorem Rnd_DN_pt_unique :
forall F : R -> Prop,
forall x f1 f2 : R,
Rnd_DN_pt F x f1 -> Rnd_DN_pt F x f2 ->
f1 = f2.
Proof.
intros F.
-apply round_unicity.
+apply round_unique.
apply Rnd_DN_pt_monotone.
Qed.
-Theorem Rnd_DN_unicity :
+Theorem Rnd_DN_unique :
forall F : R -> Prop,
forall rnd1 rnd2 : R -> R,
Rnd_DN F rnd1 -> Rnd_DN F rnd2 ->
forall x, rnd1 x = rnd2 x.
Proof.
intros F rnd1 rnd2 H1 H2 x.
-now eapply Rnd_DN_pt_unicity.
+now eapply Rnd_DN_pt_unique.
Qed.
Theorem Rnd_UP_pt_monotone :
@@ -118,28 +135,28 @@ apply Hy1.
now apply Rle_trans with (1 := Hxy).
Qed.
-Theorem Rnd_UP_pt_unicity :
+Theorem Rnd_UP_pt_unique :
forall F : R -> Prop,
forall x f1 f2 : R,
Rnd_UP_pt F x f1 -> Rnd_UP_pt F x f2 ->
f1 = f2.
Proof.
intros F.
-apply round_unicity.
+apply round_unique.
apply Rnd_UP_pt_monotone.
Qed.
-Theorem Rnd_UP_unicity :
+Theorem Rnd_UP_unique :
forall F : R -> Prop,
forall rnd1 rnd2 : R -> R,
Rnd_UP F rnd1 -> Rnd_UP F rnd2 ->
forall x, rnd1 x = rnd2 x.
Proof.
intros F rnd1 rnd2 H1 H2 x.
-now eapply Rnd_UP_pt_unicity.
+now eapply Rnd_UP_pt_unique.
Qed.
-Theorem Rnd_DN_UP_pt_sym :
+Theorem Rnd_UP_pt_opp :
forall F : R -> Prop,
( forall x, F x -> F (- x) ) ->
forall x f : R,
@@ -160,7 +177,7 @@ now apply HF.
now apply Ropp_le_cancel.
Qed.
-Theorem Rnd_UP_DN_pt_sym :
+Theorem Rnd_DN_pt_opp :
forall F : R -> Prop,
( forall x, F x -> F (- x) ) ->
forall x f : R,
@@ -181,7 +198,7 @@ now apply HF.
now apply Ropp_le_cancel.
Qed.
-Theorem Rnd_DN_UP_sym :
+Theorem Rnd_DN_opp :
forall F : R -> Prop,
( forall x, F x -> F (- x) ) ->
forall rnd1 rnd2 : R -> R,
@@ -191,10 +208,10 @@ Proof.
intros F HF rnd1 rnd2 H1 H2 x.
rewrite <- (Ropp_involutive (rnd1 (-x))).
apply f_equal.
-apply (Rnd_UP_unicity F (fun x => - rnd1 (-x))) ; trivial.
+apply (Rnd_UP_unique F (fun x => - rnd1 (-x))) ; trivial.
intros y.
pattern y at 1 ; rewrite <- Ropp_involutive.
-apply Rnd_DN_UP_pt_sym.
+apply Rnd_UP_pt_opp.
apply HF.
apply H1.
Qed.
@@ -303,18 +320,17 @@ apply Rle_refl.
(* . *)
destruct (Rle_or_lt 0 x).
(* positive *)
-rewrite Rabs_right.
-rewrite Rabs_right; auto with real.
+rewrite Rabs_pos_eq with (1 := H1).
+rewrite Rabs_pos_eq.
now apply (proj1 (H x)).
-apply Rle_ge.
now apply (proj1 (H x)).
(* negative *)
+apply Rlt_le in H1.
+rewrite Rabs_left1 with (1 := H1).
rewrite Rabs_left1.
-rewrite Rabs_left1 ; auto with real.
apply Ropp_le_contravar.
-apply (proj2 (H x)).
-auto with real.
-apply (proj2 (H x)) ; auto with real.
+now apply (proj2 (H x)).
+now apply (proj2 (H x)).
Qed.
Theorem Rnd_ZR_pt_monotone :
@@ -385,12 +401,12 @@ Proof.
intros F x fd fu f Hd Hu Hf.
destruct (Rnd_N_pt_DN_or_UP F x f Hf) as [H|H].
left.
-apply Rnd_DN_pt_unicity with (1 := H) (2 := Hd).
+apply Rnd_DN_pt_unique with (1 := H) (2 := Hd).
right.
-apply Rnd_UP_pt_unicity with (1 := H) (2 := Hu).
+apply Rnd_UP_pt_unique with (1 := H) (2 := Hu).
Qed.
-Theorem Rnd_N_pt_sym :
+Theorem Rnd_N_pt_opp_inv :
forall F : R -> Prop,
( forall x, F x -> F (- x) ) ->
forall x f : R,
@@ -449,7 +465,7 @@ apply Rminus_lt.
ring_simplify.
apply Rlt_minus.
apply Rmult_lt_compat_l.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
exact Hxy.
now apply Rlt_minus.
apply Rle_0_minus.
@@ -460,7 +476,7 @@ now apply Rlt_le.
now apply Rlt_minus.
Qed.
-Theorem Rnd_N_pt_unicity :
+Theorem Rnd_N_pt_unique :
forall F : R -> Prop,
forall x d u f1 f2 : R,
Rnd_DN_pt F x d ->
@@ -476,10 +492,10 @@ clear f1 f2. intros f1 f2 Hf1 Hf2 H12.
destruct (Rnd_N_pt_DN_or_UP F x f1 Hf1) as [Hd1|Hu1] ;
destruct (Rnd_N_pt_DN_or_UP F x f2 Hf2) as [Hd2|Hu2].
apply Rlt_not_eq with (1 := H12).
-now apply Rnd_DN_pt_unicity with (1 := Hd1).
+now apply Rnd_DN_pt_unique with (1 := Hd1).
apply Hdu.
-rewrite Rnd_DN_pt_unicity with (1 := Hd) (2 := Hd1).
-rewrite Rnd_UP_pt_unicity with (1 := Hu) (2 := Hu2).
+rewrite Rnd_DN_pt_unique with (1 := Hd) (2 := Hd1).
+rewrite Rnd_UP_pt_unique with (1 := Hu) (2 := Hu2).
rewrite <- (Rabs_pos_eq (x - f1)).
rewrite <- (Rabs_pos_eq (f2 - x)).
rewrite Rabs_minus_sym.
@@ -495,7 +511,7 @@ apply Rle_trans with x.
apply Hd2.
apply Hu1.
apply Rgt_not_eq with (1 := H12).
-now apply Rnd_UP_pt_unicity with (1 := Hu2).
+now apply Rnd_UP_pt_unique with (1 := Hu2).
intros Hf1 Hf2.
now apply Rle_antisym ; apply Rnot_lt_le ; refine (H _ _ _ _).
Qed.
@@ -547,7 +563,7 @@ rewrite 2!Rminus_0_r, Rabs_R0.
apply Rabs_pos.
Qed.
-Theorem Rnd_N_pt_pos :
+Theorem Rnd_N_pt_ge_0 :
forall F : R -> Prop, F 0 ->
forall x f, 0 <= x ->
Rnd_N_pt F x f ->
@@ -563,7 +579,7 @@ now rewrite Hx.
exact HF.
Qed.
-Theorem Rnd_N_pt_neg :
+Theorem Rnd_N_pt_le_0 :
forall F : R -> Prop, F 0 ->
forall x f, x <= 0 ->
Rnd_N_pt F x f ->
@@ -589,20 +605,20 @@ intros F HF0 HF x f Hxf.
unfold Rabs at 1.
destruct (Rcase_abs x) as [Hx|Hx].
rewrite Rabs_left1.
-apply Rnd_N_pt_sym.
+apply Rnd_N_pt_opp_inv.
exact HF.
now rewrite 2!Ropp_involutive.
-apply Rnd_N_pt_neg with (3 := Hxf).
+apply Rnd_N_pt_le_0 with (3 := Hxf).
exact HF0.
now apply Rlt_le.
rewrite Rabs_pos_eq.
exact Hxf.
-apply Rnd_N_pt_pos with (3 := Hxf).
+apply Rnd_N_pt_ge_0 with (3 := Hxf).
exact HF0.
now apply Rge_le.
Qed.
-Theorem Rnd_DN_UP_pt_N :
+Theorem Rnd_N_pt_DN_UP :
forall F : R -> Prop,
forall x d u f : R,
F f ->
@@ -635,7 +651,7 @@ apply Rle_trans with (2 := Hgu).
apply Hxu.
Qed.
-Theorem Rnd_DN_pt_N :
+Theorem Rnd_N_pt_DN :
forall F : R -> Prop,
forall x d u : R,
Rnd_DN_pt F x d ->
@@ -649,14 +665,14 @@ rewrite Rabs_minus_sym.
apply Rabs_pos_eq.
apply Rle_0_minus.
apply Hd.
-apply Rnd_DN_UP_pt_N with (2 := Hd) (3 := Hu).
+apply Rnd_N_pt_DN_UP with (2 := Hd) (3 := Hu).
apply Hd.
rewrite Hdx.
apply Rle_refl.
now rewrite Hdx.
Qed.
-Theorem Rnd_UP_pt_N :
+Theorem Rnd_N_pt_UP :
forall F : R -> Prop,
forall x d u : R,
Rnd_DN_pt F x d ->
@@ -669,22 +685,22 @@ assert (Hux: (Rabs (u - x) = u - x)%R).
apply Rabs_pos_eq.
apply Rle_0_minus.
apply Hu.
-apply Rnd_DN_UP_pt_N with (2 := Hd) (3 := Hu).
+apply Rnd_N_pt_DN_UP with (2 := Hd) (3 := Hu).
apply Hu.
now rewrite Hux.
rewrite Hux.
apply Rle_refl.
Qed.
-Definition Rnd_NG_pt_unicity_prop F P :=
+Definition Rnd_NG_pt_unique_prop F P :=
forall x d u,
Rnd_DN_pt F x d -> Rnd_N_pt F x d ->
Rnd_UP_pt F x u -> Rnd_N_pt F x u ->
P x d -> P x u -> d = u.
-Theorem Rnd_NG_pt_unicity :
+Theorem Rnd_NG_pt_unique :
forall (F : R -> Prop) (P : R -> R -> Prop),
- Rnd_NG_pt_unicity_prop F P ->
+ Rnd_NG_pt_unique_prop F P ->
forall x f1 f2 : R,
Rnd_NG_pt F P x f1 -> Rnd_NG_pt F P x f2 ->
f1 = f2.
@@ -694,11 +710,11 @@ destruct H1b as [H1b|H1b].
destruct H2b as [H2b|H2b].
destruct (Rnd_N_pt_DN_or_UP _ _ _ H1a) as [H1c|H1c] ;
destruct (Rnd_N_pt_DN_or_UP _ _ _ H2a) as [H2c|H2c].
-eapply Rnd_DN_pt_unicity ; eassumption.
+eapply Rnd_DN_pt_unique ; eassumption.
now apply (HP x f1 f2).
apply sym_eq.
now apply (HP x f2 f1 H2c H2a H1c H1a).
-eapply Rnd_UP_pt_unicity ; eassumption.
+eapply Rnd_UP_pt_unique ; eassumption.
now apply H2b.
apply sym_eq.
now apply H1b.
@@ -706,14 +722,14 @@ Qed.
Theorem Rnd_NG_pt_monotone :
forall (F : R -> Prop) (P : R -> R -> Prop),
- Rnd_NG_pt_unicity_prop F P ->
+ Rnd_NG_pt_unique_prop F P ->
round_pred_monotone (Rnd_NG_pt F P).
Proof.
intros F P HP x y f g (Hf,Hx) (Hg,Hy) [Hxy|Hxy].
now apply Rnd_N_pt_monotone with F x y.
apply Req_le.
rewrite <- Hxy in Hg, Hy.
-eapply Rnd_NG_pt_unicity ; try split ; eassumption.
+eapply Rnd_NG_pt_unique ; try split ; eassumption.
Qed.
Theorem Rnd_NG_pt_refl :
@@ -728,7 +744,7 @@ intros f2 Hf2.
now apply Rnd_N_pt_idempotent with F.
Qed.
-Theorem Rnd_NG_pt_sym :
+Theorem Rnd_NG_pt_opp_inv :
forall (F : R -> Prop) (P : R -> R -> Prop),
( forall x, F x -> F (-x) ) ->
( forall x f, P x f -> P (-x) (-f) ) ->
@@ -737,7 +753,7 @@ Theorem Rnd_NG_pt_sym :
Proof.
intros F P HF HP x f (H1,H2).
split.
-now apply Rnd_N_pt_sym.
+now apply Rnd_N_pt_opp_inv.
destruct H2 as [H2|H2].
left.
rewrite <- (Ropp_involutive x), <- (Ropp_involutive f).
@@ -748,20 +764,20 @@ rewrite <- (Ropp_involutive f).
rewrite <- H2 with (-f2).
apply sym_eq.
apply Ropp_involutive.
-apply Rnd_N_pt_sym.
+apply Rnd_N_pt_opp_inv.
exact HF.
now rewrite 2!Ropp_involutive.
Qed.
-Theorem Rnd_NG_unicity :
+Theorem Rnd_NG_unique :
forall (F : R -> Prop) (P : R -> R -> Prop),
- Rnd_NG_pt_unicity_prop F P ->
+ Rnd_NG_pt_unique_prop F P ->
forall rnd1 rnd2 : R -> R,
Rnd_NG F P rnd1 -> Rnd_NG F P rnd2 ->
forall x, rnd1 x = rnd2 x.
Proof.
intros F P HP rnd1 rnd2 H1 H2 x.
-now apply Rnd_NG_pt_unicity with F P x.
+now apply Rnd_NG_pt_unique with F P x.
Qed.
Theorem Rnd_NA_NG_pt :
@@ -775,7 +791,7 @@ destruct (Rle_or_lt 0 x) as [Hx|Hx].
(* *)
split ; intros (H1, H2).
(* . *)
-assert (Hf := Rnd_N_pt_pos F HF x f Hx H1).
+assert (Hf := Rnd_N_pt_ge_0 F HF x f Hx H1).
split.
exact H1.
destruct (Rnd_N_pt_DN_or_UP _ _ _ H1) as [H3|H3].
@@ -784,12 +800,12 @@ right.
intros f2 Hxf2.
specialize (H2 _ Hxf2).
destruct (Rnd_N_pt_DN_or_UP _ _ _ Hxf2) as [H4|H4].
-eapply Rnd_DN_pt_unicity ; eassumption.
+eapply Rnd_DN_pt_unique ; eassumption.
apply Rle_antisym.
rewrite Rabs_pos_eq with (1 := Hf) in H2.
rewrite Rabs_pos_eq in H2.
exact H2.
-now apply Rnd_N_pt_pos with F x.
+now apply Rnd_N_pt_ge_0 with F x.
apply Rle_trans with x.
apply H3.
apply H4.
@@ -803,8 +819,8 @@ split.
exact H1.
intros f2 Hxf2.
destruct H2 as [H2|H2].
-assert (Hf := Rnd_N_pt_pos F HF x f Hx H1).
-assert (Hf2 := Rnd_N_pt_pos F HF x f2 Hx Hxf2).
+assert (Hf := Rnd_N_pt_ge_0 F HF x f Hx H1).
+assert (Hf2 := Rnd_N_pt_ge_0 F HF x f2 Hx Hxf2).
rewrite 2!Rabs_pos_eq ; trivial.
rewrite 2!Rabs_pos_eq in H2 ; trivial.
destruct (Rnd_N_pt_DN_or_UP _ _ _ Hxf2) as [H3|H3].
@@ -820,7 +836,7 @@ assert (Hx' := Rlt_le _ _ Hx).
clear Hx. rename Hx' into Hx.
split ; intros (H1, H2).
(* . *)
-assert (Hf := Rnd_N_pt_neg F HF x f Hx H1).
+assert (Hf := Rnd_N_pt_le_0 F HF x f Hx H1).
split.
exact H1.
destruct (Rnd_N_pt_DN_or_UP _ _ _ H1) as [H3|H3].
@@ -842,15 +858,15 @@ apply H3.
rewrite Rabs_left1 with (1 := Hf) in H2.
rewrite Rabs_left1 in H2.
now apply Ropp_le_cancel.
-now apply Rnd_N_pt_neg with F x.
-eapply Rnd_UP_pt_unicity ; eassumption.
+now apply Rnd_N_pt_le_0 with F x.
+eapply Rnd_UP_pt_unique ; eassumption.
(* . *)
split.
exact H1.
intros f2 Hxf2.
destruct H2 as [H2|H2].
-assert (Hf := Rnd_N_pt_neg F HF x f Hx H1).
-assert (Hf2 := Rnd_N_pt_neg F HF x f2 Hx Hxf2).
+assert (Hf := Rnd_N_pt_le_0 F HF x f Hx H1).
+assert (Hf2 := Rnd_N_pt_le_0 F HF x f2 Hx Hxf2).
rewrite 2!Rabs_left1 ; trivial.
rewrite 2!Rabs_left1 in H2 ; trivial.
apply Ropp_le_contravar.
@@ -865,10 +881,10 @@ rewrite (H2 _ Hxf2).
apply Rle_refl.
Qed.
-Theorem Rnd_NA_pt_unicity_prop :
+Lemma Rnd_NA_pt_unique_prop :
forall F : R -> Prop,
F 0 ->
- Rnd_NG_pt_unicity_prop F (fun a b => (Rabs a <= Rabs b)%R).
+ Rnd_NG_pt_unique_prop F (fun a b => (Rabs a <= Rabs b)%R).
Proof.
intros F HF x d u Hxd1 Hxd2 Hxu1 Hxu2 Hd Hu.
apply Rle_antisym.
@@ -892,7 +908,7 @@ apply HF.
now apply Rlt_le.
Qed.
-Theorem Rnd_NA_pt_unicity :
+Theorem Rnd_NA_pt_unique :
forall F : R -> Prop,
F 0 ->
forall x f1 f2 : R,
@@ -900,12 +916,12 @@ Theorem Rnd_NA_pt_unicity :
f1 = f2.
Proof.
intros F HF x f1 f2 H1 H2.
-apply (Rnd_NG_pt_unicity F _ (Rnd_NA_pt_unicity_prop F HF) x).
+apply (Rnd_NG_pt_unique F _ (Rnd_NA_pt_unique_prop F HF) x).
now apply -> Rnd_NA_NG_pt.
now apply -> Rnd_NA_NG_pt.
Qed.
-Theorem Rnd_NA_N_pt :
+Theorem Rnd_NA_pt_N :
forall F : R -> Prop,
F 0 ->
forall x f : R,
@@ -936,29 +952,29 @@ destruct (Rle_lt_dec 0 x) as [Hx|Hx].
(* . *)
revert Hxf.
rewrite Rabs_pos_eq with (1 := Hx).
-rewrite 2!Rabs_pos_eq ; try ( apply (Rnd_N_pt_pos F HF x) ; assumption ).
+rewrite 2!Rabs_pos_eq ; try ( apply (Rnd_N_pt_ge_0 F HF x) ; assumption ).
intros Hxf.
rewrite H0.
apply Rplus_le_reg_r with f.
ring_simplify.
apply Rmult_le_compat_l with (2 := Hxf).
-now apply (Z2R_le 0 2).
+now apply IZR_le.
(* . *)
revert Hxf.
apply Rlt_le in Hx.
rewrite Rabs_left1 with (1 := Hx).
-rewrite 2!Rabs_left1 ; try ( apply (Rnd_N_pt_neg F HF x) ; assumption ).
+rewrite 2!Rabs_left1 ; try ( apply (Rnd_N_pt_le_0 F HF x) ; assumption ).
intros Hxf.
rewrite H0.
apply Ropp_le_contravar.
apply Rplus_le_reg_r with f.
ring_simplify.
apply Rmult_le_compat_l.
-now apply (Z2R_le 0 2).
+now apply IZR_le.
now apply Ropp_le_cancel.
Qed.
-Theorem Rnd_NA_unicity :
+Theorem Rnd_NA_unique :
forall (F : R -> Prop),
F 0 ->
forall rnd1 rnd2 : R -> R,
@@ -966,7 +982,7 @@ Theorem Rnd_NA_unicity :
forall x, rnd1 x = rnd2 x.
Proof.
intros F HF rnd1 rnd2 H1 H2 x.
-now apply Rnd_NA_pt_unicity with F x.
+now apply Rnd_NA_pt_unique with F x.
Qed.
Theorem Rnd_NA_pt_monotone :
@@ -975,7 +991,7 @@ Theorem Rnd_NA_pt_monotone :
round_pred_monotone (Rnd_NA_pt F).
Proof.
intros F HF x y f g Hxf Hyg Hxy.
-apply (Rnd_NG_pt_monotone F _ (Rnd_NA_pt_unicity_prop F HF) x y).
+apply (Rnd_NG_pt_monotone F _ (Rnd_NA_pt_unique_prop F HF) x y).
now apply -> Rnd_NA_NG_pt.
now apply -> Rnd_NA_NG_pt.
exact Hxy.
@@ -1165,7 +1181,7 @@ intros x.
destruct (proj1 (satisfies_any_imp_DN F Hany) (-x)) as (f, Hf).
exists (-f).
rewrite <- (Ropp_involutive x).
-apply Rnd_DN_UP_pt_sym.
+apply Rnd_UP_pt_opp.
apply Hany.
exact Hf.
apply Rnd_UP_pt_monotone.
diff --git a/flocq/Core/Fcore_ulp.v b/flocq/Core/Ulp.v
index 4fdd319e..4f4a5674 100644
--- a/flocq/Core/Fcore_ulp.v
+++ b/flocq/Core/Ulp.v
@@ -2,9 +2,9 @@
This file is part of the Flocq formalization of floating-point
arithmetic in Coq: http://flocq.gforge.inria.fr/
-Copyright (C) 2010-2013 Sylvie Boldo
+Copyright (C) 2009-2018 Sylvie Boldo
#<br />#
-Copyright (C) 2010-2013 Guillaume Melquiond
+Copyright (C) 2009-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
@@ -19,11 +19,7 @@ COPYING file for more details.
(** * Unit in the Last Place: our definition using fexp and its properties, successor and predecessor *)
Require Import Reals Psatz.
-Require Import Fcore_Raux.
-Require Import Fcore_defs.
-Require Import Fcore_rnd.
-Require Import Fcore_generic_fmt.
-Require Import Fcore_float_prop.
+Require Import Raux Defs Round_pred Generic_fmt Float_prop.
Section Fcore_ulp.
@@ -97,10 +93,12 @@ Definition ulp x := match Req_bool x 0 with
| Some n => bpow (fexp n)
| None => 0%R
end
- | false => bpow (canonic_exp beta fexp x)
+ | false => bpow (cexp beta fexp x)
end.
-Lemma ulp_neq_0 : forall x:R, (x <> 0)%R -> ulp x = bpow (canonic_exp beta fexp x).
+Lemma ulp_neq_0 :
+ forall x, x <> 0%R ->
+ ulp x = bpow (cexp beta fexp x).
Proof.
intros x Hx.
unfold ulp; case (Req_bool_spec x); trivial.
@@ -118,7 +116,7 @@ case Req_bool_spec; intros H1.
rewrite Req_bool_true; trivial.
rewrite <- (Ropp_involutive x), H1; ring.
rewrite Req_bool_false.
-now rewrite canonic_exp_opp.
+now rewrite cexp_opp.
intros H2; apply H1; rewrite H2; ring.
Qed.
@@ -130,7 +128,7 @@ unfold ulp; case (Req_bool_spec x 0); intros H1.
rewrite Req_bool_true; trivial.
now rewrite H1, Rabs_R0.
rewrite Req_bool_false.
-now rewrite canonic_exp_abs.
+now rewrite cexp_abs.
now apply Rabs_no_R0.
Qed.
@@ -159,9 +157,8 @@ rewrite ulp_neq_0.
unfold F2R; simpl.
apply Rmult_le_compat_r.
apply bpow_ge_0.
-apply (Z2R_le (Zsucc 0)).
-apply Zlt_le_succ.
-apply F2R_gt_0_reg with beta (canonic_exp beta fexp x).
+apply IZR_le, (Zlt_le_succ 0).
+apply gt_0_F2R with beta (cexp beta fexp x).
now rewrite <- Fx.
Qed.
@@ -178,8 +175,6 @@ now apply Rabs_pos_lt.
now apply generic_format_abs.
Qed.
-
-(* was ulp_DN_UP *)
Theorem round_UP_DN_ulp :
forall x, ~ F x ->
round beta fexp Zceil x = (round beta fexp Zfloor x + ulp x)%R.
@@ -189,13 +184,13 @@ rewrite ulp_neq_0.
unfold round. simpl.
unfold F2R. simpl.
rewrite Zceil_floor_neq.
-rewrite Z2R_plus. simpl.
+rewrite plus_IZR. simpl.
ring.
intros H.
apply Fx.
unfold generic_format, F2R. simpl.
rewrite <- H.
-rewrite Ztrunc_Z2R.
+rewrite Ztrunc_IZR.
rewrite H.
now rewrite scaled_mantissa_mult_bpow.
intros V; apply Fx.
@@ -210,7 +205,7 @@ Proof.
intros e.
rewrite ulp_neq_0.
apply f_equal.
-apply canonic_exp_fexp.
+apply cexp_fexp.
rewrite Rabs_pos_eq.
split.
ring_simplify (e + 1 - 1)%Z.
@@ -222,7 +217,7 @@ apply Rgt_not_eq, Rlt_gt, bpow_gt_0.
Qed.
-Lemma generic_format_ulp_0:
+Lemma generic_format_ulp_0 :
F (ulp 0).
Proof.
unfold ulp.
@@ -234,8 +229,9 @@ apply generic_format_bpow.
now apply valid_exp.
Qed.
-Lemma generic_format_bpow_ge_ulp_0: forall e,
- (ulp 0 <= bpow e)%R -> F (bpow e).
+Lemma generic_format_bpow_ge_ulp_0 :
+ forall e, (ulp 0 <= bpow e)%R ->
+ F (bpow e).
Proof.
intros e; unfold ulp.
rewrite Req_bool_true; trivial.
@@ -248,7 +244,7 @@ apply generic_format_bpow.
case (Zle_or_lt (e+1) (fexp (e+1))); intros H4.
absurd (e+1 <= e)%Z.
omega.
-apply Zle_trans with (1:=H4).
+apply Z.le_trans with (1:=H4).
replace (fexp (e+1)) with (fexp n).
now apply le_bpow with beta.
now apply fexp_negligible_exp_eq.
@@ -258,33 +254,36 @@ Qed.
(** The three following properties are equivalent:
[Exp_not_FTZ] ; forall x, F (ulp x) ; forall x, ulp 0 <= ulp x *)
-Lemma generic_format_ulp: Exp_not_FTZ fexp ->
- forall x, F (ulp x).
+Lemma generic_format_ulp :
+ Exp_not_FTZ fexp ->
+ forall x, F (ulp x).
Proof.
unfold Exp_not_FTZ; intros H x.
case (Req_dec x 0); intros Hx.
rewrite Hx; apply generic_format_ulp_0.
rewrite (ulp_neq_0 _ Hx).
-apply generic_format_bpow; unfold canonic_exp.
+apply generic_format_bpow.
apply H.
Qed.
-Lemma not_FTZ_generic_format_ulp:
- (forall x, F (ulp x)) -> Exp_not_FTZ fexp.
+Lemma not_FTZ_generic_format_ulp :
+ (forall x, F (ulp x)) ->
+ Exp_not_FTZ fexp.
Proof.
intros H e.
specialize (H (bpow (e-1))).
rewrite ulp_neq_0 in H.
2: apply Rgt_not_eq, bpow_gt_0.
-unfold canonic_exp in H.
-rewrite ln_beta_bpow in H.
-apply generic_format_bpow_inv' in H...
+unfold cexp in H.
+rewrite mag_bpow in H.
+apply generic_format_bpow_inv' in H.
now replace (e-1+1)%Z with e in H by ring.
Qed.
-Lemma ulp_ge_ulp_0: Exp_not_FTZ fexp ->
- forall x, (ulp 0 <= ulp x)%R.
+Lemma ulp_ge_ulp_0 :
+ Exp_not_FTZ fexp ->
+ forall x, (ulp 0 <= ulp x)%R.
Proof.
unfold Exp_not_FTZ; intros H x.
case (Req_dec x 0); intros Hx.
@@ -295,20 +294,21 @@ case negligible_exp_spec'.
intros (H1,H2); rewrite H1; apply ulp_ge_0.
intros (n,(H1,H2)); rewrite H1.
rewrite ulp_neq_0; trivial.
-apply bpow_le; unfold canonic_exp.
-generalize (ln_beta beta x); intros l.
+apply bpow_le; unfold cexp.
+generalize (mag beta x); intros l.
case (Zle_or_lt l (fexp l)); intros Hl.
-rewrite (fexp_negligible_exp_eq n l); trivial; apply Zle_refl.
+rewrite (fexp_negligible_exp_eq n l); trivial; apply Z.le_refl.
case (Zle_or_lt (fexp n) (fexp l)); trivial; intros K.
absurd (fexp n <= fexp l)%Z.
omega.
-apply Zle_trans with (2:= H _).
+apply Z.le_trans with (2:= H _).
apply Zeq_le, sym_eq, valid_exp; trivial.
omega.
Qed.
Lemma not_FTZ_ulp_ge_ulp_0:
- (forall x, (ulp 0 <= ulp x)%R) -> Exp_not_FTZ fexp.
+ (forall x, (ulp 0 <= ulp x)%R) ->
+ Exp_not_FTZ fexp.
Proof.
intros H e.
apply generic_format_bpow_inv' with beta.
@@ -318,9 +318,7 @@ rewrite <- ulp_bpow.
apply H.
Qed.
-
-
-Theorem ulp_le_pos :
+Lemma ulp_le_pos :
forall { Hm : Monotone_exp fexp },
forall x y: R,
(0 <= x)%R -> (x <= y)%R ->
@@ -332,7 +330,7 @@ rewrite ulp_neq_0.
rewrite ulp_neq_0.
apply bpow_le.
apply Hm.
-now apply ln_beta_le.
+now apply mag_le.
apply Rgt_not_eq, Rlt_gt.
now apply Rlt_le_trans with (1:=Hx).
now apply Rgt_not_eq.
@@ -341,7 +339,6 @@ apply ulp_ge_ulp_0.
apply monotone_exp_not_FTZ...
Qed.
-
Theorem ulp_le :
forall { Hm : Monotone_exp fexp },
forall x y: R,
@@ -355,26 +352,49 @@ apply ulp_le_pos; trivial.
apply Rabs_pos.
Qed.
+(** Properties when there is no minimal exponent *)
+Theorem eq_0_round_0_negligible_exp :
+ negligible_exp = None -> forall rnd {Vr: Valid_rnd rnd} x,
+ round beta fexp rnd x = 0%R -> x = 0%R.
+Proof.
+intros H rnd Vr x Hx.
+case (Req_dec x 0); try easy; intros Hx2.
+absurd (Rabs (round beta fexp rnd x) = 0%R).
+2: rewrite Hx, Rabs_R0; easy.
+apply Rgt_not_eq.
+apply Rlt_le_trans with (bpow (mag beta x - 1)).
+apply bpow_gt_0.
+apply abs_round_ge_generic; try assumption.
+apply generic_format_bpow.
+case negligible_exp_spec'; [intros (K1,K2)|idtac].
+ring_simplify (mag beta x-1+1)%Z.
+specialize (K2 (mag beta x)); now auto with zarith.
+intros (n,(Hn1,Hn2)).
+rewrite Hn1 in H; discriminate.
+now apply bpow_mag_le.
+Qed.
+
(** Definition and properties of pred and succ *)
Definition pred_pos x :=
- if Req_bool x (bpow (ln_beta beta x - 1)) then
- (x - bpow (fexp (ln_beta beta x - 1)))%R
+ if Req_bool x (bpow (mag beta x - 1)) then
+ (x - bpow (fexp (mag beta x - 1)))%R
else
(x - ulp x)%R.
Definition succ x :=
- if (Rle_bool 0 x) then
- (x+ulp x)%R
- else
- (- pred_pos (-x))%R.
+ if (Rle_bool 0 x) then
+ (x+ulp x)%R
+ else
+ (- pred_pos (-x))%R.
Definition pred x := (- succ (-x))%R.
-Theorem pred_eq_pos:
- forall x, (0 <= x)%R -> (pred x = pred_pos x)%R.
+Theorem pred_eq_pos :
+ forall x, (0 <= x)%R ->
+ pred x = pred_pos x.
Proof.
intros x Hx; unfold pred, succ.
case Rle_bool_spec; intros Hx'.
@@ -389,39 +409,29 @@ rewrite Ropp_0; ring.
now rewrite 2!Ropp_involutive.
Qed.
-Theorem succ_eq_pos:
- forall x, (0 <= x)%R -> (succ x = x + ulp x)%R.
+Theorem succ_eq_pos :
+ forall x, (0 <= x)%R ->
+ succ x = (x + ulp x)%R.
Proof.
intros x Hx; unfold succ.
now rewrite Rle_bool_true.
Qed.
-Lemma pred_eq_opp_succ_opp: forall x, pred x = (- succ (-x))%R.
+Theorem succ_opp :
+ forall x, succ (-x) = (- pred x)%R.
Proof.
-reflexivity.
-Qed.
-
-Lemma succ_eq_opp_pred_opp: forall x, succ x = (- pred (-x))%R.
-Proof.
-intros x; unfold pred.
-now rewrite 2!Ropp_involutive.
-Qed.
-
-Lemma succ_opp: forall x, (succ (-x) = - pred x)%R.
-Proof.
-intros x; rewrite succ_eq_opp_pred_opp.
-now rewrite Ropp_involutive.
+intros x.
+now apply sym_eq, Ropp_involutive.
Qed.
-Lemma pred_opp: forall x, (pred (-x) = - succ x)%R.
+Theorem pred_opp :
+ forall x, pred (-x) = (- succ x)%R.
Proof.
-intros x; rewrite pred_eq_opp_succ_opp.
+intros x.
+unfold pred.
now rewrite Ropp_involutive.
Qed.
-
-
-
(** pred and succ are in the format *)
(* cannont be x <> ulp 0, due to the counter-example 1-bit FP format fexp: e -> e-1 *)
@@ -436,7 +446,7 @@ intros x e Fx Hx' Hx.
(* *)
assert (1 <= Ztrunc (scaled_mantissa beta fexp x))%Z.
assert (0 < Ztrunc (scaled_mantissa beta fexp x))%Z.
-apply F2R_gt_0_reg with beta (canonic_exp beta fexp x).
+apply gt_0_F2R with beta (cexp beta fexp x).
rewrite <- Fx.
apply Rle_lt_trans with (2:=Hx).
apply bpow_ge_0.
@@ -446,12 +456,11 @@ case (Zle_lt_or_eq _ _ H); intros Hm.
pattern x at 1 ; rewrite Fx.
rewrite ulp_neq_0.
unfold F2R. simpl.
-pattern (bpow (canonic_exp beta fexp x)) at 2 ; rewrite <- Rmult_1_l.
+pattern (bpow (cexp beta fexp x)) at 2 ; rewrite <- Rmult_1_l.
rewrite <- Rmult_minus_distr_r.
-change 1%R with (Z2R 1).
-rewrite <- Z2R_minus.
-change (bpow e <= F2R (Float beta (Ztrunc (scaled_mantissa beta fexp x) - 1) (canonic_exp beta fexp x)))%R.
-apply bpow_le_F2R_m1; trivial.
+rewrite <- minus_IZR.
+apply bpow_le_F2R_m1.
+easy.
now rewrite <- Fx.
apply Rgt_not_eq, Rlt_gt.
apply Rlt_trans with (2:=Hx), bpow_gt_0.
@@ -476,27 +485,23 @@ intros x e Zx Fx Hx.
pattern x at 1 ; rewrite Fx.
rewrite ulp_neq_0.
unfold F2R. simpl.
-pattern (bpow (canonic_exp beta fexp x)) at 2 ; rewrite <- Rmult_1_l.
+pattern (bpow (cexp beta fexp x)) at 2 ; rewrite <- Rmult_1_l.
rewrite <- Rmult_plus_distr_r.
-change 1%R with (Z2R 1).
-rewrite <- Z2R_plus.
-change (F2R (Float beta (Ztrunc (scaled_mantissa beta fexp x) + 1) (canonic_exp beta fexp x)) <= bpow e)%R.
+rewrite <- plus_IZR.
apply F2R_p1_le_bpow.
-apply F2R_gt_0_reg with beta (canonic_exp beta fexp x).
+apply gt_0_F2R with beta (cexp beta fexp x).
now rewrite <- Fx.
now rewrite <- Fx.
now apply Rgt_not_eq.
Qed.
-
-
Lemma generic_format_pred_aux1:
forall x, (0 < x)%R -> F x ->
- x <> bpow (ln_beta beta x - 1) ->
+ x <> bpow (mag beta x - 1) ->
F (x - ulp x).
Proof.
intros x Zx Fx Hx.
-destruct (ln_beta beta x) as (ex, Ex).
+destruct (mag beta x) as (ex, Ex).
simpl in Hx.
specialize (Ex (Rgt_not_eq _ _ Zx)).
assert (Ex' : (bpow (ex - 1) < x < bpow ex)%R).
@@ -504,20 +509,20 @@ rewrite Rabs_pos_eq in Ex.
destruct Ex as (H,H'); destruct H; split; trivial.
contradict Hx; easy.
now apply Rlt_le.
-unfold generic_format, scaled_mantissa, canonic_exp.
-rewrite ln_beta_unique with beta (x - ulp x)%R ex.
+unfold generic_format, scaled_mantissa, cexp.
+rewrite mag_unique with beta (x - ulp x)%R ex.
pattern x at 1 3 ; rewrite Fx.
rewrite ulp_neq_0.
unfold scaled_mantissa.
-rewrite canonic_exp_fexp with (1 := Ex).
+rewrite cexp_fexp with (1 := Ex).
unfold F2R. simpl.
rewrite Rmult_minus_distr_r.
rewrite Rmult_assoc.
rewrite <- bpow_plus, Zplus_opp_r, Rmult_1_r.
-change (bpow 0) with (Z2R 1).
-rewrite <- Z2R_minus.
-rewrite Ztrunc_Z2R.
-rewrite Z2R_minus.
+change (bpow 0) with 1%R.
+rewrite <- minus_IZR.
+rewrite Ztrunc_IZR.
+rewrite minus_IZR.
rewrite Rmult_minus_distr_r.
now rewrite Rmult_1_l.
now apply Rgt_not_eq.
@@ -526,7 +531,7 @@ split.
apply id_m_ulp_ge_bpow; trivial.
rewrite ulp_neq_0.
intro H.
-assert (ex-1 < canonic_exp beta fexp x < ex)%Z.
+assert (ex-1 < cexp beta fexp x < ex)%Z.
split ; apply (lt_bpow beta) ; rewrite <- H ; easy.
clear -H0. omega.
now apply Rgt_not_eq.
@@ -541,13 +546,12 @@ apply Rle_0_minus.
pattern x at 2; rewrite Fx.
rewrite ulp_neq_0.
unfold F2R; simpl.
-pattern (bpow (canonic_exp beta fexp x)) at 1; rewrite <- Rmult_1_l.
+pattern (bpow (cexp beta fexp x)) at 1; rewrite <- Rmult_1_l.
apply Rmult_le_compat_r.
apply bpow_ge_0.
-replace 1%R with (Z2R 1) by reflexivity.
-apply Z2R_le.
+apply IZR_le.
assert (0 < Ztrunc (scaled_mantissa beta fexp x))%Z.
-apply F2R_gt_0_reg with beta (canonic_exp beta fexp x).
+apply gt_0_F2R with beta (cexp beta fexp x).
rewrite <- Fx.
apply Rle_lt_trans with (2:=proj1 Ex').
apply bpow_ge_0.
@@ -557,8 +561,8 @@ Qed.
Lemma generic_format_pred_aux2 :
forall x, (0 < x)%R -> F x ->
- let e := ln_beta_val beta x (ln_beta beta x) in
- x = bpow (e - 1) ->
+ let e := mag_val beta x (mag beta x) in
+ x = bpow (e - 1) ->
F (x - bpow (fexp (e - 1))).
Proof.
intros x Zx Fx e Hx.
@@ -571,7 +575,7 @@ case (Zle_lt_or_eq _ _ He); clear He; intros He.
assert (f = F2R (Float beta (Zpower beta (e-1-(fexp (e-1))) -1) (fexp (e-1))))%R.
unfold f; rewrite Hx.
unfold F2R; simpl.
-rewrite Z2R_minus, Z2R_Zpower.
+rewrite minus_IZR, IZR_Zpower.
rewrite Rmult_minus_distr_r, Rmult_1_l.
rewrite <- bpow_plus.
now replace (e - 1 - fexp (e - 1) + fexp (e - 1))%Z with (e-1)%Z by ring.
@@ -580,7 +584,7 @@ rewrite H.
apply generic_format_F2R.
intros _.
apply Zeq_le.
-apply canonic_exp_fexp.
+apply cexp_fexp.
rewrite <- H.
unfold f; rewrite Hx.
rewrite Rabs_right.
@@ -593,9 +597,8 @@ apply Rle_trans with (2*bpow (e - 2))%R;[right; ring|idtac].
apply Rle_trans with (bpow 1*bpow (e - 2))%R.
apply Rmult_le_compat_r.
apply bpow_ge_0.
-replace 2%R with (Z2R 2) by reflexivity.
-replace (bpow 1) with (Z2R beta).
-apply Z2R_le.
+replace (bpow 1) with (IZR beta).
+apply IZR_le.
apply <- Zle_is_le_bool.
now destruct beta.
simpl.
@@ -619,31 +622,30 @@ rewrite Hx, He.
ring.
Qed.
-
-Theorem generic_format_succ_aux1 :
+Lemma generic_format_succ_aux1 :
forall x, (0 < x)%R -> F x ->
F (x + ulp x).
Proof.
intros x Zx Fx.
-destruct (ln_beta beta x) as (ex, Ex).
+destruct (mag beta x) as (ex, Ex).
specialize (Ex (Rgt_not_eq _ _ Zx)).
assert (Ex' := Ex).
rewrite Rabs_pos_eq in Ex'.
destruct (id_p_ulp_le_bpow x ex) ; try easy.
-unfold generic_format, scaled_mantissa, canonic_exp.
-rewrite ln_beta_unique with beta (x + ulp x)%R ex.
+unfold generic_format, scaled_mantissa, cexp.
+rewrite mag_unique with beta (x + ulp x)%R ex.
pattern x at 1 3 ; rewrite Fx.
rewrite ulp_neq_0.
unfold scaled_mantissa.
-rewrite canonic_exp_fexp with (1 := Ex).
+rewrite cexp_fexp with (1 := Ex).
unfold F2R. simpl.
rewrite Rmult_plus_distr_r.
rewrite Rmult_assoc.
rewrite <- bpow_plus, Zplus_opp_r, Rmult_1_r.
-change (bpow 0) with (Z2R 1).
-rewrite <- Z2R_plus.
-rewrite Ztrunc_Z2R.
-rewrite Z2R_plus.
+change (bpow 0) with 1%R.
+rewrite <- plus_IZR.
+rewrite Ztrunc_IZR.
+rewrite plus_IZR.
rewrite Rmult_plus_distr_r.
now rewrite Rmult_1_l.
now apply Rgt_not_eq.
@@ -667,7 +669,7 @@ replace (Ztrunc (scaled_mantissa beta fexp x)) with Z0.
rewrite F2R_0.
apply Rle_refl.
unfold scaled_mantissa.
-rewrite canonic_exp_fexp with (1 := Ex).
+rewrite cexp_fexp with (1 := Ex).
destruct (mantissa_small_pos beta fexp x ex) ; trivial.
rewrite Ztrunc_floor.
apply sym_eq.
@@ -679,7 +681,7 @@ now apply Rlt_le.
now apply Rlt_le.
Qed.
-Theorem generic_format_pred_pos :
+Lemma generic_format_pred_pos :
forall x, F x -> (0 < x)%R ->
F (pred_pos x).
Proof.
@@ -689,7 +691,6 @@ now apply generic_format_pred_aux2.
now apply generic_format_pred_aux1.
Qed.
-
Theorem generic_format_succ :
forall x, F x ->
F (succ x).
@@ -717,9 +718,7 @@ apply generic_format_succ.
now apply generic_format_opp.
Qed.
-
-
-Theorem pred_pos_lt_id :
+Lemma pred_pos_lt_id :
forall x, (x <> 0)%R ->
(pred_pos x < x)%R.
Proof.
@@ -754,7 +753,7 @@ apply bpow_gt_0.
pattern x at 1; rewrite <- (Ropp_involutive x).
apply Ropp_lt_contravar.
apply pred_pos_lt_id.
-now auto with real.
+auto with real.
Qed.
@@ -766,7 +765,7 @@ intros x Zx; unfold pred.
pattern x at 2; rewrite <- (Ropp_involutive x).
apply Ropp_lt_contravar.
apply succ_gt_id.
-now auto with real.
+auto with real.
Qed.
Theorem succ_ge_id :
@@ -781,7 +780,7 @@ Qed.
Theorem pred_le_id :
- forall x, (pred x <= x)%R.
+ forall x, (pred x <= x)%R.
Proof.
intros x; unfold pred.
pattern x at 2; rewrite <- (Ropp_involutive x).
@@ -790,7 +789,7 @@ apply succ_ge_id.
Qed.
-Theorem pred_pos_ge_0 :
+Lemma pred_pos_ge_0 :
forall x,
(0 < x)%R -> F x -> (0 <= pred_pos x)%R.
Proof.
@@ -801,8 +800,8 @@ case Req_bool_spec; intros H.
apply Rle_0_minus.
rewrite H.
apply bpow_le.
-destruct (ln_beta beta x) as (ex,Ex) ; simpl.
-rewrite ln_beta_bpow.
+destruct (mag beta x) as (ex,Ex) ; simpl.
+rewrite mag_bpow.
ring_simplify (ex - 1 + 1 - 1)%Z.
apply generic_format_bpow_inv with beta; trivial.
simpl in H.
@@ -824,36 +823,35 @@ Qed.
Lemma pred_pos_plus_ulp_aux1 :
forall x, (0 < x)%R -> F x ->
- x <> bpow (ln_beta beta x - 1) ->
+ x <> bpow (mag beta x - 1) ->
((x - ulp x) + ulp (x-ulp x) = x)%R.
Proof.
intros x Zx Fx Hx.
replace (ulp (x - ulp x)) with (ulp x).
ring.
-assert (H:(x <> 0)%R) by auto with real.
-assert (H':(x <> bpow (canonic_exp beta fexp x))%R).
-unfold canonic_exp; intros M.
-case_eq (ln_beta beta x); intros ex Hex T.
-assert (Lex:(ln_beta_val beta x (ln_beta beta x) = ex)%Z).
+assert (H : x <> 0%R) by now apply Rgt_not_eq.
+assert (H' : x <> bpow (cexp beta fexp x)).
+unfold cexp ; intros M.
+case_eq (mag beta x); intros ex Hex T.
+assert (Lex:(mag_val beta x (mag beta x) = ex)%Z).
rewrite T; reflexivity.
rewrite Lex in *.
clear T; simpl in *; specialize (Hex H).
-rewrite Rabs_right in Hex.
-2: apply Rle_ge; apply Rlt_le; easy.
-assert (ex-1 < fexp ex < ex)%Z.
-split ; apply (lt_bpow beta); rewrite <- M;[idtac|easy].
-destruct (proj1 Hex);[trivial|idtac].
-contradict Hx; auto with real.
+rewrite Rabs_pos_eq in Hex by now apply Rlt_le.
+assert (ex - 1 < fexp ex < ex)%Z.
+ split ; apply (lt_bpow beta) ; rewrite <- M by easy.
+ lra.
+ apply Hex.
omega.
-rewrite 2!ulp_neq_0; try auto with real.
+rewrite 2!ulp_neq_0 by lra.
apply f_equal.
-unfold canonic_exp; apply f_equal.
-case_eq (ln_beta beta x); intros ex Hex T.
-assert (Lex:(ln_beta_val beta x (ln_beta beta x) = ex)%Z).
+unfold cexp ; apply f_equal.
+case_eq (mag beta x); intros ex Hex T.
+assert (Lex:(mag_val beta x (mag beta x) = ex)%Z).
rewrite T; reflexivity.
rewrite Lex in *; simpl in *; clear T.
specialize (Hex H).
-apply sym_eq, ln_beta_unique.
+apply sym_eq, mag_unique.
rewrite Rabs_right.
rewrite Rabs_right in Hex.
2: apply Rle_ge; apply Rlt_le; easy.
@@ -863,8 +861,8 @@ apply Rle_trans with (x-ulp x)%R.
apply id_m_ulp_ge_bpow; trivial.
rewrite ulp_neq_0; trivial.
rewrite ulp_neq_0; trivial.
-right; unfold canonic_exp; now rewrite Lex.
-contradict Hx; auto with real.
+right; unfold cexp; now rewrite Lex.
+lra.
apply Rle_lt_trans with (2:=proj2 Hex).
rewrite <- Rplus_0_r.
apply Rplus_le_compat_l.
@@ -874,22 +872,19 @@ apply bpow_ge_0.
apply Rle_ge.
apply Rle_0_minus.
rewrite Fx.
-unfold F2R, canonic_exp; simpl.
+unfold F2R, cexp; simpl.
rewrite Lex.
pattern (bpow (fexp ex)) at 1; rewrite <- Rmult_1_l.
apply Rmult_le_compat_r.
apply bpow_ge_0.
-replace 1%R with (Z2R (Zsucc 0)) by reflexivity.
-apply Z2R_le.
-apply Zlt_le_succ.
-apply F2R_gt_0_reg with beta (canonic_exp beta fexp x).
+apply IZR_le, (Zlt_le_succ 0).
+apply gt_0_F2R with beta (cexp beta fexp x).
now rewrite <- Fx.
Qed.
-
Lemma pred_pos_plus_ulp_aux2 :
forall x, (0 < x)%R -> F x ->
- let e := ln_beta_val beta x (ln_beta beta x) in
+ let e := mag_val beta x (mag beta x) in
x = bpow (e - 1) ->
(x - bpow (fexp (e-1)) <> 0)%R ->
((x - bpow (fexp (e-1))) + ulp (x - bpow (fexp (e-1))) = x)%R.
@@ -904,9 +899,9 @@ case (Zle_lt_or_eq _ _ He); clear He; intros He.
(* *)
rewrite ulp_neq_0; trivial.
apply f_equal.
-unfold canonic_exp; apply f_equal.
+unfold cexp ; apply f_equal.
apply sym_eq.
-apply ln_beta_unique.
+apply mag_unique.
rewrite Rabs_right.
split.
apply Rplus_le_reg_l with (bpow (fexp (e-1))).
@@ -917,9 +912,8 @@ apply Rle_trans with (2*bpow (e - 2))%R;[right; ring|idtac].
apply Rle_trans with (bpow 1*bpow (e - 2))%R.
apply Rmult_le_compat_r.
apply bpow_ge_0.
-replace 2%R with (Z2R 2) by reflexivity.
-replace (bpow 1) with (Z2R beta).
-apply Z2R_le.
+replace (bpow 1) with (IZR beta).
+apply IZR_le.
apply <- Zle_is_le_bool.
now destruct beta.
simpl.
@@ -944,7 +938,7 @@ Qed.
Lemma pred_pos_plus_ulp_aux3 :
forall x, (0 < x)%R -> F x ->
- let e := ln_beta_val beta x (ln_beta beta x) in
+ let e := mag_val beta x (mag beta x) in
x = bpow (e - 1) ->
(x - bpow (fexp (e-1)) = 0)%R ->
(ulp 0 = x)%R.
@@ -967,40 +961,44 @@ apply valid_exp; omega.
apply sym_eq, valid_exp; omega.
Qed.
-
-
-
(** The following one is false for x = 0 in FTZ *)
-Theorem pred_pos_plus_ulp :
+Lemma pred_pos_plus_ulp :
forall x, (0 < x)%R -> F x ->
(pred_pos x + ulp (pred_pos x) = x)%R.
Proof.
intros x Zx Fx.
unfold pred_pos.
case Req_bool_spec; intros H.
-case (Req_EM_T (x - bpow (fexp (ln_beta_val beta x (ln_beta beta x) -1))) 0); intros H1.
+case (Req_EM_T (x - bpow (fexp (mag_val beta x (mag beta x) -1))) 0); intros H1.
rewrite H1, Rplus_0_l.
now apply pred_pos_plus_ulp_aux3.
now apply pred_pos_plus_ulp_aux2.
now apply pred_pos_plus_ulp_aux1.
Qed.
-
-
+Theorem pred_plus_ulp :
+ forall x, (0 < x)%R -> F x ->
+ (pred x + ulp (pred x))%R = x.
+Proof.
+intros x Hx Fx.
+rewrite pred_eq_pos.
+now apply pred_pos_plus_ulp.
+now apply Rlt_le.
+Qed.
(** Rounding x + small epsilon *)
-Theorem ln_beta_plus_eps:
+Theorem mag_plus_eps :
forall x, (0 < x)%R -> F x ->
forall eps, (0 <= eps < ulp x)%R ->
- ln_beta beta (x + eps) = ln_beta beta x :> Z.
+ mag beta (x + eps) = mag beta x :> Z.
Proof.
intros x Zx Fx eps Heps.
-destruct (ln_beta beta x) as (ex, He).
+destruct (mag beta x) as (ex, He).
simpl.
specialize (He (Rgt_not_eq _ _ Zx)).
-apply ln_beta_unique.
+apply mag_unique.
rewrite Rabs_pos_eq.
rewrite Rabs_pos_eq in He.
split.
@@ -1012,13 +1010,11 @@ now apply Rplus_lt_compat_l.
pattern x at 1 ; rewrite Fx.
rewrite ulp_neq_0.
unfold F2R. simpl.
-pattern (bpow (canonic_exp beta fexp x)) at 2 ; rewrite <- Rmult_1_l.
+pattern (bpow (cexp beta fexp x)) at 2 ; rewrite <- Rmult_1_l.
rewrite <- Rmult_plus_distr_r.
-change 1%R with (Z2R 1).
-rewrite <- Z2R_plus.
-change (F2R (Float beta (Ztrunc (scaled_mantissa beta fexp x) + 1) (canonic_exp beta fexp x)) <= bpow ex)%R.
+rewrite <- plus_IZR.
apply F2R_p1_le_bpow.
-apply F2R_gt_0_reg with beta (canonic_exp beta fexp x).
+apply gt_0_F2R with beta (cexp beta fexp x).
now rewrite <- Fx.
now rewrite <- Fx.
now apply Rgt_not_eq.
@@ -1028,7 +1024,7 @@ now apply Rlt_le.
apply Heps.
Qed.
-Theorem round_DN_plus_eps_pos:
+Theorem round_DN_plus_eps_pos :
forall x, (0 <= x)%R -> F x ->
forall eps, (0 <= eps < ulp x)%R ->
round beta fexp Zfloor (x + eps) = x.
@@ -1039,8 +1035,8 @@ destruct Zx as [Zx|Zx].
pattern x at 2 ; rewrite Fx.
unfold round.
unfold scaled_mantissa. simpl.
-unfold canonic_exp at 1 2.
-rewrite ln_beta_plus_eps ; trivial.
+unfold cexp at 1 2.
+rewrite mag_plus_eps ; trivial.
apply (f_equal (fun m => F2R (Float beta m _))).
rewrite Ztrunc_floor.
apply Zfloor_imp.
@@ -1050,12 +1046,12 @@ apply Rmult_le_compat_r.
apply bpow_ge_0.
pattern x at 1 ; rewrite <- Rplus_0_r.
now apply Rplus_le_compat_l.
-apply Rlt_le_trans with ((x + ulp x) * bpow (- canonic_exp beta fexp x))%R.
+apply Rlt_le_trans with ((x + ulp x) * bpow (- cexp beta fexp x))%R.
apply Rmult_lt_compat_r.
apply bpow_gt_0.
now apply Rplus_lt_compat_l.
rewrite Rmult_plus_distr_r.
-rewrite Z2R_plus.
+rewrite plus_IZR.
apply Rplus_le_compat.
pattern x at 1 3 ; rewrite Fx.
unfold F2R. simpl.
@@ -1063,7 +1059,7 @@ rewrite Rmult_assoc.
rewrite <- bpow_plus.
rewrite Zplus_opp_r.
rewrite Rmult_1_r.
-rewrite Zfloor_Z2R.
+rewrite Zfloor_IZR.
apply Rle_refl.
rewrite ulp_neq_0.
2: now apply Rgt_not_eq.
@@ -1076,24 +1072,23 @@ apply bpow_ge_0.
(* . x=0 *)
rewrite <- Zx, Rplus_0_l; rewrite <- Zx in Heps.
case (proj1 Heps); intros P.
-unfold round, scaled_mantissa, canonic_exp.
+unfold round, scaled_mantissa, cexp.
revert Heps; unfold ulp.
rewrite Req_bool_true; trivial.
case negligible_exp_spec.
intros _ (H1,H2).
-absurd (0 < 0)%R; auto with real.
-now apply Rle_lt_trans with (1:=H1).
+exfalso ; lra.
intros n Hn H.
-assert (fexp (ln_beta beta eps) = fexp n).
+assert (fexp (mag beta eps) = fexp n).
apply valid_exp; try assumption.
-assert(ln_beta beta eps-1 < fexp n)%Z;[idtac|omega].
+assert(mag beta eps-1 < fexp n)%Z;[idtac|omega].
apply lt_bpow with beta.
apply Rle_lt_trans with (2:=proj2 H).
-destruct (ln_beta beta eps) as (e,He).
+destruct (mag beta eps) as (e,He).
simpl; rewrite Rabs_pos_eq in He.
now apply He, Rgt_not_eq.
now left.
-replace (Zfloor (eps * bpow (- fexp (ln_beta beta eps)))) with 0%Z.
+replace (Zfloor (eps * bpow (- fexp (mag beta eps)))) with 0%Z.
unfold F2R; simpl; ring.
apply sym_eq, Zfloor_imp.
split.
@@ -1128,8 +1123,8 @@ assert (Hd := round_DN_plus_eps_pos x Zx Fx eps Heps).
rewrite round_UP_DN_ulp.
rewrite Hd.
rewrite 2!ulp_neq_0.
-unfold canonic_exp.
-now rewrite ln_beta_plus_eps.
+unfold cexp.
+now rewrite mag_plus_eps.
now apply Rgt_not_eq.
now apply Rgt_not_eq, Rplus_lt_0_compat.
intros Fs.
@@ -1144,24 +1139,22 @@ now apply generic_format_succ_aux1.
rewrite <- Zx1, 2!Rplus_0_l.
intros Heps.
case (proj2 Heps).
-unfold round, scaled_mantissa, canonic_exp.
+unfold round, scaled_mantissa, cexp.
unfold ulp.
rewrite Req_bool_true; trivial.
case negligible_exp_spec.
-intros H2.
-intros J; absurd (0 < 0)%R; auto with real.
-apply Rlt_trans with eps; try assumption; apply Heps.
+lra.
intros n Hn H.
-assert (fexp (ln_beta beta eps) = fexp n).
+assert (fexp (mag beta eps) = fexp n).
apply valid_exp; try assumption.
-assert(ln_beta beta eps-1 < fexp n)%Z;[idtac|omega].
+assert(mag beta eps-1 < fexp n)%Z;[idtac|omega].
apply lt_bpow with beta.
apply Rle_lt_trans with (2:=H).
-destruct (ln_beta beta eps) as (e,He).
+destruct (mag beta eps) as (e,He).
simpl; rewrite Rabs_pos_eq in He.
now apply He, Rgt_not_eq.
now left.
-replace (Zceil (eps * bpow (- fexp (ln_beta beta eps)))) with 1%Z.
+replace (Zceil (eps * bpow (- fexp (mag beta eps)))) with 1%Z.
unfold F2R; simpl; rewrite H0; ring.
apply sym_eq, Zceil_imp.
split.
@@ -1316,7 +1309,7 @@ destruct Zp; trivial.
generalize H0.
rewrite pred_eq_pos;[idtac|now left].
unfold pred_pos.
-destruct (ln_beta beta y) as (ey,Hey); simpl.
+destruct (mag beta y) as (ey,Hey); simpl.
case Req_bool_spec; intros Hy2.
(* . *)
intros Hy3.
@@ -1326,7 +1319,7 @@ rewrite <- Hy2, <- Rplus_0_l, Hy3.
ring.
assert (Zx: (x <> 0)%R).
now apply Rgt_not_eq.
-destruct (ln_beta beta x) as (ex,Hex).
+destruct (mag beta x) as (ex,Hex).
specialize (Hex Zx).
assert (ex <= ey)%Z.
apply bpow_lt_bpow with beta.
@@ -1347,16 +1340,16 @@ omega.
absurd (0 < Ztrunc (scaled_mantissa beta fexp x) < 1)%Z.
omega.
split.
-apply F2R_gt_0_reg with beta (canonic_exp beta fexp x).
+apply gt_0_F2R with beta (cexp beta fexp x).
now rewrite <- Fx.
-apply lt_Z2R.
-apply Rmult_lt_reg_r with (bpow (canonic_exp beta fexp x)).
+apply lt_IZR.
+apply Rmult_lt_reg_r with (bpow (cexp beta fexp x)).
apply bpow_gt_0.
-replace (Z2R (Ztrunc (scaled_mantissa beta fexp x)) *
- bpow (canonic_exp beta fexp x))%R with x.
+replace (IZR (Ztrunc (scaled_mantissa beta fexp x)) *
+ bpow (cexp beta fexp x))%R with x.
rewrite Rmult_1_l.
-unfold canonic_exp.
-rewrite ln_beta_unique with beta x ex.
+unfold cexp.
+rewrite mag_unique with beta x ex.
rewrite H3,<-H1, <- Hy2.
apply H.
exact Hex.
@@ -1373,8 +1366,8 @@ assert (y = bpow (fexp ey))%R.
apply Rminus_diag_uniq.
rewrite Hy3.
rewrite ulp_neq_0;[idtac|now apply Rgt_not_eq].
-unfold canonic_exp.
-rewrite (ln_beta_unique beta y ey); trivial.
+unfold cexp.
+rewrite (mag_unique beta y ey); trivial.
apply Hey.
now apply Rgt_not_eq.
contradict Hy2.
@@ -1382,8 +1375,8 @@ rewrite H1.
apply f_equal.
apply Zplus_reg_l with 1%Z.
ring_simplify.
-apply trans_eq with (ln_beta beta y).
-apply sym_eq; apply ln_beta_unique.
+apply trans_eq with (mag beta y).
+apply sym_eq; apply mag_unique.
rewrite H1, Rabs_right.
split.
apply bpow_le.
@@ -1391,7 +1384,7 @@ omega.
apply bpow_lt.
omega.
apply Rle_ge; apply bpow_ge_0.
-apply ln_beta_unique.
+apply mag_unique.
apply Hey.
now apply Rgt_not_eq.
(* *)
@@ -1418,7 +1411,7 @@ rewrite <- V; apply pred_pos_ge_0; trivial.
apply Rle_lt_trans with (1:=proj1 H); apply H.
Qed.
-Theorem succ_le_lt_aux:
+Lemma succ_le_lt_aux:
forall x y,
F x -> F y ->
(0 <= x)%R -> (x < y)%R ->
@@ -1468,7 +1461,7 @@ now apply generic_format_opp.
rewrite Ropp_0; now left.
Qed.
-Theorem le_pred_lt :
+Theorem pred_ge_gt :
forall x y,
F x -> F y ->
(x < y)%R ->
@@ -1483,7 +1476,7 @@ now apply generic_format_opp.
now apply Ropp_lt_contravar.
Qed.
-Theorem lt_succ_le :
+Theorem succ_gt_ge :
forall x y,
(y <> 0)%R ->
(x <= y)%R ->
@@ -1505,12 +1498,12 @@ apply Rlt_le_trans with (2 := Hxy).
now apply pred_lt_id.
Qed.
-Theorem succ_pred_aux : forall x, F x -> (0 < x)%R -> succ (pred x)=x.
+Lemma succ_pred_pos :
+ forall x, F x -> (0 < x)%R -> succ (pred x) = x.
Proof.
intros x Fx Hx.
-rewrite pred_eq_pos;[idtac|now left].
-rewrite succ_eq_pos.
-2: now apply pred_pos_ge_0.
+rewrite pred_eq_pos by now left.
+rewrite succ_eq_pos by now apply pred_pos_ge_0.
now apply pred_pos_plus_ulp.
Qed.
@@ -1530,7 +1523,7 @@ rewrite H1; ring.
(* *)
intros (n,(H1,H2)); rewrite H1.
unfold pred_pos.
-rewrite ln_beta_bpow.
+rewrite mag_bpow.
replace (fexp n + 1 - 1)%Z with (fexp n) by ring.
rewrite Req_bool_true; trivial.
apply Rminus_diag_eq, f_equal.
@@ -1554,7 +1547,7 @@ rewrite <- Ropp_0 at 1.
apply pred_opp.
Qed.
-Theorem pred_succ_aux :
+Lemma pred_succ_pos :
forall x, F x -> (0 < x)%R ->
pred (succ x) = x.
Proof.
@@ -1570,7 +1563,7 @@ apply Rle_antisym.
apply Rlt_le_trans with (1 := Hx).
apply succ_ge_id.
now apply generic_format_pred, generic_format_succ.
-- apply le_pred_lt with (1 := Fx).
+- apply pred_ge_gt with (1 := Fx).
now apply generic_format_succ.
apply succ_gt_id.
now apply Rgt_not_eq.
@@ -1582,12 +1575,12 @@ Theorem succ_pred :
Proof.
intros x Fx.
destruct (Rle_or_lt 0 x) as [[Hx|Hx]|Hx].
-now apply succ_pred_aux.
+now apply succ_pred_pos.
rewrite <- Hx.
rewrite pred_0, succ_opp, pred_ulp_0.
apply Ropp_0.
-rewrite pred_eq_opp_succ_opp, succ_opp.
-rewrite pred_succ_aux.
+unfold pred.
+rewrite succ_opp, pred_succ_pos.
apply Ropp_involutive.
now apply generic_format_opp.
now apply Ropp_0_gt_lt_contravar.
@@ -1606,8 +1599,8 @@ Qed.
Theorem round_UP_pred_plus_eps :
forall x, F x ->
- forall eps, (0 < eps <= if (Rle_bool x 0) then (ulp x)
- else (ulp (pred x)))%R ->
+ forall eps, (0 < eps <= if Rle_bool x 0 then ulp x
+ else ulp (pred x))%R ->
round beta fexp Zceil (pred x + eps) = x.
Proof.
intros x Fx eps Heps.
@@ -1636,7 +1629,6 @@ now apply pred_ge_0.
now apply generic_format_opp.
Qed.
-
Theorem round_DN_minus_eps:
forall x, F x ->
forall eps, (0 < eps <= if (Rle_bool x 0) then (ulp x)
@@ -1676,7 +1668,7 @@ Qed.
(* was ulp_error *)
Theorem error_lt_ulp :
forall rnd { Zrnd : Valid_rnd rnd } x,
- (x <> 0)%R ->
+ (x <> 0)%R ->
(Rabs (round beta fexp rnd x - x) < ulp x)%R.
Proof with auto with typeclass_instances.
intros rnd Zrnd x Zx.
@@ -1734,7 +1726,6 @@ intros Zx; left.
now apply error_lt_ulp.
Qed.
-(* was ulp_half_error *)
Theorem error_le_half_ulp :
forall choice x,
(Rabs (round beta fexp (Znearest choice) x - x) <= /2 * ulp x)%R.
@@ -1748,7 +1739,7 @@ rewrite Rplus_opp_r, Rabs_R0.
apply Rmult_le_pos.
apply Rlt_le.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
apply ulp_ge_0.
(* x <> rnd x *)
set (d := round beta fexp Zfloor x).
@@ -1761,7 +1752,7 @@ apply (round_DN_pt beta fexp x).
rewrite Rabs_left1.
rewrite Ropp_minus_distr.
apply Rmult_le_reg_r with 2%R.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
apply Rplus_le_reg_r with (d - x)%R.
ring_simplify.
apply Rle_trans with (1 := H).
@@ -1778,7 +1769,7 @@ rewrite Hu.
apply (round_UP_pt beta fexp x).
rewrite Rabs_pos_eq.
apply Rmult_le_reg_r with 2%R.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
apply Rplus_le_reg_r with (- (d + ulp x - x))%R.
ring_simplify.
apply Rlt_le.
@@ -1789,28 +1780,40 @@ rewrite Hu.
apply (round_UP_pt beta fexp x).
Qed.
-
Theorem ulp_DN :
- forall x,
- (0 < round beta fexp Zfloor x)%R ->
+ forall x, (0 <= x)%R ->
ulp (round beta fexp Zfloor x) = ulp x.
Proof with auto with typeclass_instances.
-intros x Hd.
-rewrite 2!ulp_neq_0.
-now rewrite canonic_exp_DN with (2 := Hd).
-intros T; contradict Hd; rewrite T, round_0...
-apply Rlt_irrefl.
-now apply Rgt_not_eq.
-Qed.
-
-Theorem round_neq_0_negligible_exp:
- negligible_exp=None -> forall rnd { Zrnd : Valid_rnd rnd } x,
- (x <> 0)%R -> (round beta fexp rnd x <> 0)%R.
+intros x [Hx|Hx].
+- rewrite (ulp_neq_0 x) by now apply Rgt_not_eq.
+ destruct (round_ge_generic beta fexp Zfloor 0 x) as [Hd|Hd].
+ apply generic_format_0.
+ now apply Rlt_le.
+ + rewrite ulp_neq_0 by now apply Rgt_not_eq.
+ now rewrite cexp_DN with (2 := Hd).
+ + rewrite <- Hd.
+ unfold cexp.
+ destruct (mag beta x) as [e He].
+ simpl.
+ specialize (He (Rgt_not_eq _ _ Hx)).
+ apply sym_eq in Hd.
+ assert (H := exp_small_round_0 _ _ _ _ _ He Hd).
+ unfold ulp.
+ rewrite Req_bool_true by easy.
+ destruct negligible_exp_spec as [H0|k Hk].
+ now elim Zlt_not_le with (1 := H0 e).
+ now apply f_equal, fexp_negligible_exp_eq.
+- rewrite <- Hx, round_0...
+Qed.
+
+Theorem round_neq_0_negligible_exp :
+ negligible_exp = None -> forall rnd { Zrnd : Valid_rnd rnd } x,
+ (x <> 0)%R -> (round beta fexp rnd x <> 0)%R.
Proof with auto with typeclass_instances.
intros H rndn Hrnd x Hx K.
case negligible_exp_spec'.
intros (_,Hn).
-destruct (ln_beta beta x) as (e,He).
+destruct (mag beta x) as (e,He).
absurd (fexp e < e)%Z.
apply Zle_not_lt.
apply exp_small_round_0 with beta rndn x...
@@ -1819,12 +1822,10 @@ intros (n,(H1,_)).
rewrite H in H1; discriminate.
Qed.
-
(** allows rnd x to be 0 *)
-(* was ulp_error_f *)
Theorem error_lt_ulp_round :
forall { Hm : Monotone_exp fexp } rnd { Zrnd : Valid_rnd rnd } x,
- ( x <> 0)%R ->
+ (x <> 0)%R ->
(Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R.
Proof with auto with typeclass_instances.
intros Hm.
@@ -1847,72 +1848,34 @@ now apply valid_rnd_opp.
now apply Ropp_0_gt_lt_contravar.
(* 0 < x *)
intros rnd Hrnd x Hx.
-case (Rle_lt_or_eq_dec 0 (round beta fexp Zfloor x)).
-apply round_ge_generic...
-apply generic_format_0.
-now left.
-(* . 0 < round Zfloor x *)
-intros Hx2.
apply Rlt_le_trans with (ulp x).
apply error_lt_ulp...
now apply Rgt_not_eq.
rewrite <- ulp_DN; trivial.
apply ulp_le_pos.
-now left.
+apply round_ge_generic...
+apply generic_format_0.
+now apply Rlt_le.
case (round_DN_or_UP beta fexp rnd x); intros V; rewrite V.
apply Rle_refl.
apply Rle_trans with x.
apply round_DN_pt...
apply round_UP_pt...
-(* . 0 = round Zfloor x *)
-intros Hx2.
-case (round_DN_or_UP beta fexp rnd x); intros V; rewrite V; clear V.
-(* .. round down -- difficult case *)
-rewrite <- Hx2.
-unfold Rminus; rewrite Rplus_0_l, Rabs_Ropp.
-unfold ulp; rewrite Req_bool_true; trivial.
-case negligible_exp_spec.
-(* without minimal exponent *)
-intros K; contradict Hx2.
-apply Rlt_not_eq.
-apply F2R_gt_0_compat; simpl.
-apply Zlt_le_trans with 1%Z.
-apply Pos2Z.is_pos.
-apply Zfloor_lub.
-simpl; unfold scaled_mantissa, canonic_exp.
-destruct (ln_beta beta x) as (e,He); simpl.
-apply Rle_trans with (bpow (e-1) * bpow (- fexp e))%R.
-rewrite <- bpow_plus.
-replace 1%R with (bpow 0) by reflexivity.
-apply bpow_le.
-specialize (K e); omega.
-apply Rmult_le_compat_r.
-apply bpow_ge_0.
-rewrite <- (Rabs_pos_eq x).
-now apply He, Rgt_not_eq.
-now left.
-(* with a minimal exponent *)
-intros n Hn.
-rewrite Rabs_pos_eq;[idtac|now left].
-case (Rle_or_lt (bpow (fexp n)) x); trivial.
-intros K; contradict Hx2.
-apply Rlt_not_eq.
-apply Rlt_le_trans with (bpow (fexp n)).
-apply bpow_gt_0.
-apply round_ge_generic...
-apply generic_format_bpow.
-now apply valid_exp.
-(* .. round up *)
-apply Rlt_le_trans with (ulp x).
-apply error_lt_ulp...
-now apply Rgt_not_eq.
-apply ulp_le_pos.
-now left.
-apply round_UP_pt...
+now apply Rlt_le.
+Qed.
+
+Lemma error_le_ulp_round :
+ forall { Hm : Monotone_exp fexp } rnd { Zrnd : Valid_rnd rnd } x,
+ (Rabs (round beta fexp rnd x - x) <= ulp (round beta fexp rnd x))%R.
+Proof.
+intros Mexp rnd Vrnd x.
+destruct (Req_dec x 0) as [Zx|Nzx].
+{ rewrite Zx, round_0; [|exact Vrnd].
+ unfold Rminus; rewrite Ropp_0, Rplus_0_l, Rabs_R0; apply ulp_ge_0. }
+now apply Rlt_le, error_lt_ulp_round.
Qed.
(** allows both x and rnd x to be 0 *)
-(* was ulp_half_error_f *)
Theorem error_le_half_ulp_round :
forall { Hm : Monotone_exp fexp },
forall choice x,
@@ -1939,11 +1902,11 @@ apply Rle_trans with (1:=N).
right; apply f_equal.
rewrite ulp_neq_0; trivial.
apply f_equal.
-unfold canonic_exp.
+unfold cexp.
apply valid_exp; trivial.
-assert (ln_beta beta x -1 < fexp n)%Z;[idtac|omega].
+assert (mag beta x -1 < fexp n)%Z;[idtac|omega].
apply lt_bpow with beta.
-destruct (ln_beta beta x) as (e,He).
+destruct (mag beta x) as (e,He).
simpl.
apply Rle_lt_trans with (Rabs x).
now apply He.
@@ -1958,42 +1921,29 @@ now right.
(* *)
case (round_DN_or_UP beta fexp (Znearest choice) x); intros Hx.
(* . *)
-case (Rle_or_lt 0 (round beta fexp Zfloor x)).
-intros H; destruct H.
+destruct (Rle_or_lt 0 x) as [H|H].
rewrite Hx at 2.
-rewrite ulp_DN; trivial.
+rewrite ulp_DN by easy.
apply error_le_half_ulp.
-rewrite Hx in Hfx; contradict Hfx; auto with real.
-intros H.
apply Rle_trans with (1:=error_le_half_ulp _ _).
apply Rmult_le_compat_l.
apply Rlt_le, pos_half_prf.
apply ulp_le.
-rewrite Hx; rewrite (Rabs_left1 x), Rabs_left; try assumption.
+rewrite Rabs_left1 by now apply Rlt_le.
+rewrite Hx.
+rewrite Rabs_left1.
apply Ropp_le_contravar.
-apply (round_DN_pt beta fexp x).
-case (Rle_or_lt x 0); trivial.
-intros H1; contradict H.
-apply Rle_not_lt.
-apply round_ge_generic...
+apply round_DN_pt...
+apply round_le_generic...
apply generic_format_0.
-now left.
+now apply Rlt_le.
(* . *)
-case (Rle_or_lt 0 (round beta fexp Zceil x)).
-intros H; destruct H.
+destruct (Rle_or_lt 0 x) as [H|H].
apply Rle_trans with (1:=error_le_half_ulp _ _).
apply Rmult_le_compat_l.
apply Rlt_le, pos_half_prf.
apply ulp_le_pos; trivial.
-case (Rle_or_lt 0 x); trivial.
-intros H1; contradict H.
-apply Rle_not_lt.
-apply round_le_generic...
-apply generic_format_0.
-now left.
rewrite Hx; apply (round_UP_pt beta fexp x).
-rewrite Hx in Hfx; contradict Hfx; auto with real.
-intros H.
rewrite Hx at 2; rewrite <- (ulp_opp (round beta fexp Zceil x)).
rewrite <- round_DN_opp.
rewrite ulp_DN; trivial.
@@ -2002,7 +1952,9 @@ rewrite round_N_opp.
unfold Rminus.
rewrite <- Ropp_plus_distr, Rabs_Ropp.
apply error_le_half_ulp.
-rewrite round_DN_opp; apply Ropp_0_gt_lt_contravar; apply Rlt_gt; assumption.
+rewrite <- Ropp_0.
+apply Ropp_le_contravar.
+now apply Rlt_le.
Qed.
Theorem pred_le :
@@ -2011,18 +1963,22 @@ Theorem pred_le :
Proof.
intros x y Fx Fy [Hxy| ->].
2: apply Rle_refl.
-apply le_pred_lt with (2 := Fy).
+apply pred_ge_gt with (2 := Fy).
now apply generic_format_pred.
apply Rle_lt_trans with (2 := Hxy).
apply pred_le_id.
Qed.
-Theorem succ_le: forall x y,
- F x -> F y -> (x <= y)%R -> (succ x <= succ y)%R.
+Theorem succ_le :
+ forall x y, F x -> F y -> (x <= y)%R ->
+ (succ x <= succ y)%R.
Proof.
intros x y Fx Fy Hxy.
-rewrite 2!succ_eq_opp_pred_opp.
-apply Ropp_le_contravar, pred_le; try apply generic_format_opp; try assumption.
+apply Ropp_le_cancel.
+rewrite <- 2!pred_opp.
+apply pred_le.
+now apply generic_format_opp.
+now apply generic_format_opp.
now apply Ropp_le_contravar.
Qed.
@@ -2064,8 +2020,95 @@ apply Rgt_not_le with (1 := Hxy).
now apply succ_le_inv.
Qed.
-(* was lt_UP_le_DN *)
-Theorem le_round_DN_lt_UP :
+(** Adding [ulp] is a, somewhat reasonable, overapproximation of [succ]. *)
+Lemma succ_le_plus_ulp :
+ forall { Hm : Monotone_exp fexp } x,
+ (succ x <= x + ulp x)%R.
+Proof.
+intros Mexp x.
+destruct (Rle_or_lt 0 x) as [Px|Nx]; [now right; apply succ_eq_pos|].
+replace (_ + _)%R with (- (-x - ulp x))%R by ring.
+unfold succ; rewrite (Rle_bool_false _ _ Nx), <-ulp_opp.
+apply Ropp_le_contravar; unfold pred_pos.
+destruct (Req_dec (-x) (bpow (mag beta (-x) - 1))) as [Hx|Hx].
+{ rewrite (Req_bool_true _ _ Hx).
+ apply (Rplus_le_reg_r x); ring_simplify; apply Ropp_le_contravar.
+ unfold ulp; rewrite Req_bool_false; [|lra].
+ apply bpow_le, Mexp; lia. }
+ now rewrite (Req_bool_false _ _ Hx); right.
+Qed.
+
+(** And it also lies in the format. *)
+Lemma generic_format_plus_ulp :
+ forall { Hm : Monotone_exp fexp } x,
+ generic_format beta fexp x ->
+ generic_format beta fexp (x + ulp x).
+Proof.
+intros Mexp x Fx.
+destruct (Rle_or_lt 0 x) as [Px|Nx].
+{ now rewrite <-(succ_eq_pos _ Px); apply generic_format_succ. }
+apply generic_format_opp in Fx.
+replace (_ + _)%R with (- (-x - ulp x))%R by ring.
+apply generic_format_opp; rewrite <-ulp_opp.
+destruct (Req_dec (-x) (bpow (mag beta (-x) - 1))) as [Hx|Hx].
+{ unfold ulp; rewrite Req_bool_false; [|lra].
+ rewrite Hx at 1.
+ unfold cexp.
+ set (e := mag _ _).
+ assert (Hfe : (fexp e < e)%Z).
+ { now apply mag_generic_gt; [|lra|]. }
+ replace (e - 1)%Z with (e - 1 - fexp e + fexp e)%Z by ring.
+ rewrite bpow_plus.
+ set (m := bpow (_ - _)).
+ replace (_ - _)%R with ((m - 1) * bpow (fexp e))%R; [|unfold m; ring].
+ case_eq (e - 1 - fexp e)%Z.
+ { intro He; unfold m; rewrite He; simpl; ring_simplify (1 - 1)%R.
+ rewrite Rmult_0_l; apply generic_format_0. }
+ { intros p Hp; unfold m; rewrite Hp; simpl.
+ pose (f := {| Defs.Fnum := (Z.pow_pos beta p - 1)%Z;
+ Defs.Fexp := fexp e |} : Defs.float beta).
+ apply (generic_format_F2R' _ _ _ f); [|intro Hm'; unfold f; simpl].
+ { now unfold Defs.F2R; simpl; rewrite minus_IZR. }
+ unfold cexp.
+ replace (IZR _) with (bpow (Z.pos p)); [|now simpl].
+ rewrite <-Hp.
+ assert (He : (1 <= e - 1 - fexp e)%Z); [lia|].
+ set (e' := mag _ (_ * _)).
+ assert (H : (e' = e - 1 :> Z)%Z); [|rewrite H; apply Mexp; lia].
+ unfold e'; apply mag_unique.
+ rewrite Rabs_mult, (Rabs_pos_eq (bpow _)); [|apply bpow_ge_0].
+ rewrite Rabs_pos_eq;
+ [|apply (Rplus_le_reg_r 1); ring_simplify;
+ change 1%R with (bpow 0); apply bpow_le; lia].
+ assert (beta_pos : (0 < IZR beta)%R).
+ { apply (Rlt_le_trans _ 2); [lra|].
+ apply IZR_le, Z.leb_le, radix_prop. }
+ split.
+ { replace (e - 1 - 1)%Z with (e - 1 - fexp e + -1 + fexp e)%Z by ring.
+ rewrite bpow_plus.
+ apply Rmult_le_compat_r; [apply bpow_ge_0|].
+ rewrite bpow_plus; simpl; unfold Z.pow_pos; simpl.
+ rewrite Zmult_1_r.
+ apply (Rmult_le_reg_r _ _ _ beta_pos).
+ rewrite Rmult_assoc, Rinv_l; [|lra]; rewrite Rmult_1_r.
+ apply (Rplus_le_reg_r (IZR beta)); ring_simplify.
+ apply (Rle_trans _ (2 * bpow (e - 1 - fexp e))).
+ { change 2%R with (1 + 1)%R; rewrite Rmult_plus_distr_r, Rmult_1_l.
+ apply Rplus_le_compat_l.
+ rewrite <-bpow_1; apply bpow_le; lia. }
+ rewrite Rmult_comm; apply Rmult_le_compat_l; [apply bpow_ge_0|].
+ apply IZR_le, Z.leb_le, radix_prop. }
+ apply (Rmult_lt_reg_r (bpow (- fexp e))); [apply bpow_gt_0|].
+ rewrite Rmult_assoc, <-!bpow_plus.
+ replace (fexp e + - fexp e)%Z with 0%Z by ring; simpl.
+ rewrite Rmult_1_r; unfold Zminus; lra. }
+ intros p Hp; exfalso; lia. }
+replace (_ - _)%R with (pred_pos (-x)).
+{ now apply generic_format_pred_pos; [|lra]. }
+now unfold pred_pos; rewrite Req_bool_false.
+Qed.
+
+Theorem round_DN_ge_UP_gt :
forall x y, F y ->
(y < round beta fexp Zceil x -> y <= round beta fexp Zfloor x)%R.
Proof with auto with typeclass_instances.
@@ -2078,10 +2121,9 @@ apply round_UP_pt...
now apply Rlt_le.
Qed.
-(* was lt_DN_le_UP *)
-Theorem round_UP_le_gt_DN :
+Theorem round_UP_le_DN_lt :
forall x y, F y ->
- (round beta fexp Zfloor x < y -> round beta fexp Zceil x <= y)%R.
+ (round beta fexp Zfloor x < y -> round beta fexp Zceil x <= y)%R.
Proof with auto with typeclass_instances.
intros x y Fy Hlt.
apply round_UP_pt...
@@ -2092,8 +2134,6 @@ apply round_DN_pt...
now apply Rlt_le.
Qed.
-
-
Theorem pred_UP_le_DN :
forall x, (pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R.
Proof with auto with typeclass_instances.
@@ -2115,16 +2155,26 @@ absurd (round beta fexp Zceil x <= - bpow (fexp n))%R.
apply Rlt_not_le.
rewrite Zx, <- Ropp_0.
apply Ropp_lt_contravar, bpow_gt_0.
-apply round_UP_le_gt_DN; try assumption.
+apply round_UP_le_DN_lt; try assumption.
apply generic_format_opp, generic_format_bpow.
now apply valid_exp.
assert (let u := round beta fexp Zceil x in pred u < u)%R as Hup.
now apply pred_lt_id.
-apply le_round_DN_lt_UP...
+apply round_DN_ge_UP_gt...
apply generic_format_pred...
now apply round_UP_pt.
Qed.
+Theorem UP_le_succ_DN :
+ forall x, (round beta fexp Zceil x <= succ (round beta fexp Zfloor x))%R.
+Proof.
+intros x.
+rewrite <- (Ropp_involutive x).
+rewrite round_DN_opp, round_UP_opp, succ_opp.
+apply Ropp_le_contravar.
+apply pred_UP_le_DN.
+Qed.
+
Theorem pred_UP_eq_DN :
forall x, ~ F x ->
(pred (round beta fexp Zceil x) = round beta fexp Zfloor x)%R.
@@ -2132,7 +2182,7 @@ Proof with auto with typeclass_instances.
intros x Fx.
apply Rle_antisym.
now apply pred_UP_le_DN.
-apply le_pred_lt; try apply generic_format_round...
+apply pred_ge_gt; try apply generic_format_round...
pose proof round_DN_UP_lt _ _ _ Fx as HE.
now apply Rlt_trans with (1 := proj1 HE) (2 := proj2 HE).
Qed.
@@ -2147,11 +2197,9 @@ rewrite succ_pred; trivial.
apply generic_format_round...
Qed.
-
-(* was betw_eq_DN *)
-Theorem round_DN_eq_betw: forall x d, F d
- -> (d <= x < succ d)%R
- -> round beta fexp Zfloor x = d.
+Theorem round_DN_eq :
+ forall x d, F d -> (d <= x < succ d)%R ->
+ round beta fexp Zfloor x = d.
Proof with auto with typeclass_instances.
intros x d Fd (Hxd1,Hxd2).
generalize (round_DN_pt beta fexp x); intros (T1,(T2,T3)).
@@ -2169,25 +2217,161 @@ apply generic_format_succ...
now left.
Qed.
-(* was betw_eq_UP *)
-Theorem round_UP_eq_betw: forall x u, F u
- -> (pred u < x <= u)%R
- -> round beta fexp Zceil x = u.
+Theorem round_UP_eq :
+ forall x u, F u -> (pred u < x <= u)%R ->
+ round beta fexp Zceil x = u.
Proof with auto with typeclass_instances.
intros x u Fu Hux.
rewrite <- (Ropp_involutive (round beta fexp Zceil x)).
rewrite <- round_DN_opp.
rewrite <- (Ropp_involutive u).
apply f_equal.
-apply round_DN_eq_betw; try assumption.
+apply round_DN_eq; try assumption.
now apply generic_format_opp.
split;[now apply Ropp_le_contravar|idtac].
rewrite succ_opp.
now apply Ropp_lt_contravar.
Qed.
+Lemma ulp_ulp_0 : forall {H : Exp_not_FTZ fexp},
+ ulp (ulp 0) = ulp 0.
+Proof.
+intros H; case (negligible_exp_spec').
+intros (K1,K2).
+replace (ulp 0) with 0%R at 1; try easy.
+apply sym_eq; unfold ulp; rewrite Req_bool_true; try easy.
+now rewrite K1.
+intros (n,(Hn1,Hn2)).
+apply Rle_antisym.
+replace (ulp 0) with (bpow (fexp n)).
+rewrite ulp_bpow.
+apply bpow_le.
+now apply valid_exp.
+unfold ulp; rewrite Req_bool_true; try easy.
+rewrite Hn1; easy.
+now apply ulp_ge_ulp_0.
+Qed.
+Lemma ulp_succ_pos : forall x, F x -> (0 < x)%R ->
+ ulp (succ x) = ulp x \/ succ x = bpow (mag beta x).
+Proof with auto with typeclass_instances.
+intros x Fx Hx.
+generalize (Rlt_le _ _ Hx); intros Hx'.
+rewrite succ_eq_pos;[idtac|now left].
+destruct (mag beta x) as (e,He); simpl.
+rewrite Rabs_pos_eq in He; try easy.
+specialize (He (Rgt_not_eq _ _ Hx)).
+assert (H:(x+ulp x <= bpow e)%R).
+apply id_p_ulp_le_bpow; try assumption.
+apply He.
+destruct H;[left|now right].
+rewrite ulp_neq_0 at 1.
+2: apply Rgt_not_eq, Rgt_lt, Rlt_le_trans with x...
+2: rewrite <- (Rplus_0_r x) at 1; apply Rplus_le_compat_l.
+2: apply ulp_ge_0.
+rewrite ulp_neq_0 at 2.
+2: now apply Rgt_not_eq.
+f_equal; unfold cexp; f_equal.
+apply trans_eq with e.
+apply mag_unique_pos; split; try assumption.
+apply Rle_trans with (1:=proj1 He).
+rewrite <- (Rplus_0_r x) at 1; apply Rplus_le_compat_l.
+apply ulp_ge_0.
+now apply sym_eq, mag_unique_pos.
+Qed.
+
+
+Lemma ulp_round_pos :
+ forall { Not_FTZ_ : Exp_not_FTZ fexp},
+ forall rnd { Zrnd : Valid_rnd rnd } x,
+ (0 < x)%R -> ulp (round beta fexp rnd x) = ulp x
+ \/ round beta fexp rnd x = bpow (mag beta x).
+Proof with auto with typeclass_instances.
+intros Not_FTZ_ rnd Zrnd x Hx.
+case (generic_format_EM beta fexp x); intros Fx.
+rewrite round_generic...
+case (round_DN_or_UP beta fexp rnd x); intros Hr; rewrite Hr.
+left.
+apply ulp_DN; now left...
+assert (M:(0 <= round beta fexp Zfloor x)%R).
+apply round_ge_generic...
+apply generic_format_0...
+apply Rlt_le...
+destruct M as [M|M].
+rewrite <- (succ_DN_eq_UP x)...
+case (ulp_succ_pos (round beta fexp Zfloor x)); try intros Y.
+apply generic_format_round...
+assumption.
+rewrite ulp_DN in Y...
+now apply Rlt_le.
+right; rewrite Y.
+apply f_equal, mag_DN...
+left; rewrite <- (succ_DN_eq_UP x)...
+rewrite <- M, succ_0.
+rewrite ulp_ulp_0...
+case (negligible_exp_spec').
+intros (K1,K2).
+absurd (x = 0)%R.
+now apply Rgt_not_eq.
+apply eq_0_round_0_negligible_exp with Zfloor...
+intros (n,(Hn1,Hn2)).
+replace (ulp 0) with (bpow (fexp n)).
+2: unfold ulp; rewrite Req_bool_true; try easy.
+2: now rewrite Hn1.
+rewrite ulp_neq_0.
+2: apply Rgt_not_eq...
+unfold cexp; f_equal.
+destruct (mag beta x) as (e,He); simpl.
+apply sym_eq, valid_exp...
+assert (e <= fexp e)%Z.
+apply exp_small_round_0_pos with beta Zfloor x...
+rewrite <- (Rabs_pos_eq x).
+apply He, Rgt_not_eq...
+apply Rlt_le...
+replace (fexp n) with (fexp e); try assumption.
+now apply fexp_negligible_exp_eq.
+Qed.
+
+
+Theorem ulp_round : forall { Not_FTZ_ : Exp_not_FTZ fexp},
+ forall rnd { Zrnd : Valid_rnd rnd } x,
+ ulp (round beta fexp rnd x) = ulp x
+ \/ Rabs (round beta fexp rnd x) = bpow (mag beta x).
+Proof with auto with typeclass_instances.
+intros Not_FTZ_ rnd Zrnd x.
+case (Rtotal_order x 0); intros Zx.
+case (ulp_round_pos (Zrnd_opp rnd) (-x)).
+now apply Ropp_0_gt_lt_contravar.
+rewrite ulp_opp, <- ulp_opp.
+rewrite <- round_opp, Ropp_involutive.
+intros Y;now left.
+rewrite mag_opp.
+intros Y; right.
+rewrite <- (Ropp_involutive x) at 1.
+rewrite round_opp, Y.
+rewrite Rabs_Ropp, Rabs_right...
+apply Rle_ge, bpow_ge_0.
+destruct Zx as [Zx|Zx].
+left; rewrite Zx; rewrite round_0...
+rewrite Rabs_right.
+apply ulp_round_pos...
+apply Rle_ge, round_ge_generic...
+apply generic_format_0...
+now apply Rlt_le.
+Qed.
+
+Lemma succ_round_ge_id :
+ forall rnd { Zrnd : Valid_rnd rnd } x,
+ (x <= succ (round beta fexp rnd x))%R.
+Proof.
+intros rnd Vrnd x.
+apply (Rle_trans _ (round beta fexp Raux.Zceil x)).
+{ now apply round_UP_pt. }
+destruct (round_DN_or_UP beta fexp rnd x) as [Hr|Hr]; rewrite Hr.
+{ now apply UP_le_succ_DN. }
+apply succ_ge_id.
+Qed.
(** Properties of rounding to nearest and ulp *)
@@ -2215,14 +2399,14 @@ assert (T: (u < (u + succ u) / 2 < succ u)%R) by lra.
destruct T as (T1,T2).
apply Rnd_N_pt_monotone with F v ((u + succ u) / 2)%R...
apply round_N_pt...
-apply Rnd_DN_pt_N with (succ u)%R.
+apply Rnd_N_pt_DN with (succ u)%R.
pattern u at 3; replace u with (round beta fexp Zfloor ((u + succ u) / 2)).
apply round_DN_pt...
-apply round_DN_eq_betw; trivial.
+apply round_DN_eq; trivial.
split; try left; assumption.
pattern (succ u) at 2; replace (succ u) with (round beta fexp Zceil ((u + succ u) / 2)).
apply round_UP_pt...
-apply round_UP_eq_betw; trivial.
+apply round_UP_eq; trivial.
apply generic_format_succ...
rewrite pred_succ; trivial.
split; try left; assumption.
@@ -2275,12 +2459,12 @@ Lemma round_N_eq_DN_pt: forall choice x d u,
Proof with auto with typeclass_instances.
intros choice x d u Hd Hu H.
assert (H0:(d = round beta fexp Zfloor x)%R).
-apply Rnd_DN_pt_unicity with (1:=Hd).
+apply Rnd_DN_pt_unique with (1:=Hd).
apply round_DN_pt...
rewrite H0.
apply round_N_eq_DN.
rewrite <- H0.
-rewrite Rnd_UP_pt_unicity with F x (round beta fexp Zceil x) u; try assumption.
+rewrite Rnd_UP_pt_unique with F x (round beta fexp Zceil x) u; try assumption.
apply round_UP_pt...
Qed.
@@ -2310,13 +2494,28 @@ Lemma round_N_eq_UP_pt: forall choice x d u,
Proof with auto with typeclass_instances.
intros choice x d u Hd Hu H.
assert (H0:(u = round beta fexp Zceil x)%R).
-apply Rnd_UP_pt_unicity with (1:=Hu).
+apply Rnd_UP_pt_unique with (1:=Hu).
apply round_UP_pt...
rewrite H0.
apply round_N_eq_UP.
rewrite <- H0.
-rewrite Rnd_DN_pt_unicity with F x (round beta fexp Zfloor x) d; try assumption.
+rewrite Rnd_DN_pt_unique with F x (round beta fexp Zfloor x) d; try assumption.
apply round_DN_pt...
Qed.
+Lemma round_N_plus_ulp_ge :
+ forall { Hm : Monotone_exp fexp } choice1 choice2 x,
+ let rx := round beta fexp (Znearest choice2) x in
+ (x <= round beta fexp (Znearest choice1) (rx + ulp rx))%R.
+Proof.
+intros Hm choice1 choice2 x.
+simpl.
+set (rx := round _ _ _ x).
+assert (Vrnd1 : Valid_rnd (Znearest choice1)) by now apply valid_rnd_N.
+assert (Vrnd2 : Valid_rnd (Znearest choice2)) by now apply valid_rnd_N.
+apply (Rle_trans _ (succ rx)); [now apply succ_round_ge_id|].
+rewrite round_generic; [now apply succ_le_plus_ulp|now simpl|].
+now apply generic_format_plus_ulp, generic_format_round.
+Qed.
+
End Fcore_ulp.
diff --git a/flocq/Core/Fcore_Zaux.v b/flocq/Core/Zaux.v
index f6731b4c..e21d93a4 100644
--- a/flocq/Core/Fcore_Zaux.v
+++ b/flocq/Core/Zaux.v
@@ -2,9 +2,9 @@
This file is part of the Flocq formalization of floating-point
arithmetic in Coq: http://flocq.gforge.inria.fr/
-Copyright (C) 2011-2013 Sylvie Boldo
+Copyright (C) 2011-2018 Sylvie Boldo
#<br />#
-Copyright (C) 2011-2013 Guillaume Melquiond
+Copyright (C) 2011-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or
modify it under the terms of the GNU Lesser General Public
@@ -17,7 +17,7 @@ MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
COPYING file for more details.
*)
-Require Import ZArith.
+Require Import ZArith Omega.
Require Import Zquot.
Section Zmissing.
@@ -25,7 +25,7 @@ Section Zmissing.
(** About Z *)
Theorem Zopp_le_cancel :
forall x y : Z,
- (-y <= -x)%Z -> Zle x y.
+ (-y <= -x)%Z -> Z.le x y.
Proof.
intros x y Hxy.
apply Zplus_le_reg_r with (-x - y)%Z.
@@ -37,7 +37,7 @@ Theorem Zgt_not_eq :
(y < x)%Z -> (x <> y)%Z.
Proof.
intros x y H Hn.
-apply Zlt_irrefl with x.
+apply Z.lt_irrefl with x.
now rewrite Hn at 1.
Qed.
@@ -69,29 +69,8 @@ End Proof_Irrelevance.
Section Even_Odd.
-(** Zeven, used for rounding to nearest, ties to even *)
-Definition Zeven (n : Z) :=
- match n with
- | Zpos (xO _) => true
- | Zneg (xO _) => true
- | Z0 => true
- | _ => false
- end.
-
-Theorem Zeven_mult :
- forall x y, Zeven (x * y) = orb (Zeven x) (Zeven y).
-Proof.
-now intros [|[xp|xp|]|[xp|xp|]] [|[yp|yp|]|[yp|yp|]].
-Qed.
-
-Theorem Zeven_opp :
- forall x, Zeven (- x) = Zeven x.
-Proof.
-now intros [|[n|n|]|[n|n|]].
-Qed.
-
Theorem Zeven_ex :
- forall x, exists p, x = (2 * p + if Zeven x then 0 else 1)%Z.
+ forall x, exists p, x = (2 * p + if Z.even x then 0 else 1)%Z.
Proof.
intros [|[n|n|]|[n|n|]].
now exists Z0.
@@ -105,37 +84,6 @@ now exists (Zneg n).
now exists (-1)%Z.
Qed.
-Theorem Zeven_2xp1 :
- forall n, Zeven (2 * n + 1) = false.
-Proof.
-intros n.
-destruct (Zeven_ex (2 * n + 1)) as (p, Hp).
-revert Hp.
-case (Zeven (2 * n + 1)) ; try easy.
-intros H.
-apply False_ind.
-omega.
-Qed.
-
-Theorem Zeven_plus :
- forall x y, Zeven (x + y) = Bool.eqb (Zeven x) (Zeven y).
-Proof.
-intros x y.
-destruct (Zeven_ex x) as (px, Hx).
-rewrite Hx at 1.
-destruct (Zeven_ex y) as (py, Hy).
-rewrite Hy at 1.
-replace (2 * px + (if Zeven x then 0 else 1) + (2 * py + (if Zeven y then 0 else 1)))%Z
- with (2 * (px + py) + ((if Zeven x then 0 else 1) + (if Zeven y then 0 else 1)))%Z by ring.
-case (Zeven x) ; case (Zeven y).
-rewrite Zplus_0_r.
-now rewrite Zeven_mult.
-apply Zeven_2xp1.
-apply Zeven_2xp1.
-replace (2 * (px + py) + (1 + 1))%Z with (2 * (px + py + 1))%Z by ring.
-now rewrite Zeven_mult.
-Qed.
-
End Even_Odd.
Section Zpower.
@@ -145,12 +93,12 @@ Theorem Zpower_plus :
Zpower n (k1 + k2) = (Zpower n k1 * Zpower n k2)%Z.
Proof.
intros n k1 k2 H1 H2.
-now apply Zpower_exp ; apply Zle_ge.
+now apply Zpower_exp ; apply Z.le_ge.
Qed.
Theorem Zpower_Zpower_nat :
forall b e, (0 <= e)%Z ->
- Zpower b e = Zpower_nat b (Zabs_nat e).
+ Zpower b e = Zpower_nat b (Z.abs_nat e).
Proof.
intros b [|e|e] He.
apply refl_equal.
@@ -181,40 +129,14 @@ rewrite Zpower_nat_S.
now apply Zmult_lt_0_compat.
Qed.
-Theorem Zeven_Zpower :
- forall b e, (0 < e)%Z ->
- Zeven (Zpower b e) = Zeven b.
-Proof.
-intros b e He.
-case_eq (Zeven b) ; intros Hb.
-(* b even *)
-replace e with (e - 1 + 1)%Z by ring.
-rewrite Zpower_exp.
-rewrite Zeven_mult.
-replace (Zeven (b ^ 1)) with true.
-apply Bool.orb_true_r.
-unfold Zpower, Zpower_pos. simpl.
-now rewrite Zmult_1_r.
-omega.
-discriminate.
-(* b odd *)
-rewrite Zpower_Zpower_nat.
-induction (Zabs_nat e).
-easy.
-unfold Zpower_nat. simpl.
-rewrite Zeven_mult.
-now rewrite Hb.
-now apply Zlt_le_weak.
-Qed.
-
Theorem Zeven_Zpower_odd :
- forall b e, (0 <= e)%Z -> Zeven b = false ->
- Zeven (Zpower b e) = false.
+ forall b e, (0 <= e)%Z -> Z.even b = false ->
+ Z.even (Zpower b e) = false.
Proof.
intros b e He Hb.
destruct (Z_le_lt_eq_dec _ _ He) as [He'|He'].
rewrite <- Hb.
-now apply Zeven_Zpower.
+now apply Z.even_pow.
now rewrite <- He'.
Qed.
@@ -239,7 +161,7 @@ Variable r : radix.
Theorem radix_gt_0 : (0 < r)%Z.
Proof.
-apply Zlt_le_trans with 2%Z.
+apply Z.lt_le_trans with 2%Z.
easy.
apply Zle_bool_imp_le.
apply r.
@@ -248,7 +170,7 @@ Qed.
Theorem radix_gt_1 : (1 < r)%Z.
Proof.
destruct r as (v, Hr). simpl.
-apply Zlt_le_trans with 2%Z.
+apply Z.lt_le_trans with 2%Z.
easy.
now apply Zle_bool_imp_le.
Qed.
@@ -273,7 +195,7 @@ easy.
rewrite Zpower_nat_S.
apply Zmult_lt_0_compat with (2 := IHn).
apply radix_gt_0.
-apply Zle_lt_trans with (1 * Zpower_nat r n)%Z.
+apply Z.le_lt_trans with (1 * Zpower_nat r n)%Z.
rewrite Zmult_1_l.
now apply (Zlt_le_succ 0).
apply Zmult_lt_compat_r with (1 := H).
@@ -287,7 +209,7 @@ Theorem Zpower_gt_0 :
Proof.
intros p Hp.
rewrite Zpower_Zpower_nat with (1 := Hp).
-induction (Zabs_nat p).
+induction (Z.abs_nat p).
easy.
rewrite Zpower_nat_S.
apply Zmult_lt_0_compat with (2 := IHn).
@@ -336,7 +258,7 @@ rewrite <- (Zmult_1_r (r ^ e1)) at 1.
apply Zmult_lt_compat2.
split.
now apply Zpower_gt_0.
-apply Zle_refl.
+apply Z.le_refl.
split.
easy.
apply Zpower_gt_1.
@@ -363,6 +285,36 @@ apply Zpower_le.
clear -H ; omega.
Qed.
+Theorem Zpower_gt_id :
+ forall n, (n < Zpower r n)%Z.
+Proof.
+intros [|n|n] ; try easy.
+simpl.
+rewrite Zpower_pos_nat.
+rewrite Zpos_eq_Z_of_nat_o_nat_of_P.
+induction (nat_of_P n).
+easy.
+rewrite inj_S.
+change (Zpower_nat r (S n0)) with (r * Zpower_nat r n0)%Z.
+unfold Z.succ.
+apply Z.lt_le_trans with (r * (Z_of_nat n0 + 1))%Z.
+clear.
+apply Zlt_0_minus_lt.
+replace (r * (Z_of_nat n0 + 1) - (Z_of_nat n0 + 1))%Z with ((r - 1) * (Z_of_nat n0 + 1))%Z by ring.
+apply Zmult_lt_0_compat.
+cut (2 <= r)%Z. omega.
+apply Zle_bool_imp_le.
+apply r.
+apply (Zle_lt_succ 0).
+apply Zle_0_nat.
+apply Zmult_le_compat_l.
+now apply Zlt_le_succ.
+apply Z.le_trans with 2%Z.
+easy.
+apply Zle_bool_imp_le.
+apply r.
+Qed.
+
End Zpower.
Section Div_Mod.
@@ -380,7 +332,7 @@ rewrite Zopp_mult_distr_l.
apply Z_mod_plus.
easy.
apply Zmult_gt_0_compat.
-now apply Zlt_gt.
+now apply Z.lt_gt.
easy.
now elim Hb.
Qed.
@@ -411,7 +363,7 @@ Qed.
Theorem Zdiv_mod_mult :
forall n a b, (0 <= a)%Z -> (0 <= b)%Z ->
- (Zdiv (Zmod n (a * b)) a) = Zmod (Zdiv n a) b.
+ (Z.div (Zmod n (a * b)) a) = Zmod (Z.div n a) b.
Proof.
intros n a b Ha Hb.
destruct (Zle_lt_or_eq _ _ Ha) as [Ha'|Ha'].
@@ -421,12 +373,12 @@ rewrite (Zmult_comm a b) at 2.
rewrite Zmult_assoc.
unfold Zminus.
rewrite Zopp_mult_distr_l.
-rewrite Z_div_plus by now apply Zlt_gt.
+rewrite Z_div_plus by now apply Z.lt_gt.
rewrite <- Zdiv_Zdiv by easy.
apply sym_eq.
apply Zmod_eq.
-now apply Zlt_gt.
-now apply Zmult_gt_0_compat ; apply Zlt_gt.
+now apply Z.lt_gt.
+now apply Zmult_gt_0_compat ; apply Z.lt_gt.
rewrite <- Hb'.
rewrite Zmult_0_r, 2!Zmod_0_r.
apply Zdiv_0_l.
@@ -439,7 +391,7 @@ Theorem ZOdiv_mod_mult :
(Z.quot (Z.rem n (a * b)) a) = Z.rem (Z.quot n a) b.
Proof.
intros n a b.
-destruct (Z_eq_dec a 0) as [Za|Za].
+destruct (Z.eq_dec a 0) as [Za|Za].
rewrite Za.
now rewrite 2!Zquot_0_r, Zrem_0_l.
assert (Z.rem n (a * b) = n + - (Z.quot (Z.quot n a) b * b) * a)%Z.
@@ -456,34 +408,34 @@ Qed.
Theorem ZOdiv_small_abs :
forall a b,
- (Zabs a < b)%Z -> Z.quot a b = Z0.
+ (Z.abs a < b)%Z -> Z.quot a b = Z0.
Proof.
intros a b Ha.
destruct (Zle_or_lt 0 a) as [H|H].
-apply Zquot_small.
+apply Z.quot_small.
split.
exact H.
-now rewrite Zabs_eq in Ha.
-apply Zopp_inj.
-rewrite <- Zquot_opp_l, Zopp_0.
-apply Zquot_small.
+now rewrite Z.abs_eq in Ha.
+apply Z.opp_inj.
+rewrite <- Zquot_opp_l, Z.opp_0.
+apply Z.quot_small.
generalize (Zabs_non_eq a).
omega.
Qed.
Theorem ZOmod_small_abs :
forall a b,
- (Zabs a < b)%Z -> Z.rem a b = a.
+ (Z.abs a < b)%Z -> Z.rem a b = a.
Proof.
intros a b Ha.
destruct (Zle_or_lt 0 a) as [H|H].
-apply Zrem_small.
+apply Z.rem_small.
split.
exact H.
-now rewrite Zabs_eq in Ha.
-apply Zopp_inj.
+now rewrite Z.abs_eq in Ha.
+apply Z.opp_inj.
rewrite <- Zrem_opp_l.
-apply Zrem_small.
+apply Z.rem_small.
generalize (Zabs_non_eq a).
omega.
Qed.
@@ -493,7 +445,7 @@ Theorem ZOdiv_plus :
(Z.quot (a + b) c = Z.quot a c + Z.quot b c + Z.quot (Z.rem a c + Z.rem b c) c)%Z.
Proof.
intros a b c Hab.
-destruct (Z_eq_dec c 0) as [Zc|Zc].
+destruct (Z.eq_dec c 0) as [Zc|Zc].
now rewrite Zc, 4!Zquot_0_r.
apply Zmult_reg_r with (1 := Zc).
rewrite 2!Zmult_plus_distr_l.
@@ -632,8 +584,8 @@ Proof.
intros x y Hxy.
generalize (Zle_cases x y).
case Zle_bool ; intros H.
-elim (Zlt_irrefl x).
-now apply Zle_lt_trans with y.
+elim (Z.lt_irrefl x).
+now apply Z.le_lt_trans with y.
apply refl_equal.
Qed.
@@ -672,8 +624,8 @@ Proof.
intros x y Hxy.
generalize (Zlt_cases x y).
case Zlt_bool ; intros H.
-elim (Zlt_irrefl x).
-now apply Zlt_le_trans with y.
+elim (Z.lt_irrefl x).
+now apply Z.lt_le_trans with y.
apply refl_equal.
Qed.
@@ -707,32 +659,32 @@ Inductive Zcompare_prop (x y : Z) : comparison -> Prop :=
| Zcompare_Gt_ : (y < x)%Z -> Zcompare_prop x y Gt.
Theorem Zcompare_spec :
- forall x y, Zcompare_prop x y (Zcompare x y).
+ forall x y, Zcompare_prop x y (Z.compare x y).
Proof.
intros x y.
destruct (Z_dec x y) as [[H|H]|H].
generalize (Zlt_compare _ _ H).
-case (Zcompare x y) ; try easy.
+case (Z.compare x y) ; try easy.
now constructor.
generalize (Zgt_compare _ _ H).
-case (Zcompare x y) ; try easy.
+case (Z.compare x y) ; try easy.
constructor.
-now apply Zgt_lt.
+now apply Z.gt_lt.
generalize (proj2 (Zcompare_Eq_iff_eq _ _) H).
-case (Zcompare x y) ; try easy.
+case (Z.compare x y) ; try easy.
now constructor.
Qed.
Theorem Zcompare_Lt :
forall x y,
- (x < y)%Z -> Zcompare x y = Lt.
+ (x < y)%Z -> Z.compare x y = Lt.
Proof.
easy.
Qed.
Theorem Zcompare_Eq :
forall x y,
- (x = y)%Z -> Zcompare x y = Eq.
+ (x = y)%Z -> Z.compare x y = Eq.
Proof.
intros x y.
apply <- Zcompare_Eq_iff_eq.
@@ -740,21 +692,29 @@ Qed.
Theorem Zcompare_Gt :
forall x y,
- (y < x)%Z -> Zcompare x y = Gt.
+ (y < x)%Z -> Z.compare x y = Gt.
Proof.
intros x y.
-apply Zlt_gt.
+apply Z.lt_gt.
Qed.
End Zcompare.
Section cond_Zopp.
-Definition cond_Zopp (b : bool) m := if b then Zopp m else m.
+Definition cond_Zopp (b : bool) m := if b then Z.opp m else m.
+
+Theorem cond_Zopp_negb :
+ forall x y, cond_Zopp (negb x) y = Z.opp (cond_Zopp x y).
+Proof.
+intros [|] y.
+apply sym_eq, Z.opp_involutive.
+easy.
+Qed.
Theorem abs_cond_Zopp :
forall b m,
- Zabs (cond_Zopp b m) = Zabs m.
+ Z.abs (cond_Zopp b m) = Z.abs m.
Proof.
intros [|] m.
apply Zabs_Zopp.
@@ -763,14 +723,14 @@ Qed.
Theorem cond_Zopp_Zlt_bool :
forall m,
- cond_Zopp (Zlt_bool m 0) m = Zabs m.
+ cond_Zopp (Zlt_bool m 0) m = Z.abs m.
Proof.
intros m.
apply sym_eq.
case Zlt_bool_spec ; intros Hm.
apply Zabs_non_eq.
now apply Zlt_le_weak.
-now apply Zabs_eq.
+now apply Z.abs_eq.
Qed.
End cond_Zopp.
@@ -808,11 +768,11 @@ Section faster_div.
Lemma Zdiv_eucl_unique :
forall a b,
- Zdiv_eucl a b = (Zdiv a b, Zmod a b).
+ Z.div_eucl a b = (Z.div a b, Zmod a b).
Proof.
intros a b.
-unfold Zdiv, Zmod.
-now case Zdiv_eucl.
+unfold Z.div, Zmod.
+now case Z.div_eucl.
Qed.
Fixpoint Zpos_div_eucl_aux1 (a b : positive) {struct b} :=
@@ -835,7 +795,7 @@ intros a b.
revert a.
induction b ; intros a.
- easy.
-- change (Z.pos_div_eucl a (Zpos b~0)) with (Zdiv_eucl (Zpos a) (Zpos b~0)).
+- change (Z.pos_div_eucl a (Zpos b~0)) with (Z.div_eucl (Zpos a) (Zpos b~0)).
rewrite Zdiv_eucl_unique.
change (Zpos b~0) with (2 * Zpos b)%Z.
rewrite Z.rem_mul_r by easy.
@@ -843,7 +803,7 @@ induction b ; intros a.
destruct a as [a|a|].
+ change (Zpos_div_eucl_aux1 a~1 b~0) with (let (q, r) := Zpos_div_eucl_aux1 a b in (q, 2 * r + 1)%Z).
rewrite IHb. clear IHb.
- change (Z.pos_div_eucl a (Zpos b)) with (Zdiv_eucl (Zpos a) (Zpos b)).
+ change (Z.pos_div_eucl a (Zpos b)) with (Z.div_eucl (Zpos a) (Zpos b)).
rewrite Zdiv_eucl_unique.
change (Zpos a~1) with (1 + 2 * Zpos a)%Z.
rewrite (Zmult_comm 2 (Zpos a)).
@@ -853,7 +813,7 @@ induction b ; intros a.
apply Zplus_comm.
+ change (Zpos_div_eucl_aux1 a~0 b~0) with (let (q, r) := Zpos_div_eucl_aux1 a b in (q, 2 * r)%Z).
rewrite IHb. clear IHb.
- change (Z.pos_div_eucl a (Zpos b)) with (Zdiv_eucl (Zpos a) (Zpos b)).
+ change (Z.pos_div_eucl a (Zpos b)) with (Z.div_eucl (Zpos a) (Zpos b)).
rewrite Zdiv_eucl_unique.
change (Zpos a~0) with (2 * Zpos a)%Z.
rewrite (Zmult_comm 2 (Zpos a)).
@@ -861,7 +821,7 @@ induction b ; intros a.
apply f_equal.
now rewrite Z_mod_mult.
+ easy.
-- change (Z.pos_div_eucl a 1) with (Zdiv_eucl (Zpos a) 1).
+- change (Z.pos_div_eucl a 1) with (Z.div_eucl (Zpos a) 1).
rewrite Zdiv_eucl_unique.
now rewrite Zdiv_1_r, Zmod_1_r.
Qed.
@@ -879,13 +839,13 @@ Lemma Zpos_div_eucl_aux_correct :
Proof.
intros a b.
unfold Zpos_div_eucl_aux.
-change (Z.pos_div_eucl a (Zpos b)) with (Zdiv_eucl (Zpos a) (Zpos b)).
+change (Z.pos_div_eucl a (Zpos b)) with (Z.div_eucl (Zpos a) (Zpos b)).
rewrite Zdiv_eucl_unique.
case Pos.compare_spec ; intros H.
now rewrite H, Z_div_same, Z_mod_same.
now rewrite Zdiv_small, Zmod_small by (split ; easy).
rewrite Zpos_div_eucl_aux1_correct.
-change (Z.pos_div_eucl a (Zpos b)) with (Zdiv_eucl (Zpos a) (Zpos b)).
+change (Z.pos_div_eucl a (Zpos b)) with (Z.div_eucl (Zpos a) (Zpos b)).
apply Zdiv_eucl_unique.
Qed.
@@ -920,7 +880,7 @@ Definition Zfast_div_eucl (a b : Z) :=
Theorem Zfast_div_eucl_correct :
forall a b : Z,
- Zfast_div_eucl a b = Zdiv_eucl a b.
+ Zfast_div_eucl a b = Z.div_eucl a b.
Proof.
unfold Zfast_div_eucl.
intros [|a|a] [|b|b] ; try rewrite Zpos_div_eucl_aux_correct ; easy.