aboutsummaryrefslogtreecommitdiffstats
path: root/flocq/Prop
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/Prop
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/Prop')
-rw-r--r--flocq/Prop/Div_sqrt_error.v872
-rw-r--r--flocq/Prop/Double_rounding.v4545
-rw-r--r--flocq/Prop/Fprop_div_sqrt_error.v300
-rw-r--r--flocq/Prop/Mult_error.v (renamed from flocq/Prop/Fprop_mult_error.v)175
-rw-r--r--flocq/Prop/Plus_error.v (renamed from flocq/Prop/Fprop_plus_error.v)394
-rw-r--r--flocq/Prop/Relative.v (renamed from flocq/Prop/Fprop_relative.v)505
-rw-r--r--flocq/Prop/Round_odd.v1220
-rw-r--r--flocq/Prop/Sterbenz.v (renamed from flocq/Prop/Fprop_Sterbenz.v)64
8 files changed, 7392 insertions, 683 deletions
diff --git a/flocq/Prop/Div_sqrt_error.v b/flocq/Prop/Div_sqrt_error.v
new file mode 100644
index 00000000..76c7af95
--- /dev/null
+++ b/flocq/Prop/Div_sqrt_error.v
@@ -0,0 +1,872 @@
+(**
+This file is part of the Flocq formalization of floating-point
+arithmetic in Coq: http://flocq.gforge.inria.fr/
+
+Copyright (C) 2010-2018 Sylvie Boldo
+#<br />#
+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
+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.
+*)
+
+(** * Remainder of the division and square root are in the FLX format *)
+
+Require Import Psatz.
+Require Import Core Operations Relative Sterbenz Mult_error.
+
+Section Fprop_divsqrt_error.
+
+Variable beta : radix.
+Notation bpow e := (bpow beta e).
+
+Variable prec : Z.
+
+Lemma generic_format_plus_prec :
+ forall fexp, (forall e, (fexp e <= e - prec)%Z) ->
+ forall x y (fx fy: float beta),
+ (x = F2R fx)%R -> (y = F2R fy)%R -> (Rabs (x+y) < bpow (prec+Fexp fx))%R ->
+ (Rabs (x+y) < bpow (prec+Fexp fy))%R ->
+ generic_format beta fexp (x+y)%R.
+Proof.
+intros fexp Hfexp x y fx fy Hx Hy H1 H2.
+case (Req_dec (x+y) 0); intros H.
+rewrite H; apply generic_format_0.
+rewrite Hx, Hy, <- F2R_plus.
+apply generic_format_F2R.
+intros _.
+case_eq (Fplus fx fy).
+intros mz ez Hz.
+rewrite <- Hz.
+apply Z.le_trans with (Z.min (Fexp fx) (Fexp fy)).
+rewrite F2R_plus, <- Hx, <- Hy.
+unfold cexp.
+apply Z.le_trans with (1:=Hfexp _).
+apply Zplus_le_reg_l with prec; ring_simplify.
+apply mag_le_bpow with (1 := H).
+now apply Z.min_case.
+rewrite <- Fexp_Fplus, Hz.
+apply Z.le_refl.
+Qed.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+
+Notation format := (generic_format beta (FLX_exp prec)).
+Notation cexp := (cexp beta (FLX_exp prec)).
+
+Variable choice : Z -> bool.
+
+
+(** Remainder of the division in FLX *)
+Theorem div_error_FLX :
+ forall rnd { Zrnd : Valid_rnd rnd } x y,
+ format x -> format y ->
+ format (x - round beta (FLX_exp prec) rnd (x/y) * y)%R.
+Proof with auto with typeclass_instances.
+intros rnd Zrnd x y Hx Hy.
+destruct (Req_dec y 0) as [Zy|Zy].
+now rewrite Zy, Rmult_0_r, Rminus_0_r.
+destruct (Req_dec (round beta (FLX_exp prec) rnd (x/y)) 0) as [Hr|Hr].
+rewrite Hr; ring_simplify (x-0*y)%R; assumption.
+assert (Zx: x <> R0).
+contradict Hr.
+rewrite Hr.
+unfold Rdiv.
+now rewrite Rmult_0_l, round_0.
+destruct (canonical_generic_format _ _ x Hx) as (fx,(Hx1,Hx2)).
+destruct (canonical_generic_format _ _ y Hy) as (fy,(Hy1,Hy2)).
+destruct (canonical_generic_format beta (FLX_exp prec) (round beta (FLX_exp prec) rnd (x / y))) as (fr,(Hr1,Hr2)).
+apply generic_format_round...
+unfold Rminus; apply generic_format_plus_prec with fx (Fopp (Fmult fr fy)); trivial.
+intros e; apply Z.le_refl.
+now rewrite F2R_opp, F2R_mult, <- Hr1, <- Hy1.
+(* *)
+destruct (relative_error_FLX_ex beta prec (prec_gt_0 prec) rnd (x / y)%R) as (eps,(Heps1,Heps2)).
+rewrite Heps2.
+rewrite <- Rabs_Ropp.
+replace (-(x + - (x / y * (1 + eps) * y)))%R with (x * eps)%R by now field.
+rewrite Rabs_mult.
+apply Rlt_le_trans with (Rabs x * 1)%R.
+apply Rmult_lt_compat_l.
+now apply Rabs_pos_lt.
+apply Rlt_le_trans with (1 := Heps1).
+change 1%R with (bpow 0).
+apply bpow_le.
+generalize (prec_gt_0 prec).
+clear ; omega.
+rewrite Rmult_1_r.
+rewrite Hx2, <- Hx1.
+unfold cexp.
+destruct (mag beta x) as (ex, Hex).
+simpl.
+specialize (Hex Zx).
+apply Rlt_le.
+apply Rlt_le_trans with (1 := proj2 Hex).
+apply bpow_le.
+unfold FLX_exp.
+ring_simplify.
+apply Z.le_refl.
+(* *)
+replace (Fexp (Fopp (Fmult fr fy))) with (Fexp fr + Fexp fy)%Z.
+2: unfold Fopp, Fmult; destruct fr; destruct fy; now simpl.
+replace (x + - (round beta (FLX_exp prec) rnd (x / y) * y))%R with
+ (y * (-(round beta (FLX_exp prec) rnd (x / y) - x/y)))%R.
+2: field; assumption.
+rewrite Rabs_mult.
+apply Rlt_le_trans with (Rabs y * bpow (Fexp fr))%R.
+apply Rmult_lt_compat_l.
+now apply Rabs_pos_lt.
+rewrite Rabs_Ropp.
+replace (bpow (Fexp fr)) with (ulp beta (FLX_exp prec) (F2R fr)).
+rewrite <- Hr1.
+apply error_lt_ulp_round...
+apply Rmult_integral_contrapositive_currified; try apply Rinv_neq_0_compat; assumption.
+rewrite ulp_neq_0.
+2: now rewrite <- Hr1.
+apply f_equal.
+now rewrite Hr2, <- Hr1.
+replace (prec+(Fexp fr+Fexp fy))%Z with ((prec+Fexp fy)+Fexp fr)%Z by ring.
+rewrite bpow_plus.
+apply Rmult_le_compat_r.
+apply bpow_ge_0.
+rewrite Hy2, <- Hy1 ; unfold cexp, FLX_exp.
+ring_simplify (prec + (mag beta y - prec))%Z.
+destruct (mag beta y); simpl.
+left; now apply a.
+Qed.
+
+(** Remainder of the square in FLX (with p>1) and rounding to nearest *)
+Variable Hp1 : Z.lt 1 prec.
+
+Theorem sqrt_error_FLX_N :
+ forall x, format x ->
+ format (x - Rsqr (round beta (FLX_exp prec) (Znearest choice) (sqrt x)))%R.
+Proof with auto with typeclass_instances.
+intros x Hx.
+destruct (total_order_T x 0) as [[Hxz|Hxz]|Hxz].
+unfold sqrt.
+destruct (Rcase_abs x).
+rewrite round_0...
+unfold Rsqr.
+now rewrite Rmult_0_l, Rminus_0_r.
+elim (Rlt_irrefl 0).
+now apply Rgt_ge_trans with x.
+rewrite Hxz, sqrt_0, round_0...
+unfold Rsqr.
+rewrite Rmult_0_l, Rminus_0_r.
+apply generic_format_0.
+case (Req_dec (round beta (FLX_exp prec) (Znearest choice) (sqrt x)) 0); intros Hr.
+rewrite Hr; unfold Rsqr; ring_simplify (x-0*0)%R; assumption.
+destruct (canonical_generic_format _ _ x Hx) as (fx,(Hx1,Hx2)).
+destruct (canonical_generic_format beta (FLX_exp prec) (round beta (FLX_exp prec) (Znearest choice) (sqrt x))) as (fr,(Hr1,Hr2)).
+apply generic_format_round...
+unfold Rminus; apply generic_format_plus_prec with fx (Fopp (Fmult fr fr)); trivial.
+intros e; apply Z.le_refl.
+unfold Rsqr; now rewrite F2R_opp,F2R_mult, <- Hr1.
+(* *)
+apply Rle_lt_trans with x.
+apply Rabs_minus_le.
+apply Rle_0_sqr.
+destruct (relative_error_N_FLX_ex beta prec (prec_gt_0 prec) choice (sqrt x)) as (eps,(Heps1,Heps2)).
+rewrite Heps2.
+rewrite Rsqr_mult, Rsqr_sqrt, Rmult_comm. 2: now apply Rlt_le.
+apply Rmult_le_compat_r.
+now apply Rlt_le.
+apply Rle_trans with (5²/4²)%R.
+rewrite <- Rsqr_div.
+apply Rsqr_le_abs_1.
+apply Rle_trans with (1 := Rabs_triang _ _).
+rewrite Rabs_R1.
+apply Rplus_le_reg_l with (-1)%R.
+replace (-1 + (1 + Rabs eps))%R with (Rabs eps) by ring.
+apply Rle_trans with (1 := Heps1).
+rewrite Rabs_pos_eq.
+apply Rmult_le_reg_l with 2%R.
+now apply IZR_lt.
+rewrite <- Rmult_assoc, Rinv_r, Rmult_1_l.
+apply Rle_trans with (bpow (-1)).
+apply bpow_le.
+omega.
+replace (2 * (-1 + 5 / 4))%R with (/2)%R by field.
+apply Rinv_le.
+now apply IZR_lt.
+apply IZR_le.
+unfold Zpower_pos. simpl.
+rewrite Zmult_1_r.
+apply Zle_bool_imp_le.
+apply beta.
+now apply IZR_neq.
+unfold Rdiv.
+apply Rmult_le_pos.
+now apply IZR_le.
+apply Rlt_le.
+apply Rinv_0_lt_compat.
+now apply IZR_lt.
+now apply IZR_neq.
+unfold Rsqr.
+replace (5 * 5 / (4 * 4))%R with (25 * /16)%R by field.
+apply Rmult_le_reg_r with 16%R.
+now apply IZR_lt.
+rewrite Rmult_assoc, Rinv_l, Rmult_1_r.
+now apply (IZR_le _ 32).
+now apply IZR_neq.
+rewrite Hx2, <- Hx1; unfold cexp, FLX_exp.
+ring_simplify (prec + (mag beta x - prec))%Z.
+destruct (mag beta x); simpl.
+rewrite <- (Rabs_right x).
+apply a.
+now apply Rgt_not_eq.
+now apply Rgt_ge.
+(* *)
+replace (Fexp (Fopp (Fmult fr fr))) with (Fexp fr + Fexp fr)%Z.
+2: unfold Fopp, Fmult; destruct fr; now simpl.
+rewrite Hr1.
+replace (x + - Rsqr (F2R fr))%R with (-((F2R fr - sqrt x)*(F2R fr + sqrt x)))%R.
+2: rewrite <- (sqrt_sqrt x) at 3; auto.
+2: unfold Rsqr; ring.
+rewrite Rabs_Ropp, Rabs_mult.
+apply Rle_lt_trans with ((/2*bpow (Fexp fr))* Rabs (F2R fr + sqrt x))%R.
+apply Rmult_le_compat_r.
+apply Rabs_pos.
+apply Rle_trans with (/2*ulp beta (FLX_exp prec) (F2R fr))%R.
+rewrite <- Hr1.
+apply error_le_half_ulp_round...
+right; rewrite ulp_neq_0.
+2: now rewrite <- Hr1.
+apply f_equal.
+rewrite Hr2, <- Hr1; trivial.
+rewrite Rmult_assoc, Rmult_comm.
+replace (prec+(Fexp fr+Fexp fr))%Z with (Fexp fr + (prec+Fexp fr))%Z by ring.
+rewrite bpow_plus, Rmult_assoc.
+apply Rmult_lt_compat_l.
+apply bpow_gt_0.
+apply Rmult_lt_reg_l with (1 := Rlt_0_2).
+apply Rle_lt_trans with (Rabs (F2R fr + sqrt x)).
+right; field.
+apply Rle_lt_trans with (1:=Rabs_triang _ _).
+(* . *)
+assert (Rabs (F2R fr) < bpow (prec + Fexp fr))%R.
+rewrite Hr2.
+unfold cexp, FLX_exp.
+ring_simplify (prec + (mag beta (F2R fr) - prec))%Z.
+destruct (mag beta (F2R fr)); simpl.
+apply a.
+rewrite <- Hr1; auto.
+(* . *)
+apply Rlt_le_trans with (bpow (prec + Fexp fr)+ Rabs (sqrt x))%R.
+now apply Rplus_lt_compat_r.
+(* . *)
+replace (2 * bpow (prec + Fexp fr))%R with (bpow (prec + Fexp fr) + bpow (prec + Fexp fr))%R by ring.
+apply Rplus_le_compat_l.
+assert (sqrt x <> 0)%R.
+apply Rgt_not_eq.
+now apply sqrt_lt_R0.
+destruct (mag beta (sqrt x)) as (es,Es).
+specialize (Es H0).
+apply Rle_trans with (bpow es).
+now apply Rlt_le.
+apply bpow_le.
+case (Zle_or_lt es (prec + Fexp fr)) ; trivial.
+intros H1.
+absurd (Rabs (F2R fr) < bpow (es - 1))%R.
+apply Rle_not_lt.
+rewrite <- Hr1.
+apply abs_round_ge_generic...
+apply generic_format_bpow.
+unfold FLX_exp; omega.
+apply Es.
+apply Rlt_le_trans with (1:=H).
+apply bpow_le.
+omega.
+now apply Rlt_le.
+Qed.
+
+Lemma sqrt_error_N_FLX_aux1 x (Fx : format x) (Px : (0 < x)%R) :
+ exists (mu : R) (e : Z), (format mu /\ x = mu * bpow (2 * e) :> R
+ /\ 1 <= mu < bpow 2)%R.
+Proof.
+set (e := ((mag beta x - 1) / 2)%Z).
+set (mu := (x * bpow (-2 * e)%Z)%R).
+assert (Hbe : (bpow (-2 * e) * bpow (2 * e) = 1)%R).
+{ now rewrite <- bpow_plus; case e; simpl; [reflexivity| |]; intro p;
+ rewrite Z.pos_sub_diag. }
+assert (Fmu : format mu); [now apply mult_bpow_exact_FLX|].
+exists mu, e; split; [exact Fmu|split; [|split]].
+{ set (e2 := (2 * e)%Z); simpl; unfold mu; rewrite Rmult_assoc.
+ now unfold e2; rewrite Hbe, Rmult_1_r. }
+{ apply (Rmult_le_reg_r (bpow (2 * e))).
+ { apply bpow_gt_0. }
+ rewrite Rmult_1_l; set (e2 := (2 * e)%Z); simpl; unfold mu.
+ unfold e2; rewrite Rmult_assoc, Hbe, Rmult_1_r.
+ apply (Rle_trans _ (bpow (mag beta x - 1))).
+ { now apply bpow_le; unfold e; apply Z_mult_div_ge. }
+ set (l := mag _ _); rewrite <- (Rabs_pos_eq _ (Rlt_le _ _ Px)).
+ unfold l; apply bpow_mag_le.
+ intro Hx; revert Px; rewrite Hx; apply Rlt_irrefl. }
+simpl; unfold mu; change (IZR _) with (bpow 2).
+apply (Rmult_lt_reg_r (bpow (2 * e))); [now apply bpow_gt_0|].
+rewrite Rmult_assoc, Hbe, Rmult_1_r.
+apply (Rlt_le_trans _ (bpow (mag beta x))).
+{ rewrite <- (Rabs_pos_eq _ (Rlt_le _ _ Px)) at 1; apply bpow_mag_gt. }
+rewrite <- bpow_plus; apply bpow_le; unfold e; set (mxm1 := (_ - 1)%Z).
+replace (_ * _)%Z with (2 * (mxm1 / 2) + mxm1 mod 2 - mxm1 mod 2)%Z by ring.
+rewrite <- Z.div_mod; [|now simpl].
+apply (Zplus_le_reg_r _ _ (mxm1 mod 2 - mag beta x)%Z).
+unfold mxm1; destruct (Z.mod_bound_or (mag beta x - 1) 2); omega.
+Qed.
+
+Notation u_ro := (u_ro beta prec).
+
+Lemma sqrt_error_N_FLX_aux2 x (Fx : format x) :
+ (1 <= x)%R ->
+ (x = 1 :> R \/ x = 1 + 2 * u_ro :> R \/ 1 + 4 * u_ro <= x)%R.
+Proof.
+intro HxGe1.
+assert (Pu_ro : (0 <= u_ro)%R); [apply Rmult_le_pos; [lra|apply bpow_ge_0]|].
+destruct (Rle_or_lt x 1) as [HxLe1|HxGt1]; [now left; apply Rle_antisym|right].
+assert (F1 : format 1); [now apply generic_format_FLX_1|].
+assert (H2eps : (2 * u_ro = bpow (-prec + 1))%R).
+{ unfold u_ro; rewrite bpow_plus; field. }
+assert (HmuGe1p2eps : (1 + 2 * u_ro <= x)%R).
+{ rewrite H2eps, <- succ_FLX_1.
+ now apply succ_le_lt; [now apply FLX_exp_valid| | |]. }
+destruct (Rle_or_lt x (1 + 2 * u_ro)) as [HxLe1p2eps|HxGt1p2eps];
+ [now left; apply Rle_antisym|right].
+assert (Hulp1p2eps : (ulp beta (FLX_exp prec) (1 + 2 * u_ro) = 2 * u_ro)%R).
+{ destruct (ulp_succ_pos _ _ _ F1 Rlt_0_1) as [Hsucc|Hsucc].
+ { now rewrite H2eps, <- succ_FLX_1, <- ulp_FLX_1. }
+ exfalso; revert Hsucc; apply Rlt_not_eq.
+ rewrite succ_FLX_1, mag_1, bpow_1, <- H2eps; simpl.
+ apply (Rlt_le_trans _ 2); [apply Rplus_lt_compat_l|].
+ { unfold u_ro; rewrite <-Rmult_assoc, Rinv_r, Rmult_1_l; [|lra].
+ change R1 with (bpow 0); apply bpow_lt; omega. }
+ apply IZR_le, Zle_bool_imp_le, radix_prop. }
+assert (Hsucc1p2eps :
+ (succ beta (FLX_exp prec) (1 + 2 * u_ro) = 1 + 4 * u_ro)%R).
+{ unfold succ; rewrite Rle_bool_true; [rewrite Hulp1p2eps; ring|].
+ apply Rplus_le_le_0_compat; lra. }
+rewrite <- Hsucc1p2eps.
+apply succ_le_lt; [now apply FLX_exp_valid| |exact Fx|now simpl].
+rewrite H2eps, <- succ_FLX_1.
+now apply generic_format_succ; [apply FLX_exp_valid|].
+Qed.
+
+Lemma sqrt_error_N_FLX_aux3 :
+ (u_ro / sqrt (1 + 4 * u_ro) <= 1 - 1 / sqrt (1 + 2 * u_ro))%R.
+Proof.
+assert (Pu_ro : (0 <= u_ro)%R); [apply Rmult_le_pos; [lra|apply bpow_ge_0]|].
+unfold Rdiv; apply (Rplus_le_reg_r (/ sqrt (1 + 2 * u_ro))); ring_simplify.
+apply (Rmult_le_reg_r (sqrt (1 + 4 * u_ro) * sqrt (1 + 2 * u_ro))).
+{ apply Rmult_lt_0_compat; apply sqrt_lt_R0; lra. }
+field_simplify; [|split; apply Rgt_not_eq, Rlt_gt, sqrt_lt_R0; lra].
+unfold Rdiv; rewrite Rinv_1, !Rmult_1_r.
+apply Rsqr_incr_0_var; [|now apply Rmult_le_pos; apply sqrt_pos].
+rewrite <-sqrt_mult; [|lra|lra].
+rewrite Rsqr_sqrt; [|apply Rmult_le_pos; lra].
+unfold Rsqr; ring_simplify; unfold pow; rewrite !Rmult_1_r.
+rewrite !sqrt_def; [|lra|lra].
+apply (Rplus_le_reg_r (-u_ro * u_ro - 1 -4 * u_ro - 2 * u_ro ^ 3)).
+ring_simplify; apply Rsqr_incr_0_var.
+{ unfold Rsqr; ring_simplify.
+ unfold pow; rewrite !Rmult_1_r, !sqrt_def; [|lra|lra].
+ apply (Rplus_le_reg_r (-32 * u_ro ^ 4 - 24 * u_ro ^ 3 - 4 * u_ro ^ 2)).
+ ring_simplify.
+ replace (_ + _)%R
+ with (((4 * u_ro ^ 2 - 28 * u_ro + 9) * u_ro + 4) * u_ro ^ 3)%R by ring.
+ apply Rmult_le_pos; [|now apply pow_le].
+ assert (Heps_le_half : (u_ro <= 1 / 2)%R).
+ { unfold u_ro, Rdiv; rewrite Rmult_comm; apply Rmult_le_compat_r; [lra|].
+ change 1%R with (bpow 0); apply bpow_le; omega. }
+ apply (Rle_trans _ (-8 * u_ro + 4)); [lra|].
+ apply Rplus_le_compat_r, Rmult_le_compat_r; [apply Pu_ro|].
+ now assert (H : (0 <= u_ro ^ 2)%R); [apply pow2_ge_0|lra]. }
+assert (H : (u_ro ^ 3 <= u_ro ^ 2)%R).
+{ unfold pow; rewrite <-!Rmult_assoc, Rmult_1_r.
+ apply Rmult_le_compat_l; [now apply Rmult_le_pos; apply Pu_ro|].
+ now apply Rlt_le, u_ro_lt_1. }
+now assert (H' : (0 <= u_ro ^ 2)%R); [apply pow2_ge_0|lra].
+Qed.
+
+Lemma om1ds1p2u_ro_pos : (0 <= 1 - 1 / sqrt (1 + 2 * u_ro))%R.
+Proof.
+unfold Rdiv; rewrite Rmult_1_l, <-Rinv_1 at 1.
+apply Rle_0_minus, Rinv_le; [lra|].
+rewrite <- sqrt_1 at 1; apply sqrt_le_1_alt.
+assert (H := u_ro_pos beta prec); lra.
+Qed.
+
+Lemma om1ds1p2u_ro_le_u_rod1pu_ro :
+ (1 - 1 / sqrt (1 + 2 * u_ro) <= u_ro / (1 + u_ro))%R.
+Proof.
+assert (Pu_ro := u_ro_pos beta prec).
+apply (Rmult_le_reg_r (sqrt (1 + 2 * u_ro) * (1 + u_ro))).
+{ apply Rmult_lt_0_compat; [apply sqrt_lt_R0|]; lra. }
+field_simplify; [|lra|intro H; apply sqrt_eq_0 in H; lra].
+unfold Rdiv, Rminus; rewrite Rinv_1, !Rmult_1_r, !Rplus_assoc.
+rewrite <-(Rplus_0_r (sqrt _ * _)) at 2; apply Rplus_le_compat_l.
+apply (Rplus_le_reg_r (1 + u_ro)); ring_simplify.
+rewrite <-(sqrt_square (_ + 1)); [|lra]; apply sqrt_le_1_alt.
+assert (H : (0 <= u_ro * u_ro)%R); [apply Rmult_le_pos|]; lra.
+Qed.
+
+Lemma s1p2u_rom1_pos : (0 <= sqrt (1 + 2 * u_ro) - 1)%R.
+apply (Rplus_le_reg_r 1); ring_simplify.
+rewrite <-sqrt_1 at 1; apply sqrt_le_1_alt.
+assert (H := u_ro_pos beta prec); lra.
+Qed.
+
+Theorem sqrt_error_N_FLX x (Fx : format x) :
+ (Rabs (round beta (FLX_exp prec) (Znearest choice) (sqrt x) - sqrt x)
+ <= (1 - 1 / sqrt (1 + 2 * u_ro)) * Rabs (sqrt x))%R.
+Proof.
+assert (Peps := u_ro_pos beta prec).
+assert (Peps' : (0 < u_ro)%R).
+{ unfold u_ro; apply Rmult_lt_0_compat; [lra|apply bpow_gt_0]. }
+assert (Pb := om1ds1p2u_ro_pos).
+assert (Pb' := s1p2u_rom1_pos).
+destruct (Rle_or_lt x 0) as [Nx|Px].
+{ rewrite (sqrt_neg _ Nx), round_0, Rabs_R0, Rmult_0_r; [|apply valid_rnd_N].
+ now unfold Rminus; rewrite Rplus_0_l, Rabs_Ropp, Rabs_R0; right. }
+destruct (sqrt_error_N_FLX_aux1 _ Fx Px)
+ as (mu, (e, (Fmu, (Hmu, (HmuGe1, HmuLtsqradix))))).
+pose (t := sqrt x).
+set (rt := round _ _ _ _).
+assert (Ht : (t = sqrt mu * bpow e)%R).
+{ unfold t; rewrite Hmu, sqrt_mult_alt; [|now apply (Rle_trans _ _ _ Rle_0_1)].
+ now rewrite sqrt_bpow. }
+destruct (sqrt_error_N_FLX_aux2 _ Fmu HmuGe1) as [Hmu'|[Hmu'|Hmu']].
+{ unfold rt; fold t; rewrite Ht, Hmu', sqrt_1, Rmult_1_l.
+ rewrite round_generic; [|now apply valid_rnd_N|].
+ { rewrite Rminus_diag_eq, Rabs_R0; [|now simpl].
+ now apply Rmult_le_pos; [|apply Rabs_pos]. }
+ apply generic_format_bpow'; [now apply FLX_exp_valid|].
+ unfold FLX_exp; omega. }
+{ assert (Hsqrtmu : (1 <= sqrt mu < 1 + u_ro)%R); [rewrite Hmu'; split|].
+ { rewrite <- sqrt_1 at 1; apply sqrt_le_1_alt; lra. }
+ { rewrite <- sqrt_square; [|lra]; apply sqrt_lt_1_alt; split; [lra|].
+ ring_simplify; assert (0 < u_ro ^ 2)%R; [apply pow_lt|]; lra. }
+ assert (Fbpowe : generic_format beta (FLX_exp prec) (bpow e)).
+ { apply generic_format_bpow; unfold FLX_exp; omega. }
+ assert (Hrt : rt = bpow e :> R).
+ { unfold rt; fold t; rewrite Ht; simpl; apply Rle_antisym.
+ { apply round_N_le_midp; [now apply FLX_exp_valid|exact Fbpowe|].
+ apply (Rlt_le_trans _ ((1 + u_ro) * bpow e)).
+ { now apply Rmult_lt_compat_r; [apply bpow_gt_0|]. }
+ unfold succ; rewrite Rle_bool_true; [|now apply bpow_ge_0].
+ rewrite ulp_bpow; unfold FLX_exp.
+ unfold Z.sub, u_ro; rewrite !bpow_plus; right; field. }
+ apply round_ge_generic;
+ [now apply FLX_exp_valid|now apply valid_rnd_N|exact Fbpowe|].
+ rewrite <- (Rmult_1_l (bpow _)) at 1.
+ now apply Rmult_le_compat_r; [apply bpow_ge_0|]. }
+ fold t; rewrite Hrt, Ht, Hmu', <-(Rabs_pos_eq _ Pb), <-Rabs_mult.
+ rewrite Rabs_minus_sym; right; f_equal; field; lra. }
+assert (Hsqrtmu : (1 + u_ro < sqrt mu)%R).
+{ apply (Rlt_le_trans _ (sqrt (1 + 4 * u_ro))); [|now apply sqrt_le_1_alt].
+ assert (P1peps : (0 <= 1 + u_ro)%R)
+ by now apply Rplus_le_le_0_compat; [lra|apply Peps].
+ rewrite <- (sqrt_square (1 + u_ro)); [|lra].
+ apply sqrt_lt_1_alt; split; [now apply Rmult_le_pos|].
+ apply (Rplus_lt_reg_r (-1 - 2 * u_ro)); ring_simplify; simpl.
+ rewrite Rmult_1_r; apply Rmult_lt_compat_r; [apply Peps'|].
+ now apply (Rlt_le_trans _ 1); [apply u_ro_lt_1|lra]. }
+assert (Hulpt : (ulp beta (FLX_exp prec) t = 2 * u_ro * bpow e)%R).
+{ unfold ulp; rewrite Req_bool_false; [|apply Rgt_not_eq, Rlt_gt].
+ { unfold u_ro; rewrite <-Rmult_assoc, Rinv_r, Rmult_1_l, <-bpow_plus; [|lra].
+ f_equal; unfold cexp, FLX_exp.
+ assert (Hmagt : (mag beta t = 1 + e :> Z)%Z).
+ { apply mag_unique.
+ unfold t; rewrite (Rabs_pos_eq _ (Rlt_le _ _ (sqrt_lt_R0 _ Px))).
+ fold t; split.
+ { rewrite Ht; replace (_ - _)%Z with e by ring.
+ rewrite <- (Rmult_1_l (bpow _)) at 1; apply Rmult_le_compat_r.
+ { apply bpow_ge_0. }
+ now rewrite <- sqrt_1; apply sqrt_le_1_alt. }
+ rewrite bpow_plus, bpow_1, Ht; simpl.
+ apply Rmult_lt_compat_r; [now apply bpow_gt_0|].
+ rewrite <- sqrt_square.
+ { apply sqrt_lt_1_alt; split; [lra|].
+ apply (Rlt_le_trans _ _ _ HmuLtsqradix); right.
+ now unfold bpow, Z.pow_pos; simpl; rewrite Zmult_1_r, mult_IZR. }
+ apply IZR_le, (Z.le_trans _ 2), Zle_bool_imp_le, radix_prop; omega. }
+ rewrite Hmagt; ring. }
+ rewrite Ht; apply Rmult_lt_0_compat; [|now apply bpow_gt_0].
+ now apply (Rlt_le_trans _ 1); [lra|rewrite <- sqrt_1; apply sqrt_le_1_alt]. }
+assert (Pt : (0 < t)%R).
+{ rewrite Ht; apply Rmult_lt_0_compat; [lra|apply bpow_gt_0]. }
+assert (H : (Rabs ((rt - sqrt x) / sqrt x)
+ <= 1 - 1 / sqrt (1 + 2 * u_ro))%R).
+{ unfold Rdiv; rewrite Rabs_mult, (Rabs_pos_eq (/ _));
+ [|now left; apply Rinv_0_lt_compat].
+ apply (Rle_trans _ ((u_ro * bpow e) / t)).
+ { unfold Rdiv; apply Rmult_le_compat_r; [now left; apply Rinv_0_lt_compat|].
+ apply (Rle_trans _ _ _ (error_le_half_ulp _ _ _ _)).
+ fold t; rewrite Hulpt; right; field. }
+ apply (Rle_trans _ (u_ro / sqrt (1 + 4 * u_ro))).
+ { apply (Rle_trans _ (u_ro * bpow e / (sqrt (1 + 4 * u_ro) * bpow e))).
+ { unfold Rdiv; apply Rmult_le_compat_l;
+ [now apply Rmult_le_pos; [apply Peps|apply bpow_ge_0]|].
+ apply Rinv_le.
+ { apply Rmult_lt_0_compat; [apply sqrt_lt_R0; lra|apply bpow_gt_0]. }
+ now rewrite Ht; apply Rmult_le_compat_r;
+ [apply bpow_ge_0|apply sqrt_le_1_alt]. }
+ right; field; split; apply Rgt_not_eq, Rlt_gt;
+ [apply sqrt_lt_R0; lra|apply bpow_gt_0]. }
+ apply sqrt_error_N_FLX_aux3. }
+revert H; unfold Rdiv; rewrite Rabs_mult, Rabs_Rinv; [|fold t; lra]; intro H.
+apply (Rmult_le_reg_r (/ Rabs (sqrt x)));
+ [apply Rinv_0_lt_compat, Rabs_pos_lt; fold t; lra|].
+apply (Rle_trans _ _ _ H); right; field; split; [apply Rabs_no_R0;fold t|]; lra.
+Qed.
+
+Theorem sqrt_error_N_FLX_ex x (Fx : format x) :
+ exists eps,
+ (Rabs eps <= 1 - 1 / sqrt (1 + 2 * u_ro))%R /\
+ round beta (FLX_exp prec) (Znearest choice) (sqrt x)
+ = (sqrt x * (1 + eps))%R.
+Proof.
+now apply relative_error_le_conversion;
+ [apply valid_rnd_N|apply om1ds1p2u_ro_pos|apply sqrt_error_N_FLX].
+Qed.
+
+Lemma sqrt_error_N_round_ex_derive :
+ forall x rx,
+ (exists eps,
+ (Rabs eps <= 1 - 1 / sqrt (1 + 2 * u_ro))%R /\ rx = (x * (1 + eps))%R) ->
+ exists eps,
+ (Rabs eps <= sqrt (1 + 2 * u_ro) - 1)%R /\ x = (rx * (1 + eps))%R.
+Proof.
+intros x rx (d, (Bd, Hd)).
+assert (H := Rabs_le_inv _ _ Bd).
+assert (H' := om1ds1p2u_ro_le_u_rod1pu_ro).
+assert (H'' := u_rod1pu_ro_le_u_ro beta prec).
+assert (H''' := u_ro_lt_1 beta prec prec_gt_0_).
+assert (Hpos := s1p2u_rom1_pos).
+destruct (Req_dec rx 0) as [Zfx|Nzfx].
+{ exists 0%R; split; [now rewrite Rabs_R0|].
+ rewrite Rplus_0_r, Rmult_1_r, Zfx; rewrite Zfx in Hd.
+ destruct (Rmult_integral _ _ (sym_eq Hd)); lra. }
+destruct (Req_dec x 0) as [Zx|Nzx].
+{ now exfalso; revert Hd; rewrite Zx; rewrite Rmult_0_l. }
+set (d' := ((x - rx) / rx)%R).
+assert (Hd' : (Rabs d' <= sqrt (1 + 2 * u_ro) - 1)%R).
+{ unfold d'; rewrite Hd.
+ replace (_ / _)%R with (- d / (1 + d))%R; [|now field; split; lra].
+ unfold Rdiv; rewrite Rabs_mult, Rabs_Ropp.
+ rewrite (Rabs_pos_eq (/ _)); [|apply Rlt_le, Rinv_0_lt_compat; lra].
+ apply (Rmult_le_reg_r (1 + d)); [lra|].
+ rewrite Rmult_assoc, Rinv_l, Rmult_1_r; [|lra].
+ apply (Rle_trans _ _ _ Bd).
+ apply (Rle_trans _ ((sqrt (1 + 2 * u_ro) - 1) * (1/sqrt (1 + 2 * u_ro))));
+ [right; field|apply Rmult_le_compat_l]; lra. }
+now exists d'; split; [exact Hd'|]; unfold d'; field.
+Qed.
+
+(** sqrt(1 + 2 u_ro) - 1 <= u_ro *)
+Theorem sqrt_error_N_FLX_round_ex :
+ forall x,
+ format x ->
+ exists eps,
+ (Rabs eps <= sqrt (1 + 2 * u_ro) - 1)%R /\
+ sqrt x = (round beta (FLX_exp prec) (Znearest choice) (sqrt x) * (1 + eps))%R.
+Proof.
+now intros x Fx; apply sqrt_error_N_round_ex_derive, sqrt_error_N_FLX_ex.
+Qed.
+
+Variable emin : Z.
+Hypothesis Hemin : (emin <= 2 * (1 - prec))%Z.
+
+Theorem sqrt_error_N_FLT_ex :
+ forall x,
+ generic_format beta (FLT_exp emin prec) x ->
+ exists eps,
+ (Rabs eps <= 1 - 1 / sqrt (1 + 2 * u_ro))%R /\
+ round beta (FLT_exp emin prec) (Znearest choice) (sqrt x)
+ = (sqrt x * (1 + eps))%R.
+Proof.
+intros x Fx.
+assert (Heps := u_ro_pos).
+assert (Pb := om1ds1p2u_ro_pos).
+destruct (Rle_or_lt x 0) as [Nx|Px].
+{ exists 0%R; split; [now rewrite Rabs_R0|].
+ now rewrite (sqrt_neg x Nx), round_0, Rmult_0_l; [|apply valid_rnd_N]. }
+assert (Fx' := generic_format_FLX_FLT _ _ _ _ Fx).
+destruct (sqrt_error_N_FLX_ex _ Fx') as (d, (Bd, Hd)).
+exists d; split; [exact Bd|]; rewrite <-Hd; apply round_FLT_FLX.
+apply (Rle_trans _ (bpow (emin / 2)%Z)).
+{ apply bpow_le, Z.div_le_lower_bound; lia. }
+apply (Rle_trans _ _ _ (sqrt_bpow_ge _ _)).
+rewrite Rabs_pos_eq; [|now apply sqrt_pos]; apply sqrt_le_1_alt.
+revert Fx; apply generic_format_ge_bpow; [|exact Px].
+intro e; unfold FLT_exp; apply Z.le_max_r.
+Qed.
+
+(** sqrt(1 + 2 u_ro) - 1 <= u_ro *)
+Theorem sqrt_error_N_FLT_round_ex :
+ forall x,
+ generic_format beta (FLT_exp emin prec) x ->
+ exists eps,
+ (Rabs eps <= sqrt (1 + 2 * u_ro) - 1)%R /\
+ sqrt x
+ = (round beta (FLT_exp emin prec) (Znearest choice) (sqrt x) * (1 + eps))%R.
+Proof.
+now intros x Fx; apply sqrt_error_N_round_ex_derive, sqrt_error_N_FLT_ex.
+Qed.
+
+End Fprop_divsqrt_error.
+
+Section format_REM_aux.
+
+Variable beta : radix.
+Notation bpow e := (bpow beta e).
+
+Variable fexp : Z -> Z.
+Context { valid_exp : Valid_exp fexp }.
+Context { monotone_exp : Monotone_exp fexp }.
+
+Variable rnd : R -> Z.
+Context { valid_rnd : Valid_rnd rnd }.
+
+Notation format := (generic_format beta fexp).
+
+Lemma format_REM_aux:
+ forall x y : R,
+ format x -> format y -> (0 <= x)%R -> (0 < y)%R ->
+ ((0 < x/y < /2)%R -> rnd (x/y) = 0%Z) ->
+ format (x - IZR (rnd (x/y))*y).
+Proof with auto with typeclass_instances.
+intros x y Fx Fy Hx Hy rnd_small.
+pose (n:=rnd (x / y)).
+assert (Hn:(IZR n = round beta (FIX_exp 0) rnd (x/y))%R).
+unfold round, FIX_exp, cexp, scaled_mantissa, F2R; simpl.
+now rewrite 2!Rmult_1_r.
+assert (H:(0 <= n)%Z).
+apply le_IZR; rewrite Hn; simpl.
+apply Rle_trans with (round beta (FIX_exp 0) rnd 0).
+right; apply sym_eq, round_0...
+apply round_le...
+apply Fourier_util.Rle_mult_inv_pos; assumption.
+case (Zle_lt_or_eq 0 n); try exact H.
+clear H; intros H.
+case (Zle_lt_or_eq 1 n).
+omega.
+clear H; intros H.
+set (ex := cexp beta fexp x).
+set (ey := cexp beta fexp y).
+set (mx := Ztrunc (scaled_mantissa beta fexp x)).
+set (my := Ztrunc (scaled_mantissa beta fexp y)).
+case (Zle_or_lt ey ex); intros Hexy.
+(* ey <= ex *)
+assert (H0:(x-IZR n *y = F2R (Float beta (mx*beta^(ex-ey) - n*my) ey))%R).
+unfold Rminus; rewrite Rplus_comm.
+replace (IZR n) with (F2R (Float beta n 0)).
+rewrite Fx, Fy.
+fold mx my ex ey.
+rewrite <- F2R_mult.
+rewrite <- F2R_opp.
+rewrite <- F2R_plus.
+unfold Fplus. simpl.
+rewrite Zle_imp_le_bool with (1 := Hexy).
+f_equal; f_equal; ring.
+unfold F2R; simpl; ring.
+fold n; rewrite H0.
+apply generic_format_F2R.
+rewrite <- H0; intros H3.
+apply monotone_exp.
+apply mag_le_abs.
+rewrite H0; apply F2R_neq_0; easy.
+apply Rmult_le_reg_l with (/Rabs y)%R.
+apply Rinv_0_lt_compat.
+apply Rabs_pos_lt.
+now apply Rgt_not_eq.
+rewrite Rinv_l.
+2: apply Rgt_not_eq, Rabs_pos_lt.
+2: now apply Rgt_not_eq.
+rewrite <- Rabs_Rinv.
+2: now apply Rgt_not_eq.
+rewrite <- Rabs_mult.
+replace (/y * (x - IZR n *y))%R with (-(IZR n - x/y))%R.
+rewrite Rabs_Ropp.
+rewrite Hn.
+apply Rle_trans with (1:= error_le_ulp beta (FIX_exp 0) _ _).
+rewrite ulp_FIX.
+simpl; apply Rle_refl.
+field.
+now apply Rgt_not_eq.
+(* ex < ey: impossible as 1 < n *)
+absurd (1 < n)%Z; try easy.
+apply Zle_not_lt.
+apply le_IZR; simpl; rewrite Hn.
+apply round_le_generic...
+apply generic_format_FIX.
+exists (Float beta 1 0); try easy.
+unfold F2R; simpl; ring.
+apply Rmult_le_reg_r with y; try easy.
+unfold Rdiv; rewrite Rmult_assoc.
+rewrite Rinv_l, Rmult_1_r, Rmult_1_l.
+2: now apply Rgt_not_eq.
+assert (mag beta x < mag beta y)%Z.
+case (Zle_or_lt (mag beta y) (mag beta x)); try easy.
+intros J; apply monotone_exp in J; clear -J Hexy.
+unfold ex, ey, cexp in Hexy; omega.
+left; apply lt_mag with beta; easy.
+(* n = 1 -> Sterbenz + rnd_small *)
+intros Hn'; fold n; rewrite <- Hn'.
+rewrite Rmult_1_l.
+case Hx; intros Hx'.
+assert (J:(0 < x/y)%R).
+apply Fourier_util.Rlt_mult_inv_pos; assumption.
+apply sterbenz...
+assert (H0:(Rabs (1 - x/y) < 1)%R).
+rewrite Hn', Hn.
+apply Rlt_le_trans with (ulp beta (FIX_exp 0) (round beta (FIX_exp 0) rnd (x / y)))%R.
+apply error_lt_ulp_round...
+now apply Rgt_not_eq.
+rewrite ulp_FIX.
+rewrite <- Hn, <- Hn'.
+apply Rle_refl.
+apply Rabs_lt_inv in H0.
+split; apply Rmult_le_reg_l with (/y)%R; try now apply Rinv_0_lt_compat.
+unfold Rdiv; rewrite <- Rmult_assoc.
+rewrite Rinv_l.
+2: now apply Rgt_not_eq.
+rewrite Rmult_1_l, Rmult_comm; fold (x/y)%R.
+case (Rle_or_lt (/2) (x/y)); try easy.
+intros K.
+elim Zlt_not_le with (1 := H).
+apply Zeq_le.
+apply rnd_small.
+now split.
+apply Ropp_le_cancel; apply Rplus_le_reg_l with 1%R.
+apply Rle_trans with (1-x/y)%R.
+2: right; unfold Rdiv; ring.
+left; apply Rle_lt_trans with (2:=proj1 H0).
+right; field.
+now apply Rgt_not_eq.
+rewrite <- Hx', Rminus_0_l.
+now apply generic_format_opp.
+(* n = 0 *)
+clear H; intros H; fold n; rewrite <- H.
+now rewrite Rmult_0_l, Rminus_0_r.
+Qed.
+
+End format_REM_aux.
+
+Section format_REM.
+
+Variable beta : radix.
+Notation bpow e := (bpow beta e).
+
+Variable fexp : Z -> Z.
+Context { valid_exp : Valid_exp fexp }.
+Context { monotone_exp : Monotone_exp fexp }.
+
+Notation format := (generic_format beta fexp).
+
+Theorem format_REM :
+ forall rnd : R -> Z, Valid_rnd rnd ->
+ forall x y : R,
+ ((Rabs (x/y) < /2)%R -> rnd (x/y)%R = 0%Z) ->
+ format x -> format y ->
+ format (x - IZR (rnd (x/y)%R) * y).
+Proof with auto with typeclass_instances.
+(* assume 0 < y *)
+assert (H: forall rnd : R -> Z, Valid_rnd rnd ->
+ forall x y : R,
+ ((Rabs (x/y) < /2)%R -> rnd (x/y)%R = 0%Z) ->
+ format x -> format y -> (0 < y)%R ->
+ format (x - IZR (rnd (x/y)%R) * y)).
+intros rnd valid_rnd x y Hrnd Fx Fy Hy.
+case (Rle_or_lt 0 x); intros Hx.
+apply format_REM_aux; try easy.
+intros K.
+apply Hrnd.
+rewrite Rabs_pos_eq.
+apply K.
+apply Rlt_le, K.
+replace (x - IZR (rnd (x/y)) * y)%R with
+ (- (-x - IZR (Zrnd_opp rnd (-x/y)) * y))%R.
+apply generic_format_opp.
+apply format_REM_aux; try easy...
+now apply generic_format_opp.
+apply Ropp_le_cancel; rewrite Ropp_0, Ropp_involutive; now left.
+replace (- x / y)%R with (- (x/y))%R by (unfold Rdiv; ring).
+intros K.
+unfold Zrnd_opp.
+rewrite Ropp_involutive, Hrnd.
+easy.
+rewrite Rabs_left.
+apply K.
+apply Ropp_lt_cancel.
+now rewrite Ropp_0.
+unfold Zrnd_opp.
+replace (- (- x / y))%R with (x / y)%R by (unfold Rdiv; ring).
+rewrite opp_IZR.
+ring.
+(* *)
+intros rnd valid_rnd x y Hrnd Fx Fy.
+case (Rle_or_lt 0 y); intros Hy.
+destruct Hy as [Hy|Hy].
+now apply H.
+now rewrite <- Hy, Rmult_0_r, Rminus_0_r.
+replace (IZR (rnd (x/y)) * y)%R with
+ (IZR ((Zrnd_opp rnd) ((x / -y))) * -y)%R.
+apply H; try easy...
+replace (x / - y)%R with (- (x/y))%R.
+intros K.
+unfold Zrnd_opp.
+rewrite Ropp_involutive, Hrnd.
+easy.
+now rewrite <- Rabs_Ropp.
+field; now apply Rlt_not_eq.
+now apply generic_format_opp.
+apply Ropp_lt_cancel; now rewrite Ropp_0, Ropp_involutive.
+unfold Zrnd_opp.
+replace (- (x / - y))%R with (x/y)%R.
+rewrite opp_IZR.
+ring.
+field; now apply Rlt_not_eq.
+Qed.
+
+Theorem format_REM_ZR:
+ forall x y : R,
+ format x -> format y ->
+ format (x - IZR (Ztrunc (x/y)) * y).
+Proof with auto with typeclass_instances.
+intros x y Fx Fy.
+apply format_REM; try easy...
+intros K.
+apply Z.abs_0_iff.
+rewrite <- Ztrunc_abs.
+rewrite Ztrunc_floor by apply Rabs_pos.
+apply Zle_antisym.
+replace 0%Z with (Zfloor (/2)).
+apply Zfloor_le.
+now apply Rlt_le.
+apply Zfloor_imp.
+simpl ; lra.
+apply Zfloor_lub.
+apply Rabs_pos.
+Qed.
+
+Theorem format_REM_N :
+ forall choice,
+ forall x y : R,
+ format x -> format y ->
+ format (x - IZR (Znearest choice (x/y)) * y).
+Proof with auto with typeclass_instances.
+intros choice x y Fx Fy.
+apply format_REM; try easy...
+intros K.
+apply Znearest_imp.
+now rewrite Rminus_0_r.
+Qed.
+
+End format_REM.
diff --git a/flocq/Prop/Double_rounding.v b/flocq/Prop/Double_rounding.v
new file mode 100644
index 00000000..055409bb
--- /dev/null
+++ b/flocq/Prop/Double_rounding.v
@@ -0,0 +1,4545 @@
+(**
+This file is part of the Flocq formalization of floating-point
+arithmetic in Coq: http://flocq.gforge.inria.fr/
+
+Copyright (C) 2014-2018 Sylvie Boldo
+#<br />#
+Copyright (C) 2014-2018 Guillaume Melquiond
+#<br />#
+Copyright (C) 2014-2018 Pierre Roux
+
+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.
+*)
+
+(** * Conditions for innocuous double rounding. *)
+
+Require Import Psatz.
+Require Import Raux Defs Generic_fmt Operations Ulp FLX FLT FTZ.
+
+Open Scope R_scope.
+
+Section Double_round.
+
+Variable beta : radix.
+Notation bpow e := (bpow beta e).
+Notation mag x := (mag beta x).
+
+Definition round_round_eq fexp1 fexp2 choice1 choice2 x :=
+ round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x)
+ = round beta fexp1 (Znearest choice1) x.
+
+(** A little tactic to simplify terms of the form [bpow a * bpow b]. *)
+Ltac bpow_simplify :=
+ (* bpow ex * bpow ey ~~> bpow (ex + ey) *)
+ repeat
+ match goal with
+ | |- context [(Raux.bpow _ _ * Raux.bpow _ _)] =>
+ rewrite <- bpow_plus
+ | |- context [(?X1 * Raux.bpow _ _ * Raux.bpow _ _)] =>
+ rewrite (Rmult_assoc X1); rewrite <- bpow_plus
+ | |- context [(?X1 * (?X2 * Raux.bpow _ _) * Raux.bpow _ _)] =>
+ rewrite <- (Rmult_assoc X1 X2); rewrite (Rmult_assoc (X1 * X2));
+ rewrite <- bpow_plus
+ end;
+ (* ring_simplify arguments of bpow *)
+ repeat
+ match goal with
+ | |- context [(Raux.bpow _ ?X)] =>
+ progress ring_simplify X
+ end;
+ (* bpow 0 ~~> 1 *)
+ change (Raux.bpow _ 0) with 1;
+ repeat
+ match goal with
+ | |- context [(_ * 1)] =>
+ rewrite Rmult_1_r
+ end.
+
+Definition midp (fexp : Z -> Z) (x : R) :=
+ round beta fexp Zfloor x + / 2 * ulp beta fexp x.
+
+Definition midp' (fexp : Z -> Z) (x : R) :=
+ round beta fexp Zceil x - / 2 * ulp beta fexp x.
+
+Lemma round_round_lt_mid_further_place' :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ forall x,
+ 0 < x ->
+ (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z ->
+ x < bpow (mag x) - / 2 * ulp beta fexp2 x ->
+ x < midp fexp1 x - / 2 * ulp beta fexp2 x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x.
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 x Px Hf2f1 Hx1.
+unfold round_round_eq.
+set (x' := round beta fexp1 Zfloor x).
+intro Hx2'.
+assert (Hx2 : x - round beta fexp1 Zfloor x
+ < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)).
+{ now apply (Rplus_lt_reg_r (round beta fexp1 Zfloor x)); ring_simplify. }
+set (x'' := round beta fexp2 (Znearest choice2) x).
+assert (Hr1 : Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))).
+apply Rle_trans with (/ 2 * ulp beta fexp2 x).
+now unfold x''; apply error_le_half_ulp...
+rewrite ulp_neq_0;[now right|now apply Rgt_not_eq].
+assert (Pxx' : 0 <= x - x').
+{ apply Rle_0_minus.
+ apply round_DN_pt.
+ exact Vfexp1. }
+rewrite 2!ulp_neq_0 in Hx2; try (apply Rgt_not_eq; assumption).
+assert (Hr2 : Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))).
+{ replace (x'' - x') with (x'' - x + (x - x')) by ring.
+ apply (Rle_lt_trans _ _ _ (Rabs_triang _ _)).
+ replace (/ 2 * _) with (/ 2 * bpow (fexp2 (mag x))
+ + (/ 2 * (bpow (fexp1 (mag x))
+ - bpow (fexp2 (mag x))))) by ring.
+ apply Rplus_le_lt_compat.
+ - exact Hr1.
+ - now rewrite Rabs_right; [|now apply Rle_ge]; apply Hx2. }
+destruct (Req_dec x'' 0) as [Zx''|Nzx''].
+- (* x'' = 0 *)
+ rewrite Zx'' in Hr1 |- *.
+ rewrite round_0; [|now apply valid_rnd_N].
+ unfold round, F2R, scaled_mantissa, cexp; simpl.
+ rewrite (Znearest_imp _ _ 0); [now simpl; rewrite Rmult_0_l|].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x)))); [now apply bpow_gt_0|].
+ rewrite <- (Rabs_right (bpow (fexp1 _))) at 1;
+ [|now apply Rle_ge; apply bpow_ge_0].
+ rewrite <- Rabs_mult; rewrite Rmult_minus_distr_r.
+ rewrite Rmult_0_l.
+ bpow_simplify.
+ rewrite Rabs_minus_sym.
+ apply (Rle_lt_trans _ _ _ Hr1).
+ apply Rmult_lt_compat_l; [lra|].
+ apply bpow_lt.
+ omega.
+- (* x'' <> 0 *)
+ assert (Lx'' : mag x'' = mag x :> Z).
+ { apply Zle_antisym.
+ - apply mag_le_bpow; [exact Nzx''|].
+ replace x'' with (x'' - x + x) by ring.
+ apply (Rle_lt_trans _ _ _ (Rabs_triang _ _)).
+ replace (bpow _) with (/ 2 * bpow (fexp2 (mag x))
+ + (bpow (mag x)
+ - / 2 * bpow (fexp2 (mag x)))) by ring.
+ apply Rplus_le_lt_compat; [exact Hr1|].
+ rewrite ulp_neq_0 in Hx1;[idtac| now apply Rgt_not_eq].
+ now rewrite Rabs_right; [|apply Rle_ge; apply Rlt_le].
+ - unfold x'' in Nzx'' |- *.
+ now apply mag_round_ge; [|apply valid_rnd_N|]. }
+ unfold round, F2R, scaled_mantissa, cexp; simpl.
+ rewrite Lx''.
+ rewrite (Znearest_imp _ _ (Zfloor (scaled_mantissa beta fexp1 x))).
+ + rewrite (Znearest_imp _ _ (Zfloor (scaled_mantissa beta fexp1 x)));
+ [reflexivity|].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x)))); [now apply bpow_gt_0|].
+ rewrite <- (Rabs_right (bpow (fexp1 _))) at 1;
+ [|now apply Rle_ge; apply bpow_ge_0].
+ rewrite <- Rabs_mult.
+ rewrite Rmult_minus_distr_r.
+ fold x'.
+ bpow_simplify.
+ rewrite Rabs_right; [|now apply Rle_ge].
+ apply (Rlt_le_trans _ _ _ Hx2).
+ apply Rmult_le_compat_l; [lra|].
+ generalize (bpow_ge_0 beta (fexp2 (mag x))).
+ unfold ulp, cexp; lra.
+ + apply (Rmult_lt_reg_r (bpow (fexp1 (mag x)))); [now apply bpow_gt_0|].
+ rewrite <- (Rabs_right (bpow (fexp1 _))) at 1;
+ [|now apply Rle_ge; apply bpow_ge_0].
+ rewrite <- Rabs_mult.
+ rewrite Rmult_minus_distr_r.
+ fold x'.
+ now bpow_simplify.
+Qed.
+
+Lemma round_round_lt_mid_further_place :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ forall x,
+ 0 < x ->
+ (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z ->
+ (fexp1 (mag x) <= mag x)%Z ->
+ x < midp fexp1 x - / 2 * ulp beta fexp2 x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x.
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 x Px Hf2f1 Hf1.
+intro Hx2'.
+assert (Hx2 : x - round beta fexp1 Zfloor x
+ < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)).
+{ now apply (Rplus_lt_reg_r (round beta fexp1 Zfloor x)); ring_simplify. }
+revert Hx2.
+unfold round_round_eq.
+set (x' := round beta fexp1 Zfloor x).
+intro Hx2.
+assert (Pxx' : 0 <= x - x').
+{ apply Rle_0_minus.
+ apply round_DN_pt.
+ exact Vfexp1. }
+assert (x < bpow (mag x) - / 2 * bpow (fexp2 (mag x)));
+ [|apply round_round_lt_mid_further_place'; try assumption]...
+2: rewrite ulp_neq_0;[assumption|now apply Rgt_not_eq].
+destruct (Req_dec x' 0) as [Zx'|Nzx'].
+- (* x' = 0 *)
+ rewrite Zx' in Hx2; rewrite Rminus_0_r in Hx2.
+ apply (Rlt_le_trans _ _ _ Hx2).
+ rewrite Rmult_minus_distr_l.
+ rewrite 2!ulp_neq_0;[idtac|now apply Rgt_not_eq|now apply Rgt_not_eq].
+ apply Rplus_le_compat_r.
+ apply (Rmult_le_reg_r (bpow (- mag x))); [now apply bpow_gt_0|].
+ unfold ulp, cexp; bpow_simplify.
+ apply Rmult_le_reg_l with (1 := Rlt_0_2).
+ replace (2 * (/ 2 * _)) with (bpow (fexp1 (mag x) - mag x)) by field.
+ apply Rle_trans with 1; [|lra].
+ change 1 with (bpow 0); apply bpow_le.
+ omega.
+- (* x' <> 0 *)
+ assert (Px' : 0 < x').
+ { assert (0 <= x'); [|lra].
+ unfold x'.
+ apply (Rmult_le_reg_r (bpow (- fexp1 (mag x))));
+ [now apply bpow_gt_0|].
+ rewrite Rmult_0_l.
+ unfold round, F2R, cexp; simpl; bpow_simplify.
+ apply IZR_le.
+ apply Zfloor_lub.
+ rewrite <- (Rabs_right x); [|now apply Rle_ge; apply Rlt_le].
+ rewrite scaled_mantissa_abs.
+ apply Rabs_pos. }
+ assert (Hx' : x' <= bpow (mag x) - ulp beta fexp1 x).
+ { apply (Rplus_le_reg_r (ulp beta fexp1 x)); ring_simplify.
+ rewrite <- ulp_DN.
+ - change (round _ _ _ _) with x'.
+ apply id_p_ulp_le_bpow.
+ + exact Px'.
+ + change x' with (round beta fexp1 Zfloor x).
+ now apply generic_format_round; [|apply valid_rnd_DN].
+ + apply Rle_lt_trans with x.
+ * now apply round_DN_pt.
+ * rewrite <- (Rabs_right x) at 1; [|now apply Rle_ge; apply Rlt_le].
+ apply bpow_mag_gt.
+ - exact Vfexp1.
+ - now apply Rlt_le. }
+ fold (cexp beta fexp2 x); fold (ulp beta fexp2 x).
+ assert (/ 2 * ulp beta fexp1 x <= ulp beta fexp1 x).
+ rewrite <- (Rmult_1_l (ulp _ _ _)) at 2.
+ apply Rmult_le_compat_r; [|lra].
+ apply ulp_ge_0.
+ rewrite 2!ulp_neq_0 in Hx2;[|now apply Rgt_not_eq|now apply Rgt_not_eq].
+ rewrite ulp_neq_0 in Hx';[|now apply Rgt_not_eq].
+ rewrite ulp_neq_0 in H;[|now apply Rgt_not_eq].
+ lra.
+Qed.
+
+Lemma round_round_lt_mid_same_place :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 ->
+ forall (choice1 choice2 : Z -> bool),
+ forall x,
+ 0 < x ->
+ (fexp2 (mag x) = fexp1 (mag x))%Z ->
+ x < midp fexp1 x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x.
+Proof.
+intros fexp1 fexp2 Vfexp1 choice1 choice2 x Px Hf2f1.
+intro Hx'.
+assert (Hx : x - round beta fexp1 Zfloor x < / 2 * ulp beta fexp1 x).
+{ now apply (Rplus_lt_reg_r (round beta fexp1 Zfloor x)); ring_simplify. }
+revert Hx.
+unfold round_round_eq.
+set (x' := round beta fexp1 Zfloor x).
+intro Hx.
+assert (Pxx' : 0 <= x - x').
+{ apply Rle_0_minus.
+ apply round_DN_pt.
+ exact Vfexp1. }
+assert (H : Rabs (x * bpow (- fexp1 (mag x)) -
+ IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2).
+{ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x)))); [now apply bpow_gt_0|].
+ unfold scaled_mantissa, cexp in Hx.
+ rewrite <- (Rabs_right (bpow (fexp1 _))) at 1;
+ [|now apply Rle_ge; apply bpow_ge_0].
+ rewrite <- Rabs_mult.
+ rewrite Rmult_minus_distr_r.
+ bpow_simplify.
+ apply Rabs_lt.
+ change (IZR _ * _) with x'.
+ split.
+ - apply Rlt_le_trans with 0; [|exact Pxx'].
+ rewrite <- Ropp_0.
+ apply Ropp_lt_contravar.
+ rewrite <- (Rmult_0_r (/ 2)).
+ apply Rmult_lt_compat_l; [lra|].
+ apply bpow_gt_0.
+ - rewrite ulp_neq_0 in Hx;try apply Rgt_not_eq; assumption. }
+unfold round at 2.
+unfold F2R, scaled_mantissa, cexp; simpl.
+rewrite Hf2f1.
+rewrite (Znearest_imp _ _ (Zfloor (scaled_mantissa beta fexp1 x)) H).
+rewrite round_generic.
+ + unfold round, F2R, scaled_mantissa, cexp; simpl.
+ now rewrite (Znearest_imp _ _ (Zfloor (x * bpow (- fexp1 (mag x))))).
+ + now apply valid_rnd_N.
+ + fold (cexp beta fexp1 x).
+ change (IZR _ * bpow _) with (round beta fexp1 Zfloor x).
+ apply generic_format_round.
+ exact Vfexp1.
+ now apply valid_rnd_DN.
+Qed.
+
+Lemma round_round_lt_mid :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ forall x,
+ 0 < x ->
+ (fexp2 (mag x) <= fexp1 (mag x))%Z ->
+ (fexp1 (mag x) <= mag x)%Z ->
+ x < midp fexp1 x ->
+ ((fexp2 (mag x) <= fexp1 (mag x) - 1)%Z ->
+ x < midp fexp1 x - / 2 * ulp beta fexp2 x) ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x.
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 x Px Hf2f1 Hf1 Hx Hx'.
+destruct (Zle_or_lt (fexp1 (mag x)) (fexp2 (mag x))) as [Hf2'|Hf2'].
+- (* fexp1 (mag x) <= fexp2 (mag x) *)
+ assert (Hf2'' : (fexp2 (mag x) = fexp1 (mag x) :> Z)%Z); [omega|].
+ now apply round_round_lt_mid_same_place.
+- (* fexp2 (mag x) < fexp1 (mag x) *)
+ assert (Hf2'' : (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z); [omega|].
+ generalize (Hx' Hf2''); intro Hx''.
+ now apply round_round_lt_mid_further_place.
+Qed.
+
+Lemma round_round_gt_mid_further_place' :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ forall x,
+ 0 < x ->
+ (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z ->
+ round beta fexp2 (Znearest choice2) x < bpow (mag x) ->
+ midp' fexp1 x + / 2 * ulp beta fexp2 x < x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x.
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 x Px Hf2f1.
+intros Hx1 Hx2'.
+assert (Hx2 : round beta fexp1 Zceil x - x
+ < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)).
+{ apply (Rplus_lt_reg_r (- / 2 * ulp beta fexp1 x + x
+ + / 2 * ulp beta fexp2 x)); ring_simplify.
+ now unfold midp' in Hx2'. }
+revert Hx1 Hx2.
+unfold round_round_eq.
+set (x' := round beta fexp1 Zceil x).
+set (x'' := round beta fexp2 (Znearest choice2) x).
+intros Hx1 Hx2.
+assert (Hr1 : Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))).
+ apply Rle_trans with (/2* ulp beta fexp2 x).
+ now unfold x''; apply error_le_half_ulp...
+ rewrite ulp_neq_0;[now right|now apply Rgt_not_eq].
+assert (Px'x : 0 <= x' - x).
+{ apply Rle_0_minus.
+ apply round_UP_pt.
+ exact Vfexp1. }
+assert (Hr2 : Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))).
+{ replace (x'' - x') with (x'' - x + (x - x')) by ring.
+ apply (Rle_lt_trans _ _ _ (Rabs_triang _ _)).
+ replace (/ 2 * _) with (/ 2 * bpow (fexp2 (mag x))
+ + (/ 2 * (bpow (fexp1 (mag x))
+ - bpow (fexp2 (mag x))))) by ring.
+ apply Rplus_le_lt_compat.
+ - exact Hr1.
+ - rewrite Rabs_minus_sym.
+ rewrite 2!ulp_neq_0 in Hx2; try (apply Rgt_not_eq; assumption).
+ now rewrite Rabs_right; [|now apply Rle_ge]; apply Hx2. }
+destruct (Req_dec x'' 0) as [Zx''|Nzx''].
+- (* x'' = 0 *)
+ rewrite Zx'' in Hr1 |- *.
+ rewrite round_0; [|now apply valid_rnd_N].
+ unfold round, F2R, scaled_mantissa, cexp; simpl.
+ rewrite (Znearest_imp _ _ 0); [now simpl; rewrite Rmult_0_l|].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x)))); [now apply bpow_gt_0|].
+ rewrite <- (Rabs_right (bpow (fexp1 _))) at 1;
+ [|now apply Rle_ge; apply bpow_ge_0].
+ rewrite <- Rabs_mult; rewrite Rmult_minus_distr_r.
+ rewrite Rmult_0_l.
+ bpow_simplify.
+ rewrite Rabs_minus_sym.
+ apply (Rle_lt_trans _ _ _ Hr1).
+ apply Rmult_lt_compat_l; [lra|].
+ apply bpow_lt.
+ omega.
+- (* x'' <> 0 *)
+ assert (Lx'' : mag x'' = mag x :> Z).
+ { apply Zle_antisym.
+ - apply mag_le_bpow; [exact Nzx''|].
+ rewrite Rabs_right; [exact Hx1|apply Rle_ge].
+ apply round_ge_generic.
+ + exact Vfexp2.
+ + now apply valid_rnd_N.
+ + apply generic_format_0.
+ + now apply Rlt_le.
+ - unfold x'' in Nzx'' |- *.
+ now apply mag_round_ge; [|apply valid_rnd_N|]. }
+ unfold round, F2R, scaled_mantissa, cexp; simpl.
+ rewrite Lx''.
+ rewrite (Znearest_imp _ _ (Zceil (scaled_mantissa beta fexp1 x))).
+ + rewrite (Znearest_imp _ _ (Zceil (scaled_mantissa beta fexp1 x)));
+ [reflexivity|].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x)))); [now apply bpow_gt_0|].
+ rewrite <- (Rabs_right (bpow (fexp1 _))) at 1;
+ [|now apply Rle_ge; apply bpow_ge_0].
+ rewrite <- Rabs_mult.
+ rewrite Rmult_minus_distr_r.
+ fold x'.
+ bpow_simplify.
+ rewrite Rabs_minus_sym.
+ rewrite Rabs_right; [|now apply Rle_ge].
+ apply (Rlt_le_trans _ _ _ Hx2).
+ apply Rmult_le_compat_l; [lra|].
+ generalize (bpow_ge_0 beta (fexp2 (mag x))).
+ rewrite 2!ulp_neq_0; try (apply Rgt_not_eq; assumption).
+ unfold cexp; lra.
+ + apply (Rmult_lt_reg_r (bpow (fexp1 (mag x)))); [now apply bpow_gt_0|].
+ rewrite <- (Rabs_right (bpow (fexp1 _))) at 1;
+ [|now apply Rle_ge; apply bpow_ge_0].
+ rewrite <- Rabs_mult.
+ rewrite Rmult_minus_distr_r.
+ fold x'.
+ now bpow_simplify.
+Qed.
+
+Lemma round_round_gt_mid_further_place :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ forall x,
+ 0 < x ->
+ (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z ->
+ (fexp1 (mag x) <= mag x)%Z ->
+ midp' fexp1 x + / 2 * ulp beta fexp2 x < x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x.
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 x Px Hf2f1 Hf1 Hx2'.
+assert (Hx2 : round beta fexp1 Zceil x - x
+ < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)).
+{ apply (Rplus_lt_reg_r (- / 2 * ulp beta fexp1 x + x
+ + / 2 * ulp beta fexp2 x)); ring_simplify.
+ now unfold midp' in Hx2'. }
+revert Hx2.
+unfold round_round_eq.
+set (x' := round beta fexp1 Zfloor x).
+intro Hx2.
+set (x'' := round beta fexp2 (Znearest choice2) x).
+destruct (Rlt_or_le x'' (bpow (mag x))) as [Hx''|Hx''];
+ [now apply round_round_gt_mid_further_place'|].
+(* bpow (mag x) <= x'' *)
+assert (Hx''pow : x'' = bpow (mag x)).
+{ assert (H'x'' : x'' < bpow (mag x) + / 2 * ulp beta fexp2 x).
+ { apply Rle_lt_trans with (x + / 2 * ulp beta fexp2 x).
+ - apply (Rplus_le_reg_r (- x)); ring_simplify.
+ apply Rabs_le_inv.
+ apply error_le_half_ulp.
+ exact Vfexp2.
+ - apply Rplus_lt_compat_r.
+ rewrite <- Rabs_right at 1; [|now apply Rle_ge; apply Rlt_le].
+ apply bpow_mag_gt. }
+ apply Rle_antisym; [|exact Hx''].
+ unfold x'', round, F2R, scaled_mantissa, cexp; simpl.
+ apply (Rmult_le_reg_r (bpow (- fexp2 (mag x)))); [now apply bpow_gt_0|].
+ bpow_simplify.
+ rewrite <- (IZR_Zpower _ (_ - _)); [|omega].
+ apply IZR_le.
+ apply Zlt_succ_le; unfold Z.succ.
+ apply lt_IZR.
+ rewrite plus_IZR; rewrite IZR_Zpower; [|omega].
+ apply (Rmult_lt_reg_r (bpow (fexp2 (mag x)))); [now apply bpow_gt_0|].
+ rewrite Rmult_plus_distr_r; rewrite Rmult_1_l.
+ bpow_simplify.
+ apply (Rlt_le_trans _ _ _ H'x'').
+ apply Rplus_le_compat_l.
+ rewrite <- (Rmult_1_l (Raux.bpow _ _)).
+ rewrite ulp_neq_0;[idtac|now apply Rgt_not_eq].
+ apply Rmult_le_compat_r; [now apply bpow_ge_0|].
+ lra. }
+assert (Hr : Rabs (x - x'') < / 2 * ulp beta fexp1 x).
+{ apply Rle_lt_trans with (/ 2 * ulp beta fexp2 x).
+ - rewrite Rabs_minus_sym.
+ apply error_le_half_ulp.
+ exact Vfexp2.
+ - apply Rmult_lt_compat_l; [lra|].
+ rewrite 2!ulp_neq_0; try now apply Rgt_not_eq.
+ unfold cexp; apply bpow_lt.
+ omega. }
+unfold round, F2R, scaled_mantissa, cexp; simpl.
+assert (Hf : (0 <= mag x - fexp1 (mag x''))%Z).
+{ rewrite Hx''pow.
+ rewrite mag_bpow.
+ assert (fexp1 (mag x + 1) <= mag x)%Z; [|omega].
+ destruct (Zle_or_lt (mag x) (fexp1 (mag x))) as [Hle|Hlt];
+ [|now apply Vfexp1].
+ assert (H : (mag x = fexp1 (mag x) :> Z)%Z);
+ [now apply Zle_antisym|].
+ rewrite H.
+ now apply Vfexp1. }
+rewrite (Znearest_imp _ _ (beta ^ (mag x - fexp1 (mag x'')))%Z).
+- rewrite (Znearest_imp _ _ (beta ^ (mag x - fexp1 (mag x)))%Z).
+ + rewrite IZR_Zpower; [|exact Hf].
+ rewrite IZR_Zpower; [|omega].
+ now bpow_simplify.
+ + rewrite IZR_Zpower; [|omega].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x)))); [now apply bpow_gt_0|].
+ rewrite <- (Rabs_right (bpow (fexp1 _))) at 1;
+ [|now apply Rle_ge; apply bpow_ge_0].
+ rewrite <- Rabs_mult.
+ rewrite Rmult_minus_distr_r.
+ bpow_simplify.
+ rewrite ulp_neq_0 in Hr;[idtac|now apply Rgt_not_eq].
+ rewrite <- Hx''pow; exact Hr.
+- rewrite IZR_Zpower; [|exact Hf].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x'')))); [now apply bpow_gt_0|].
+ rewrite <- (Rabs_right (bpow (fexp1 _))) at 1;
+ [|now apply Rle_ge; apply bpow_ge_0].
+ rewrite <- Rabs_mult.
+ rewrite Rmult_minus_distr_r.
+ bpow_simplify.
+ rewrite Rminus_diag_eq; [|exact Hx''pow].
+ rewrite Rabs_R0.
+ rewrite <- (Rmult_0_r (/ 2)).
+ apply Rmult_lt_compat_l; [lra|apply bpow_gt_0].
+Qed.
+
+Lemma round_round_gt_mid_same_place :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 ->
+ forall (choice1 choice2 : Z -> bool),
+ forall x,
+ 0 < x ->
+ (fexp2 (mag x) = fexp1 (mag x))%Z ->
+ midp' fexp1 x < x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x.
+Proof.
+intros fexp1 fexp2 Vfexp1 choice1 choice2 x Px Hf2f1 Hx'.
+assert (Hx : round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x).
+{ apply (Rplus_lt_reg_r (- / 2 * ulp beta fexp1 x + x)); ring_simplify.
+ now unfold midp' in Hx'. }
+assert (H : Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x))))
+ - x * bpow (- fexp1 (mag x))) < / 2).
+{ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x)))); [now apply bpow_gt_0|].
+ unfold scaled_mantissa, cexp in Hx.
+ rewrite <- (Rabs_right (bpow (fexp1 _))) at 1;
+ [|now apply Rle_ge; apply bpow_ge_0].
+ rewrite <- Rabs_mult.
+ rewrite Rmult_minus_distr_r.
+ bpow_simplify.
+ apply Rabs_lt.
+ split.
+ - apply Rlt_le_trans with 0.
+ + rewrite <- Ropp_0; apply Ropp_lt_contravar.
+ rewrite <- (Rmult_0_r (/ 2)).
+ apply Rmult_lt_compat_l; [lra|].
+ apply bpow_gt_0.
+ + apply Rle_0_minus.
+ apply round_UP_pt.
+ exact Vfexp1.
+ - rewrite ulp_neq_0 in Hx;[exact Hx|now apply Rgt_not_eq]. }
+unfold round_round_eq, round at 2.
+unfold F2R, scaled_mantissa, cexp; simpl.
+rewrite Hf2f1.
+rewrite (Znearest_imp _ _ (Zceil (scaled_mantissa beta fexp1 x))).
+- rewrite round_generic.
+ + unfold round, F2R, scaled_mantissa, cexp; simpl.
+ now rewrite (Znearest_imp _ _ (Zceil (x * bpow (- fexp1 (mag x)))));
+ [|rewrite Rabs_minus_sym].
+ + now apply valid_rnd_N.
+ + fold (cexp beta fexp1 x).
+ change (IZR _ * bpow _) with (round beta fexp1 Zceil x).
+ apply generic_format_round.
+ exact Vfexp1.
+ now apply valid_rnd_UP.
+- now rewrite Rabs_minus_sym.
+Qed.
+
+Lemma round_round_gt_mid :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ forall x,
+ 0 < x ->
+ (fexp2 (mag x) <= fexp1 (mag x))%Z ->
+ (fexp1 (mag x) <= mag x)%Z ->
+ midp' fexp1 x < x ->
+ ((fexp2 (mag x) <= fexp1 (mag x) - 1)%Z ->
+ midp' fexp1 x + / 2 * ulp beta fexp2 x < x) ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x.
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 x Px Hf2f1 Hf1 Hx Hx'.
+destruct (Zle_or_lt (fexp1 (mag x)) (fexp2 (mag x))) as [Hf2'|Hf2'].
+- (* fexp1 (mag x) <= fexp2 (mag x) *)
+ assert (Hf2'' : (fexp2 (mag x) = fexp1 (mag x) :> Z)%Z); [omega|].
+ now apply round_round_gt_mid_same_place.
+- (* fexp2 (mag x) < fexp1 (mag x) *)
+ assert (Hf2'' : (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z); [omega|].
+ generalize (Hx' Hf2''); intro Hx''.
+ now apply round_round_gt_mid_further_place.
+Qed.
+
+Section Double_round_mult.
+
+Lemma mag_mult_disj :
+ forall x y,
+ x <> 0 -> y <> 0 ->
+ ((mag (x * y) = (mag x + mag y - 1)%Z :> Z)
+ \/ (mag (x * y) = (mag x + mag y)%Z :> Z)).
+Proof.
+intros x y Zx Zy.
+destruct (mag_mult beta x y Zx Zy).
+omega.
+Qed.
+
+Definition round_round_mult_hyp fexp1 fexp2 :=
+ (forall ex ey, (fexp2 (ex + ey) <= fexp1 ex + fexp1 ey)%Z)
+ /\ (forall ex ey, (fexp2 (ex + ey - 1) <= fexp1 ex + fexp1 ey)%Z).
+
+Lemma round_round_mult_aux :
+ forall (fexp1 fexp2 : Z -> Z),
+ round_round_mult_hyp fexp1 fexp2 ->
+ forall x y,
+ generic_format beta fexp1 x -> generic_format beta fexp1 y ->
+ generic_format beta fexp2 (x * y).
+Proof.
+intros fexp1 fexp2 Hfexp x y Fx Fy.
+destruct (Req_dec x 0) as [Zx|Zx].
+- (* x = 0 *)
+ rewrite Zx.
+ rewrite Rmult_0_l.
+ now apply generic_format_0.
+- (* x <> 0 *)
+ destruct (Req_dec y 0) as [Zy|Zy].
+ + (* y = 0 *)
+ rewrite Zy.
+ rewrite Rmult_0_r.
+ now apply generic_format_0.
+ + (* y <> 0 *)
+ revert Fx Fy.
+ unfold generic_format.
+ unfold cexp.
+ set (mx := Ztrunc (scaled_mantissa beta fexp1 x)).
+ set (my := Ztrunc (scaled_mantissa beta fexp1 y)).
+ unfold F2R; simpl.
+ intros Fx Fy.
+ set (fxy := Float beta (mx * my) (fexp1 (mag x) + fexp1 (mag y))).
+ assert (Hxy : x * y = F2R fxy).
+ { unfold fxy, F2R; simpl.
+ rewrite bpow_plus.
+ rewrite mult_IZR.
+ rewrite Fx, Fy at 1.
+ ring. }
+ apply generic_format_F2R' with (f := fxy); [now rewrite Hxy|].
+ intros _.
+ unfold cexp, fxy; simpl.
+ destruct Hfexp as (Hfexp1, Hfexp2).
+ now destruct (mag_mult_disj x y Zx Zy) as [Lxy|Lxy]; rewrite Lxy.
+Qed.
+
+Variable rnd : R -> Z.
+Context { valid_rnd : Valid_rnd rnd }.
+
+Theorem round_round_mult :
+ forall (fexp1 fexp2 : Z -> Z),
+ round_round_mult_hyp fexp1 fexp2 ->
+ forall x y,
+ generic_format beta fexp1 x -> generic_format beta fexp1 y ->
+ round beta fexp1 rnd (round beta fexp2 rnd (x * y))
+ = round beta fexp1 rnd (x * y).
+Proof.
+intros fexp1 fexp2 Hfexp x y Fx Fy.
+assert (Hxy : round beta fexp2 rnd (x * y) = x * y).
+{ apply round_generic; [assumption|].
+ now apply (round_round_mult_aux fexp1 fexp2). }
+now rewrite Hxy at 1.
+Qed.
+
+Section Double_round_mult_FLX.
+
+Variable prec : Z.
+Variable prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Theorem round_round_mult_FLX :
+ (2 * prec <= prec')%Z ->
+ forall x y,
+ FLX_format beta prec x -> FLX_format beta prec y ->
+ round beta (FLX_exp prec) rnd (round beta (FLX_exp prec') rnd (x * y))
+ = round beta (FLX_exp prec) rnd (x * y).
+Proof.
+intros Hprec x y Fx Fy.
+apply round_round_mult;
+ [|now apply generic_format_FLX|now apply generic_format_FLX].
+unfold round_round_mult_hyp; split; intros ex ey; unfold FLX_exp;
+omega.
+Qed.
+
+End Double_round_mult_FLX.
+
+Section Double_round_mult_FLT.
+
+Variable emin prec : Z.
+Variable emin' prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Theorem round_round_mult_FLT :
+ (emin' <= 2 * emin)%Z -> (2 * prec <= prec')%Z ->
+ forall x y,
+ FLT_format beta emin prec x -> FLT_format beta emin prec y ->
+ round beta (FLT_exp emin prec) rnd
+ (round beta (FLT_exp emin' prec') rnd (x * y))
+ = round beta (FLT_exp emin prec) rnd (x * y).
+Proof.
+intros Hemin Hprec x y Fx Fy.
+apply round_round_mult;
+ [|now apply generic_format_FLT|now apply generic_format_FLT].
+unfold round_round_mult_hyp; split; intros ex ey;
+unfold FLT_exp;
+generalize (Zmax_spec (ex + ey - prec') emin');
+generalize (Zmax_spec (ex + ey - 1 - prec') emin');
+generalize (Zmax_spec (ex - prec) emin);
+generalize (Zmax_spec (ey - prec) emin);
+omega.
+Qed.
+
+End Double_round_mult_FLT.
+
+Section Double_round_mult_FTZ.
+
+Variable emin prec : Z.
+Variable emin' prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Theorem round_round_mult_FTZ :
+ (emin' + prec' <= 2 * emin + prec)%Z ->
+ (2 * prec <= prec')%Z ->
+ forall x y,
+ FTZ_format beta emin prec x -> FTZ_format beta emin prec y ->
+ round beta (FTZ_exp emin prec) rnd
+ (round beta (FTZ_exp emin' prec') rnd (x * y))
+ = round beta (FTZ_exp emin prec) rnd (x * y).
+Proof.
+intros Hemin Hprec x y Fx Fy.
+apply round_round_mult;
+ [|now apply generic_format_FTZ|now apply generic_format_FTZ].
+unfold round_round_mult_hyp; split; intros ex ey;
+unfold FTZ_exp;
+unfold Prec_gt_0 in *;
+destruct (Z.ltb_spec (ex + ey - prec') emin');
+destruct (Z.ltb_spec (ex - prec) emin);
+destruct (Z.ltb_spec (ey - prec) emin);
+destruct (Z.ltb_spec (ex + ey - 1 - prec') emin');
+omega.
+Qed.
+
+End Double_round_mult_FTZ.
+
+End Double_round_mult.
+
+Section Double_round_plus.
+
+Lemma mag_plus_disj :
+ forall x y,
+ 0 < y -> y <= x ->
+ ((mag (x + y) = mag x :> Z)
+ \/ (mag (x + y) = (mag x + 1)%Z :> Z)).
+Proof.
+intros x y Py Hxy.
+destruct (mag_plus beta x y Py Hxy).
+omega.
+Qed.
+
+Lemma mag_plus_separated :
+ forall fexp : Z -> Z,
+ forall x y,
+ 0 < x -> 0 <= y ->
+ generic_format beta fexp x ->
+ (mag y <= fexp (mag x))%Z ->
+ (mag (x + y) = mag x :> Z).
+Proof.
+intros fexp x y Px Nny Fx Hsep.
+apply mag_plus_eps with (1 := Px) (2 := Fx).
+apply (conj Nny).
+rewrite <- Rabs_pos_eq with (1 := Nny).
+apply Rlt_le_trans with (1 := bpow_mag_gt beta _).
+rewrite ulp_neq_0 by now apply Rgt_not_eq.
+now apply bpow_le.
+Qed.
+
+Lemma mag_minus_disj :
+ forall x y,
+ 0 < x -> 0 < y ->
+ (mag y <= mag x - 2)%Z ->
+ ((mag (x - y) = mag x :> Z)
+ \/ (mag (x - y) = (mag x - 1)%Z :> Z)).
+Proof.
+intros x y Px Py Hln.
+assert (Hxy : y < x); [now apply (lt_mag beta); [ |omega]|].
+generalize (mag_minus beta x y Py Hxy); intro Hln2.
+generalize (mag_minus_lb beta x y Px Py Hln); intro Hln3.
+omega.
+Qed.
+
+Lemma mag_minus_separated :
+ forall fexp : Z -> Z, Valid_exp fexp ->
+ forall x y,
+ 0 < x -> 0 < y -> y < x ->
+ bpow (mag x - 1) < x ->
+ generic_format beta fexp x -> (mag y <= fexp (mag x))%Z ->
+ (mag (x - y) = mag x :> Z).
+Proof.
+intros fexp Vfexp x y Px Py Yltx Xgtpow Fx Ly.
+apply mag_unique.
+split.
+- apply Rabs_ge; right.
+ assert (Hy : y < ulp beta fexp (bpow (mag x - 1))).
+ { rewrite ulp_bpow.
+ replace (_ + _)%Z with (mag x : Z) by ring.
+ rewrite <- (Rabs_right y); [|now apply Rle_ge; apply Rlt_le].
+ apply Rlt_le_trans with (bpow (mag y)).
+ - apply bpow_mag_gt.
+ - now apply bpow_le. }
+ apply (Rplus_le_reg_r y); ring_simplify.
+ apply Rle_trans with (bpow (mag x - 1)
+ + ulp beta fexp (bpow (mag x - 1))).
+ + now apply Rplus_le_compat_l; apply Rlt_le.
+ + rewrite <- succ_eq_pos;[idtac|apply bpow_ge_0].
+ apply succ_le_lt; [apply Vfexp|idtac|exact Fx|assumption].
+ apply (generic_format_bpow beta fexp (mag x - 1)).
+ replace (_ + _)%Z with (mag x : Z) by ring.
+ assert (fexp (mag x) < mag x)%Z; [|omega].
+ now apply mag_generic_gt; [|now apply Rgt_not_eq|].
+- rewrite Rabs_right.
+ + apply Rlt_trans with x.
+ * rewrite <- (Rplus_0_r x) at 2.
+ apply Rplus_lt_compat_l.
+ rewrite <- Ropp_0.
+ now apply Ropp_lt_contravar.
+ * apply Rabs_lt_inv.
+ apply bpow_mag_gt.
+ + lra.
+Qed.
+
+Definition round_round_plus_hyp fexp1 fexp2 :=
+ (forall ex ey, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z)
+ /\ (forall ex ey, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z)
+ /\ (forall ex ey, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z)
+ /\ (forall ex ey, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z).
+
+Lemma round_round_plus_aux0_aux_aux :
+ forall (fexp1 fexp2 : Z -> Z),
+ forall x y,
+ (fexp1 (mag x) <= fexp1 (mag y))%Z ->
+ (fexp2 (mag (x + y))%Z <= fexp1 (mag x))%Z ->
+ (fexp2 (mag (x + y))%Z <= fexp1 (mag y))%Z ->
+ generic_format beta fexp1 x -> generic_format beta fexp1 y ->
+ generic_format beta fexp2 (x + y).
+Proof.
+intros fexp1 fexp2 x y Oxy Hlnx Hlny Fx Fy.
+destruct (Req_dec x 0) as [Zx|Nzx].
+- (* x = 0 *)
+ rewrite Zx, Rplus_0_l in Hlny |- *.
+ now apply (generic_inclusion_mag beta fexp1).
+- (* x <> 0 *)
+ destruct (Req_dec y 0) as [Zy|Nzy].
+ + (* y = 0 *)
+ rewrite Zy, Rplus_0_r in Hlnx |- *.
+ now apply (generic_inclusion_mag beta fexp1).
+ + (* y <> 0 *)
+ revert Fx Fy.
+ unfold generic_format at -3, cexp, F2R; simpl.
+ set (mx := Ztrunc (scaled_mantissa beta fexp1 x)).
+ set (my := Ztrunc (scaled_mantissa beta fexp1 y)).
+ intros Fx Fy.
+ set (fxy := Float beta (mx + my * (beta ^ (fexp1 (mag y)
+ - fexp1 (mag x))))
+ (fexp1 (mag x))).
+ assert (Hxy : x + y = F2R fxy).
+ { unfold fxy, F2R; simpl.
+ rewrite plus_IZR.
+ rewrite Rmult_plus_distr_r.
+ rewrite <- Fx.
+ rewrite mult_IZR.
+ rewrite IZR_Zpower; [|omega].
+ bpow_simplify.
+ now rewrite <- Fy. }
+ apply generic_format_F2R' with (f := fxy); [now rewrite Hxy|].
+ intros _.
+ now unfold cexp, fxy; simpl.
+Qed.
+
+Lemma round_round_plus_aux0_aux :
+ forall (fexp1 fexp2 : Z -> Z),
+ forall x y,
+ (fexp2 (mag (x + y))%Z <= fexp1 (mag x))%Z ->
+ (fexp2 (mag (x + y))%Z <= fexp1 (mag y))%Z ->
+ generic_format beta fexp1 x -> generic_format beta fexp1 y ->
+ generic_format beta fexp2 (x + y).
+Proof.
+intros fexp1 fexp2 x y Hlnx Hlny Fx Fy.
+destruct (Z.le_gt_cases (fexp1 (mag x)) (fexp1 (mag y))) as [Hle|Hgt].
+- now apply (round_round_plus_aux0_aux_aux fexp1).
+- rewrite Rplus_comm in Hlnx, Hlny |- *.
+ now apply (round_round_plus_aux0_aux_aux fexp1); [omega| | | |].
+Qed.
+
+(* fexp1 (mag x) - 1 <= mag y :
+ * addition is exact in the largest precision (fexp2). *)
+Lemma round_round_plus_aux0 :
+ forall (fexp1 fexp2 : Z -> Z), Valid_exp fexp1 ->
+ round_round_plus_hyp fexp1 fexp2 ->
+ forall x y,
+ (0 < x)%R -> (0 < y)%R -> (y <= x)%R ->
+ (fexp1 (mag x) - 1 <= mag y)%Z ->
+ generic_format beta fexp1 x -> generic_format beta fexp1 y ->
+ generic_format beta fexp2 (x + y).
+Proof.
+intros fexp1 fexp2 Vfexp1 Hexp x y Px Py Hyx Hln Fx Fy.
+assert (Nny : (0 <= y)%R); [now apply Rlt_le|].
+destruct Hexp as (_,(Hexp2,(Hexp3,Hexp4))).
+destruct (Z.le_gt_cases (mag y) (fexp1 (mag x))) as [Hle|Hgt].
+- (* mag y <= fexp1 (mag x) *)
+ assert (Lxy : mag (x + y) = mag x :> Z);
+ [now apply (mag_plus_separated fexp1)|].
+ apply (round_round_plus_aux0_aux fexp1);
+ [| |assumption|assumption]; rewrite Lxy.
+ + now apply Hexp4; omega.
+ + now apply Hexp3; omega.
+- (* fexp1 (mag x) < mag y *)
+ apply (round_round_plus_aux0_aux fexp1); [| |assumption|assumption].
+ destruct (mag_plus_disj x y Py Hyx) as [Lxy|Lxy]; rewrite Lxy.
+ + now apply Hexp4; omega.
+ + apply Hexp2; apply (mag_le beta y x Py) in Hyx.
+ replace (_ - _)%Z with (mag x : Z) by ring.
+ omega.
+ + destruct (mag_plus_disj x y Py Hyx) as [Lxy|Lxy]; rewrite Lxy.
+ * now apply Hexp3; omega.
+ * apply Hexp2.
+ replace (_ - _)%Z with (mag x : Z) by ring.
+ omega.
+Qed.
+
+Lemma round_round_plus_aux1_aux :
+ forall k, (0 < k)%Z ->
+ forall (fexp : Z -> Z),
+ forall x y,
+ 0 < x -> 0 < y ->
+ (mag y <= fexp (mag x) - k)%Z ->
+ (mag (x + y) = mag x :> Z) ->
+ generic_format beta fexp x ->
+ 0 < (x + y) - round beta fexp Zfloor (x + y) < bpow (fexp (mag x) - k).
+Proof.
+assert (Hbeta : (2 <= beta)%Z).
+{ destruct beta as (beta_val,beta_prop).
+ now apply Zle_bool_imp_le. }
+intros k Hk fexp x y Px Py Hln Hlxy Fx.
+revert Fx.
+unfold round, generic_format, F2R, scaled_mantissa, cexp; simpl.
+rewrite Hlxy.
+set (mx := Ztrunc (x * bpow (- fexp (mag x)))).
+intros Fx.
+assert (R : (x + y) * bpow (- fexp (mag x))
+ = IZR mx + y * bpow (- fexp (mag x))).
+{ rewrite Fx at 1.
+ rewrite Rmult_plus_distr_r.
+ now bpow_simplify. }
+rewrite R.
+assert (LB : 0 < y * bpow (- fexp (mag x))).
+{ rewrite <- (Rmult_0_r y).
+ now apply Rmult_lt_compat_l; [|apply bpow_gt_0]. }
+assert (UB : y * bpow (- fexp (mag x)) < / IZR (beta ^ k)).
+{ apply Rlt_le_trans with (bpow (mag y) * bpow (- fexp (mag x))).
+ - apply Rmult_lt_compat_r; [now apply bpow_gt_0|].
+ rewrite <- (Rabs_right y) at 1; [|now apply Rle_ge; apply Rlt_le].
+ apply bpow_mag_gt.
+ - apply Rle_trans with (bpow (fexp (mag x) - k)
+ * bpow (- fexp (mag x)))%R.
+ + apply Rmult_le_compat_r; [now apply bpow_ge_0|].
+ now apply bpow_le.
+ + bpow_simplify.
+ rewrite bpow_opp.
+ destruct k.
+ * omega.
+ * simpl; unfold Raux.bpow, Z.pow_pos.
+ now apply Rle_refl.
+ * casetype False; apply (Z.lt_irrefl 0).
+ apply (Z.lt_trans _ _ _ Hk).
+ apply Zlt_neg_0. }
+rewrite (Zfloor_imp mx).
+{ split; ring_simplify.
+ - apply (Rmult_lt_reg_r (bpow (- fexp (mag x)))); [now apply bpow_gt_0|].
+ rewrite Rmult_minus_distr_r, Rmult_0_l.
+ bpow_simplify.
+ rewrite R; ring_simplify.
+ now apply Rmult_lt_0_compat; [|apply bpow_gt_0].
+ - apply (Rmult_lt_reg_r (bpow (- fexp (mag x)))); [now apply bpow_gt_0|].
+ rewrite Rmult_minus_distr_r.
+ bpow_simplify.
+ rewrite R; ring_simplify.
+ apply (Rlt_le_trans _ _ _ UB).
+ rewrite bpow_opp.
+ apply Rinv_le; [now apply bpow_gt_0|].
+ now rewrite IZR_Zpower; [right|omega]. }
+split.
+- rewrite <- Rplus_0_r at 1; apply Rplus_le_compat_l.
+ now apply Rlt_le.
+- rewrite plus_IZR; apply Rplus_lt_compat_l.
+ apply (Rmult_lt_reg_r (bpow (fexp (mag x)))); [now apply bpow_gt_0|].
+ rewrite Rmult_1_l.
+ bpow_simplify.
+ apply Rlt_trans with (bpow (mag y)).
+ + rewrite <- Rabs_right at 1; [|now apply Rle_ge; apply Rlt_le].
+ apply bpow_mag_gt.
+ + apply bpow_lt; omega.
+Qed.
+
+(* mag y <= fexp1 (mag x) - 2 : round_round_lt_mid applies. *)
+Lemma round_round_plus_aux1 :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < x -> 0 < y ->
+ (mag y <= fexp1 (mag x) - 2)%Z ->
+ generic_format beta fexp1 x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x + y).
+Proof.
+assert (Hbeta : (2 <= beta)%Z).
+{ destruct beta as (beta_val,beta_prop).
+ now apply Zle_bool_imp_le. }
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Px Py Hly Fx.
+assert (Lxy : mag (x + y) = mag x :> Z);
+ [now apply (mag_plus_separated fexp1); [|apply Rlt_le| |omega]|].
+destruct Hexp as (_,(_,(_,Hexp4))).
+assert (Hf2 : (fexp2 (mag x) <= fexp1 (mag x))%Z);
+ [now apply Hexp4; omega|].
+assert (Bpow2 : bpow (- 2) <= / 2 * / 2).
+{ replace (/2 * /2) with (/4) by field.
+ rewrite (bpow_opp _ 2).
+ apply Rinv_le; [lra|].
+ apply (IZR_le (2 * 2) (beta * (beta * 1))).
+ rewrite Zmult_1_r.
+ now apply Zmult_le_compat; omega. }
+assert (P2 : (0 < 2)%Z) by omega.
+unfold round_round_eq.
+apply round_round_lt_mid.
+- exact Vfexp1.
+- exact Vfexp2.
+- lra.
+- now rewrite Lxy.
+- rewrite Lxy.
+ assert (fexp1 (mag x) < mag x)%Z; [|omega].
+ now apply mag_generic_gt; [|apply Rgt_not_eq|].
+- unfold midp.
+ apply (Rplus_lt_reg_r (- round beta fexp1 Zfloor (x + y))).
+ apply (Rlt_le_trans _ _ _ (proj2 (round_round_plus_aux1_aux 2 P2 fexp1 x y Px
+ Py Hly Lxy Fx))).
+ ring_simplify.
+ rewrite ulp_neq_0;[idtac|now apply Rgt_not_eq, Rplus_lt_0_compat].
+ unfold cexp; rewrite Lxy.
+ apply (Rmult_le_reg_r (bpow (- fexp1 (mag x)))); [now apply bpow_gt_0|].
+ bpow_simplify.
+ apply (Rle_trans _ _ _ Bpow2).
+ rewrite <- (Rmult_1_r (/ 2)) at 3.
+ apply Rmult_le_compat_l; lra.
+- rewrite ulp_neq_0;[idtac|now apply Rgt_not_eq, Rplus_lt_0_compat].
+ unfold round, F2R, scaled_mantissa, cexp; simpl; rewrite Lxy.
+ intro Hf2'.
+ apply (Rmult_lt_reg_r (bpow (- fexp1 (mag x))));
+ [now apply bpow_gt_0|].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x)))); [now apply bpow_gt_0|].
+ bpow_simplify.
+ apply (Rplus_lt_reg_r (- round beta fexp1 Zfloor (x + y))).
+ unfold midp; ring_simplify.
+ apply (Rlt_le_trans _ _ _ (proj2 (round_round_plus_aux1_aux 2 P2 fexp1 x y Px
+ Py Hly Lxy Fx))).
+ apply (Rmult_le_reg_r (bpow (- fexp1 (mag x)))); [now apply bpow_gt_0|].
+ rewrite ulp_neq_0;[idtac|now apply Rgt_not_eq, Rplus_lt_0_compat].
+ unfold cexp; rewrite Lxy, Rmult_minus_distr_r; bpow_simplify.
+ apply (Rle_trans _ _ _ Bpow2).
+ rewrite <- (Rmult_1_r (/ 2)) at 3; rewrite <- Rmult_minus_distr_l.
+ apply Rmult_le_compat_l; [lra|].
+ apply (Rplus_le_reg_r (- 1)); ring_simplify.
+ replace (_ - _) with (- (/ 2)) by lra.
+ apply Ropp_le_contravar.
+ { apply Rle_trans with (bpow (- 1)).
+ - apply bpow_le; omega.
+ - unfold Raux.bpow, Z.pow_pos; simpl.
+ apply Rinv_le; [lra|].
+ apply IZR_le; omega. }
+Qed.
+
+(* round_round_plus_aux{0,1} together *)
+Lemma round_round_plus_aux2 :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < x -> 0 < y -> y <= x ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x + y).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Px Py Hyx Fx Fy.
+unfold round_round_eq.
+destruct (Zle_or_lt (mag y) (fexp1 (mag x) - 2)) as [Hly|Hly].
+- (* mag y <= fexp1 (mag x) - 2 *)
+ now apply round_round_plus_aux1.
+- (* fexp1 (mag x) - 2 < mag y *)
+ rewrite (round_generic beta fexp2).
+ + reflexivity.
+ + now apply valid_rnd_N.
+ + assert (Hf1 : (fexp1 (mag x) - 1 <= mag y)%Z); [omega|].
+ now apply (round_round_plus_aux0 fexp1).
+Qed.
+
+Lemma round_round_plus_aux :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 <= x -> 0 <= y ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x + y).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Nnx Nny Fx Fy.
+unfold round_round_eq.
+destruct (Req_dec x 0) as [Zx|Nzx].
+- (* x = 0 *)
+ destruct Hexp as (_,(_,(_,Hexp4))).
+ rewrite Zx; rewrite Rplus_0_l.
+ rewrite (round_generic beta fexp2).
+ + reflexivity.
+ + now apply valid_rnd_N.
+ + apply (generic_inclusion_mag beta fexp1).
+ now intros _; apply Hexp4; omega.
+ exact Fy.
+- (* x <> 0 *)
+ destruct (Req_dec y 0) as [Zy|Nzy].
+ + (* y = 0 *)
+ destruct Hexp as (_,(_,(_,Hexp4))).
+ rewrite Zy; rewrite Rplus_0_r.
+ rewrite (round_generic beta fexp2).
+ * reflexivity.
+ * now apply valid_rnd_N.
+ * apply (generic_inclusion_mag beta fexp1).
+ now intros _; apply Hexp4; omega.
+ exact Fx.
+ + (* y <> 0 *)
+ assert (Px : 0 < x); [lra|].
+ assert (Py : 0 < y); [lra|].
+ destruct (Rlt_or_le x y) as [H|H].
+ * (* x < y *)
+ apply Rlt_le in H.
+ rewrite Rplus_comm.
+ now apply round_round_plus_aux2.
+ * now apply round_round_plus_aux2.
+Qed.
+
+Lemma round_round_minus_aux0_aux :
+ forall (fexp1 fexp2 : Z -> Z),
+ forall x y,
+ (fexp2 (mag (x - y))%Z <= fexp1 (mag x))%Z ->
+ (fexp2 (mag (x - y))%Z <= fexp1 (mag y))%Z ->
+ generic_format beta fexp1 x -> generic_format beta fexp1 y ->
+ generic_format beta fexp2 (x - y).
+Proof.
+intros fexp1 fexp2 x y.
+replace (x - y)%R with (x + (- y))%R; [|ring].
+intros Hlnx Hlny Fx Fy.
+rewrite <- (mag_opp beta y) in Hlny.
+apply generic_format_opp in Fy.
+now apply (round_round_plus_aux0_aux fexp1).
+Qed.
+
+(* fexp1 (mag x) - 1 <= mag y :
+ * substraction is exact in the largest precision (fexp2). *)
+Lemma round_round_minus_aux0 :
+ forall (fexp1 fexp2 : Z -> Z),
+ round_round_plus_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < y -> y < x ->
+ (fexp1 (mag x) - 1 <= mag y)%Z ->
+ generic_format beta fexp1 x -> generic_format beta fexp1 y ->
+ generic_format beta fexp2 (x - y).
+Proof.
+intros fexp1 fexp2 Hexp x y Py Hyx Hln Fx Fy.
+assert (Px := Rlt_trans 0 y x Py Hyx).
+destruct Hexp as (Hexp1,(_,(Hexp3,Hexp4))).
+assert (Lyx : (mag y <= mag x)%Z);
+ [now apply mag_le; [|apply Rlt_le]|].
+destruct (Z.lt_ge_cases (mag x - 2) (mag y)) as [Hlt|Hge].
+- (* mag x - 2 < mag y *)
+ assert (Hor : (mag y = mag x :> Z)
+ \/ (mag y = mag x - 1 :> Z)%Z); [omega|].
+ destruct Hor as [Heq|Heqm1].
+ + (* mag y = mag x *)
+ apply (round_round_minus_aux0_aux fexp1); [| |exact Fx|exact Fy].
+ * apply Hexp4.
+ apply Z.le_trans with (mag (x - y)); [omega|].
+ now apply mag_minus.
+ * rewrite Heq.
+ apply Hexp4.
+ apply Z.le_trans with (mag (x - y)); [omega|].
+ now apply mag_minus.
+ + (* mag y = mag x - 1 *)
+ apply (round_round_minus_aux0_aux fexp1); [| |exact Fx|exact Fy].
+ * apply Hexp4.
+ apply Z.le_trans with (mag (x - y)); [omega|].
+ now apply mag_minus.
+ * rewrite Heqm1.
+ apply Hexp4.
+ apply Zplus_le_compat_r.
+ now apply mag_minus.
+- (* mag y <= mag x - 2 *)
+ destruct (mag_minus_disj x y Px Py Hge) as [Lxmy|Lxmy].
+ + (* mag (x - y) = mag x *)
+ apply (round_round_minus_aux0_aux fexp1); [| |exact Fx|exact Fy].
+ * apply Hexp4.
+ omega.
+ * now rewrite Lxmy; apply Hexp3.
+ + (* mag (x - y) = mag x - 1 *)
+ apply (round_round_minus_aux0_aux fexp1); [| |exact Fx|exact Fy];
+ rewrite Lxmy.
+ * apply Hexp1.
+ replace (_ + _)%Z with (mag x : Z); [|ring].
+ now apply Z.le_trans with (mag y).
+ * apply Hexp1.
+ now replace (_ + _)%Z with (mag x : Z); [|ring].
+Qed.
+
+(* mag y <= fexp1 (mag x) - 2,
+ * fexp1 (mag (x - y)) - 1 <= mag y :
+ * substraction is exact in the largest precision (fexp2). *)
+Lemma round_round_minus_aux1 :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ round_round_plus_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < y -> y < x ->
+ (mag y <= fexp1 (mag x) - 2)%Z ->
+ (fexp1 (mag (x - y)) - 1 <= mag y)%Z ->
+ generic_format beta fexp1 x -> generic_format beta fexp1 y ->
+ generic_format beta fexp2 (x - y).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 Hexp x y Py Hyx Hln Hln' Fx Fy.
+assert (Px := Rlt_trans 0 y x Py Hyx).
+destruct Hexp as (Hexp1,(Hexp2,(Hexp3,Hexp4))).
+assert (Lyx : (mag y <= mag x)%Z);
+ [now apply mag_le; [|apply Rlt_le]|].
+assert (Hfx : (fexp1 (mag x) < mag x)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+assert (Hfy : (fexp1 (mag y) < mag y)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+apply (round_round_minus_aux0_aux fexp1); [| |exact Fx|exact Fy].
+- apply Z.le_trans with (fexp1 (mag (x - y))).
+ + apply Hexp4; omega.
+ + omega.
+- now apply Hexp3.
+Qed.
+
+Lemma round_round_minus_aux2_aux :
+ forall (fexp : Z -> Z),
+ Valid_exp fexp ->
+ forall x y,
+ 0 < y -> y < x ->
+ (mag y <= fexp (mag x) - 1)%Z ->
+ generic_format beta fexp x ->
+ generic_format beta fexp y ->
+ round beta fexp Zceil (x - y) - (x - y) <= y.
+Proof.
+intros fexp Vfexp x y Py Hxy Hly Fx Fy.
+assert (Px := Rlt_trans 0 y x Py Hxy).
+revert Fx.
+unfold generic_format, F2R, scaled_mantissa, cexp; simpl.
+set (mx := Ztrunc (x * bpow (- fexp (mag x)))).
+intro Fx.
+assert (Hfx : (fexp (mag x) < mag x)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+assert (Hfy : (fexp (mag y) < mag y)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+destruct (Rlt_or_le (bpow (mag x - 1)) x) as [Hx|Hx].
+- (* bpow (mag x - 1) < x *)
+ assert (Lxy : mag (x - y) = mag x :> Z);
+ [now apply (mag_minus_separated fexp); [| | | | | |omega]|].
+ assert (Rxy : round beta fexp Zceil (x - y) = x).
+ { unfold round, F2R, scaled_mantissa, cexp; simpl.
+ rewrite Lxy.
+ apply eq_sym; rewrite Fx at 1; apply eq_sym.
+ apply Rmult_eq_compat_r.
+ apply f_equal.
+ rewrite Fx at 1.
+ rewrite Rmult_minus_distr_r.
+ bpow_simplify.
+ apply Zceil_imp.
+ split.
+ - unfold Zminus; rewrite plus_IZR.
+ apply Rplus_lt_compat_l.
+ apply Ropp_lt_contravar; simpl.
+ apply (Rmult_lt_reg_r (bpow (fexp (mag x))));
+ [now apply bpow_gt_0|].
+ rewrite Rmult_1_l; bpow_simplify.
+ apply Rlt_le_trans with (bpow (mag y)).
+ + rewrite <- Rabs_right at 1; [|now apply Rle_ge; apply Rlt_le].
+ apply bpow_mag_gt.
+ + apply bpow_le.
+ omega.
+ - rewrite <- (Rplus_0_r (IZR _)) at 2.
+ apply Rplus_le_compat_l.
+ rewrite <- Ropp_0; apply Ropp_le_contravar.
+ rewrite <- (Rmult_0_r y).
+ apply Rmult_le_compat_l; [now apply Rlt_le|].
+ now apply bpow_ge_0. }
+ rewrite Rxy; ring_simplify.
+ apply Rle_refl.
+- (* x <= bpow (mag x - 1) *)
+ assert (Xpow : x = bpow (mag x - 1)).
+ { apply Rle_antisym; [exact Hx|].
+ destruct (mag x) as (ex, Hex); simpl.
+ rewrite <- (Rabs_right x); [|now apply Rle_ge; apply Rlt_le].
+ apply Hex.
+ now apply Rgt_not_eq. }
+ assert (Lxy : (mag (x - y) = mag x - 1 :> Z)%Z).
+ { apply Zle_antisym.
+ - apply mag_le_bpow.
+ + apply Rminus_eq_contra.
+ now intro Hx'; rewrite Hx' in Hxy; apply (Rlt_irrefl y).
+ + rewrite Rabs_right; lra.
+ - apply (mag_minus_lb beta x y Px Py).
+ omega. }
+ assert (Hfx1 : (fexp (mag x - 1) < mag x - 1)%Z);
+ [now apply (valid_exp_large fexp (mag y)); [|omega]|].
+ assert (Rxy : round beta fexp Zceil (x - y) <= x).
+ { rewrite Xpow at 2.
+ unfold round, F2R, scaled_mantissa, cexp; simpl.
+ rewrite Lxy.
+ apply (Rmult_le_reg_r (bpow (- fexp (mag x - 1)%Z)));
+ [now apply bpow_gt_0|].
+ bpow_simplify.
+ rewrite <- (IZR_Zpower beta (_ - _ - _)); [|omega].
+ apply IZR_le.
+ apply Zceil_glb.
+ rewrite IZR_Zpower; [|omega].
+ rewrite Xpow at 1.
+ rewrite Rmult_minus_distr_r.
+ bpow_simplify.
+ rewrite <- (Rplus_0_r (bpow _)) at 2.
+ apply Rplus_le_compat_l.
+ rewrite <- Ropp_0; apply Ropp_le_contravar.
+ rewrite <- (Rmult_0_r y).
+ apply Rmult_le_compat_l; [now apply Rlt_le|].
+ now apply bpow_ge_0. }
+ lra.
+Qed.
+
+(* mag y <= fexp1 (mag x) - 2 :
+ * mag y <= fexp1 (mag (x - y)) - 2 :
+ * round_round_gt_mid applies. *)
+Lemma round_round_minus_aux2 :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < y -> y < x ->
+ (mag y <= fexp1 (mag x) - 2)%Z ->
+ (mag y <= fexp1 (mag (x - y)) - 2)%Z ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x - y).
+Proof.
+assert (Hbeta : (2 <= beta)%Z).
+{ destruct beta as (beta_val,beta_prop).
+ now apply Zle_bool_imp_le. }
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Py Hxy Hly Hly' Fx Fy.
+assert (Px := Rlt_trans 0 y x Py Hxy).
+destruct Hexp as (_,(_,(_,Hexp4))).
+assert (Hf2 : (fexp2 (mag x) <= fexp1 (mag x))%Z);
+ [now apply Hexp4; omega|].
+assert (Hfx : (fexp1 (mag x) < mag x)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+assert (Bpow2 : bpow (- 2) <= / 2 * / 2).
+{ replace (/2 * /2) with (/4) by field.
+ rewrite (bpow_opp _ 2).
+ apply Rinv_le; [lra|].
+ apply (IZR_le (2 * 2) (beta * (beta * 1))).
+ rewrite Zmult_1_r.
+ now apply Zmult_le_compat; omega. }
+assert (Ly : y < bpow (mag y)).
+{ apply Rabs_lt_inv.
+ apply bpow_mag_gt. }
+unfold round_round_eq.
+apply round_round_gt_mid.
+- exact Vfexp1.
+- exact Vfexp2.
+- lra.
+- apply Hexp4; omega.
+- assert (fexp1 (mag (x - y)) < mag (x - y))%Z; [|omega].
+ apply (valid_exp_large fexp1 (mag x - 1)).
+ + apply (valid_exp_large fexp1 (mag y)); [|omega].
+ now apply mag_generic_gt; [|apply Rgt_not_eq|].
+ + now apply mag_minus_lb; [| |omega].
+- unfold midp'.
+ apply (Rplus_lt_reg_r (/ 2 * ulp beta fexp1 (x - y) - (x - y))).
+ ring_simplify.
+ replace (_ + _) with (round beta fexp1 Zceil (x - y) - (x - y)) by ring.
+ apply Rlt_le_trans with (bpow (fexp1 (mag (x - y)) - 2)).
+ + apply Rle_lt_trans with y;
+ [now apply round_round_minus_aux2_aux; try assumption; omega|].
+ apply (Rlt_le_trans _ _ _ Ly).
+ now apply bpow_le.
+ + rewrite ulp_neq_0;[idtac|now apply sym_not_eq, Rlt_not_eq, Rgt_minus].
+ unfold cexp.
+ replace (_ - 2)%Z with (fexp1 (mag (x - y)) - 1 - 1)%Z by ring.
+ unfold Zminus at 1; rewrite bpow_plus.
+ rewrite Rmult_comm.
+ apply Rmult_le_compat.
+ * now apply bpow_ge_0.
+ * now apply bpow_ge_0.
+ * unfold Raux.bpow, Z.pow_pos; simpl.
+ rewrite Zmult_1_r; apply Rinv_le.
+ lra.
+ now apply IZR_le.
+ * apply bpow_le; omega.
+- intro Hf2'.
+ unfold midp'.
+ apply (Rplus_lt_reg_r (/ 2 * ulp beta fexp1 (x - y) - (x - y)
+ - / 2 * ulp beta fexp2 (x - y))).
+ ring_simplify.
+ replace (_ + _) with (round beta fexp1 Zceil (x - y) - (x - y)) by ring.
+ apply Rle_lt_trans with y;
+ [now apply round_round_minus_aux2_aux; try assumption; omega|].
+ apply (Rlt_le_trans _ _ _ Ly).
+ apply Rle_trans with (bpow (fexp1 (mag (x - y)) - 2));
+ [now apply bpow_le|].
+ replace (_ - 2)%Z with (fexp1 (mag (x - y)) - 1 - 1)%Z by ring.
+ unfold Zminus at 1; rewrite bpow_plus.
+ rewrite <- Rmult_minus_distr_l.
+ rewrite Rmult_comm; apply Rmult_le_compat.
+ + apply bpow_ge_0.
+ + apply bpow_ge_0.
+ + unfold Raux.bpow, Z.pow_pos; simpl.
+ rewrite Zmult_1_r; apply Rinv_le; [lra|].
+ now apply IZR_le.
+ + rewrite 2!ulp_neq_0; try now apply Rgt_not_eq, Rgt_minus.
+ unfold cexp.
+ apply (Rplus_le_reg_r (bpow (fexp2 (mag (x - y))))); ring_simplify.
+ apply Rle_trans with (2 * bpow (fexp1 (mag (x - y)) - 1)).
+ * rewrite double.
+ apply Rplus_le_compat_l.
+ now apply bpow_le.
+ * unfold Zminus; rewrite bpow_plus.
+ rewrite Rmult_comm; rewrite Rmult_assoc.
+ rewrite <- Rmult_1_r.
+ apply Rmult_le_compat_l; [now apply bpow_ge_0|].
+ unfold Raux.bpow, Z.pow_pos; simpl.
+ rewrite Zmult_1_r.
+ apply IZR_le, Rinv_le in Hbeta.
+ simpl in Hbeta.
+ lra.
+ apply Rlt_0_2.
+Qed.
+
+(* round_round_minus_aux{0,1,2} together *)
+Lemma round_round_minus_aux3 :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < y -> y <= x ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x - y).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Py Hyx Fx Fy.
+assert (Px := Rlt_le_trans 0 y x Py Hyx).
+unfold round_round_eq.
+destruct (Req_dec y x) as [Hy|Hy].
+- (* y = x *)
+ rewrite Hy; replace (x - x) with 0 by ring.
+ rewrite round_0.
+ + reflexivity.
+ + now apply valid_rnd_N.
+- (* y < x *)
+ assert (Hyx' : y < x); [lra|].
+ destruct (Zle_or_lt (mag y) (fexp1 (mag x) - 2)) as [Hly|Hly].
+ + (* mag y <= fexp1 (mag x) - 2 *)
+ destruct (Zle_or_lt (mag y) (fexp1 (mag (x - y)) - 2))
+ as [Hly'|Hly'].
+ * (* mag y <= fexp1 (mag (x - y)) - 2 *)
+ now apply round_round_minus_aux2.
+ * (* fexp1 (mag (x - y)) - 2 < mag y *)
+ { rewrite (round_generic beta fexp2).
+ - reflexivity.
+ - now apply valid_rnd_N.
+ - assert (Hf1 : (fexp1 (mag (x - y)) - 1 <= mag y)%Z); [omega|].
+ now apply (round_round_minus_aux1 fexp1). }
+ + rewrite (round_generic beta fexp2).
+ * reflexivity.
+ * now apply valid_rnd_N.
+ * assert (Hf1 : (fexp1 (mag x) - 1 <= mag y)%Z); [omega|].
+ now apply (round_round_minus_aux0 fexp1).
+Qed.
+
+Lemma round_round_minus_aux :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 <= x -> 0 <= y ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x - y).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Nnx Nny Fx Fy.
+unfold round_round_eq.
+destruct (Req_dec x 0) as [Zx|Nzx].
+- (* x = 0 *)
+ rewrite Zx; unfold Rminus; rewrite Rplus_0_l.
+ do 3 rewrite round_N_opp.
+ rewrite (round_generic beta fexp2).
+ * reflexivity.
+ * now apply valid_rnd_N.
+ * apply (generic_inclusion_mag beta fexp1).
+ destruct Hexp as (_,(_,(_,Hexp4))).
+ now intros _; apply Hexp4; omega.
+ exact Fy.
+- (* x <> 0 *)
+ destruct (Req_dec y 0) as [Zy|Nzy].
+ + (* y = 0 *)
+ rewrite Zy; unfold Rminus; rewrite Ropp_0; rewrite Rplus_0_r.
+ rewrite (round_generic beta fexp2).
+ * reflexivity.
+ * now apply valid_rnd_N.
+ * apply (generic_inclusion_mag beta fexp1).
+ destruct Hexp as (_,(_,(_,Hexp4))).
+ now intros _; apply Hexp4; omega.
+ exact Fx.
+ + (* y <> 0 *)
+ assert (Px : 0 < x); [lra|].
+ assert (Py : 0 < y); [lra|].
+ destruct (Rlt_or_le x y) as [H|H].
+ * (* x < y *)
+ apply Rlt_le in H.
+ replace (x - y) with (- (y - x)) by ring.
+ do 3 rewrite round_N_opp.
+ apply Ropp_eq_compat.
+ now apply round_round_minus_aux3.
+ * (* y <= x *)
+ now apply round_round_minus_aux3.
+Qed.
+
+Lemma round_round_plus :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_hyp fexp1 fexp2 ->
+ forall x y,
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x + y).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Fx Fy.
+unfold round_round_eq.
+destruct (Rlt_or_le x 0) as [Sx|Sx]; destruct (Rlt_or_le y 0) as [Sy|Sy].
+- (* x < 0, y < 0 *)
+ replace (x + y) with (- (- x - y)); [|ring].
+ do 3 rewrite round_N_opp.
+ apply Ropp_eq_compat.
+ assert (Px : 0 <= - x); [lra|].
+ assert (Py : 0 <= - y); [lra|].
+ apply generic_format_opp in Fx.
+ apply generic_format_opp in Fy.
+ now apply round_round_plus_aux.
+- (* x < 0, 0 <= y *)
+ replace (x + y) with (y - (- x)); [|ring].
+ assert (Px : 0 <= - x); [lra|].
+ apply generic_format_opp in Fx.
+ now apply round_round_minus_aux.
+- (* 0 <= x, y < 0 *)
+ replace (x + y) with (x - (- y)); [|ring].
+ assert (Py : 0 <= - y); [lra|].
+ apply generic_format_opp in Fy.
+ now apply round_round_minus_aux.
+- (* 0 <= x, 0 <= y *)
+ now apply round_round_plus_aux.
+Qed.
+
+Lemma round_round_minus :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_hyp fexp1 fexp2 ->
+ forall x y,
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x - y).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Fx Fy.
+unfold Rminus.
+apply generic_format_opp in Fy.
+now apply round_round_plus.
+Qed.
+
+Section Double_round_plus_FLX.
+
+Variable prec : Z.
+Variable prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FLX_round_round_plus_hyp :
+ (2 * prec + 1 <= prec')%Z ->
+ round_round_plus_hyp (FLX_exp prec) (FLX_exp prec').
+Proof.
+intros Hprec.
+unfold FLX_exp.
+unfold round_round_plus_hyp; split; [|split; [|split]];
+intros ex ey; try omega.
+unfold Prec_gt_0 in prec_gt_0_.
+omega.
+Qed.
+
+Theorem round_round_plus_FLX :
+ forall choice1 choice2,
+ (2 * prec + 1 <= prec')%Z ->
+ forall x y,
+ FLX_format beta prec x -> FLX_format beta prec y ->
+ round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x + y).
+Proof.
+intros choice1 choice2 Hprec x y Fx Fy.
+apply round_round_plus.
+- now apply FLX_exp_valid.
+- now apply FLX_exp_valid.
+- now apply FLX_round_round_plus_hyp.
+- now apply generic_format_FLX.
+- now apply generic_format_FLX.
+Qed.
+
+Theorem round_round_minus_FLX :
+ forall choice1 choice2,
+ (2 * prec + 1 <= prec')%Z ->
+ forall x y,
+ FLX_format beta prec x -> FLX_format beta prec y ->
+ round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x - y).
+Proof.
+intros choice1 choice2 Hprec x y Fx Fy.
+apply round_round_minus.
+- now apply FLX_exp_valid.
+- now apply FLX_exp_valid.
+- now apply FLX_round_round_plus_hyp.
+- now apply generic_format_FLX.
+- now apply generic_format_FLX.
+Qed.
+
+End Double_round_plus_FLX.
+
+Section Double_round_plus_FLT.
+
+Variable emin prec : Z.
+Variable emin' prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FLT_round_round_plus_hyp :
+ (emin' <= emin)%Z -> (2 * prec + 1 <= prec')%Z ->
+ round_round_plus_hyp (FLT_exp emin prec) (FLT_exp emin' prec').
+Proof.
+intros Hemin Hprec.
+unfold FLT_exp.
+unfold round_round_plus_hyp; split; [|split; [|split]]; intros ex ey.
+- generalize (Zmax_spec (ex + 1 - prec) emin).
+ generalize (Zmax_spec (ex - prec') emin').
+ generalize (Zmax_spec (ey - prec) emin).
+ omega.
+- generalize (Zmax_spec (ex - 1 - prec) emin).
+ generalize (Zmax_spec (ex - prec') emin').
+ generalize (Zmax_spec (ey - prec) emin).
+ omega.
+- generalize (Zmax_spec (ex - prec) emin).
+ generalize (Zmax_spec (ex - prec') emin').
+ generalize (Zmax_spec (ey - prec) emin).
+ omega.
+- unfold Prec_gt_0 in prec_gt_0_.
+ generalize (Zmax_spec (ex - prec') emin').
+ generalize (Zmax_spec (ey - prec) emin).
+ omega.
+Qed.
+
+Theorem round_round_plus_FLT :
+ forall choice1 choice2,
+ (emin' <= emin)%Z -> (2 * prec + 1 <= prec')%Z ->
+ forall x y,
+ FLT_format beta emin prec x -> FLT_format beta emin prec y ->
+ round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec')
+ choice1 choice2 (x + y).
+Proof.
+intros choice1 choice2 Hemin Hprec x y Fx Fy.
+apply round_round_plus.
+- now apply FLT_exp_valid.
+- now apply FLT_exp_valid.
+- now apply FLT_round_round_plus_hyp.
+- now apply generic_format_FLT.
+- now apply generic_format_FLT.
+Qed.
+
+Theorem round_round_minus_FLT :
+ forall choice1 choice2,
+ (emin' <= emin)%Z -> (2 * prec + 1 <= prec')%Z ->
+ forall x y,
+ FLT_format beta emin prec x -> FLT_format beta emin prec y ->
+ round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec')
+ choice1 choice2 (x - y).
+Proof.
+intros choice1 choice2 Hemin Hprec x y Fx Fy.
+apply round_round_minus.
+- now apply FLT_exp_valid.
+- now apply FLT_exp_valid.
+- now apply FLT_round_round_plus_hyp.
+- now apply generic_format_FLT.
+- now apply generic_format_FLT.
+Qed.
+
+End Double_round_plus_FLT.
+
+Section Double_round_plus_FTZ.
+
+Variable emin prec : Z.
+Variable emin' prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FTZ_round_round_plus_hyp :
+ (emin' + prec' <= emin + 1)%Z -> (2 * prec + 1 <= prec')%Z ->
+ round_round_plus_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec').
+Proof.
+intros Hemin Hprec.
+unfold FTZ_exp.
+unfold Prec_gt_0 in *.
+unfold round_round_plus_hyp; split; [|split; [|split]]; intros ex ey.
+- destruct (Z.ltb_spec (ex + 1 - prec) emin);
+ destruct (Z.ltb_spec (ex - prec') emin');
+ destruct (Z.ltb_spec (ey - prec) emin);
+ omega.
+- destruct (Z.ltb_spec (ex - 1 - prec) emin);
+ destruct (Z.ltb_spec (ex - prec') emin');
+ destruct (Z.ltb_spec (ey - prec) emin);
+ omega.
+- destruct (Z.ltb_spec (ex - prec) emin);
+ destruct (Z.ltb_spec (ex - prec') emin');
+ destruct (Z.ltb_spec (ey - prec) emin);
+ omega.
+- destruct (Z.ltb_spec (ex - prec') emin');
+ destruct (Z.ltb_spec (ey - prec) emin);
+ omega.
+Qed.
+
+Theorem round_round_plus_FTZ :
+ forall choice1 choice2,
+ (emin' + prec' <= emin + 1)%Z -> (2 * prec + 1 <= prec')%Z ->
+ forall x y,
+ FTZ_format beta emin prec x -> FTZ_format beta emin prec y ->
+ round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec')
+ choice1 choice2 (x + y).
+Proof.
+intros choice1 choice2 Hemin Hprec x y Fx Fy.
+apply round_round_plus.
+- now apply FTZ_exp_valid.
+- now apply FTZ_exp_valid.
+- now apply FTZ_round_round_plus_hyp.
+- now apply generic_format_FTZ.
+- now apply generic_format_FTZ.
+Qed.
+
+Theorem round_round_minus_FTZ :
+ forall choice1 choice2,
+ (emin' + prec' <= emin + 1)%Z -> (2 * prec + 1 <= prec')%Z ->
+ forall x y,
+ FTZ_format beta emin prec x -> FTZ_format beta emin prec y ->
+ round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec')
+ choice1 choice2 (x - y).
+Proof.
+intros choice1 choice2 Hemin Hprec x y Fx Fy.
+apply round_round_minus.
+- now apply FTZ_exp_valid.
+- now apply FTZ_exp_valid.
+- now apply FTZ_round_round_plus_hyp.
+- now apply generic_format_FTZ.
+- now apply generic_format_FTZ.
+Qed.
+
+End Double_round_plus_FTZ.
+
+Section Double_round_plus_radix_ge_3.
+
+Definition round_round_plus_radix_ge_3_hyp fexp1 fexp2 :=
+ (forall ex ey, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z)
+ /\ (forall ex ey, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z)
+ /\ (forall ex ey, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z)
+ /\ (forall ex ey, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z).
+
+(* fexp1 (mag x) <= mag y :
+ * addition is exact in the largest precision (fexp2). *)
+Lemma round_round_plus_radix_ge_3_aux0 :
+ forall (fexp1 fexp2 : Z -> Z), Valid_exp fexp1 ->
+ round_round_plus_radix_ge_3_hyp fexp1 fexp2 ->
+ forall x y,
+ (0 < y)%R -> (y <= x)%R ->
+ (fexp1 (mag x) <= mag y)%Z ->
+ generic_format beta fexp1 x -> generic_format beta fexp1 y ->
+ generic_format beta fexp2 (x + y).
+Proof.
+intros fexp1 fexp2 Vfexp1 Hexp x y Py Hyx Hln Fx Fy.
+assert (Px := Rlt_le_trans 0 y x Py Hyx).
+assert (Nny : (0 <= y)%R); [now apply Rlt_le|].
+destruct Hexp as (_,(Hexp2,(Hexp3,Hexp4))).
+destruct (Z.le_gt_cases (mag y) (fexp1 (mag x))) as [Hle|Hgt].
+- (* mag y <= fexp1 (mag x) *)
+ assert (Lxy : mag (x + y) = mag x :> Z);
+ [now apply (mag_plus_separated fexp1)|].
+ apply (round_round_plus_aux0_aux fexp1);
+ [| |assumption|assumption]; rewrite Lxy.
+ + now apply Hexp4; omega.
+ + now apply Hexp3; omega.
+- (* fexp1 (mag x) < mag y *)
+ apply (round_round_plus_aux0_aux fexp1); [| |assumption|assumption].
+ destruct (mag_plus_disj x y Py Hyx) as [Lxy|Lxy]; rewrite Lxy.
+ + now apply Hexp4; omega.
+ + apply Hexp2; apply (mag_le beta y x Py) in Hyx.
+ replace (_ - _)%Z with (mag x : Z) by ring.
+ omega.
+ + destruct (mag_plus_disj x y Py Hyx) as [Lxy|Lxy]; rewrite Lxy.
+ * now apply Hexp3; omega.
+ * apply Hexp2.
+ replace (_ - _)%Z with (mag x : Z) by ring.
+ omega.
+Qed.
+
+(* mag y <= fexp1 (mag x) - 1 : round_round_lt_mid applies. *)
+Lemma round_round_plus_radix_ge_3_aux1 :
+ (3 <= beta)%Z ->
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_radix_ge_3_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < x -> 0 < y ->
+ (mag y <= fexp1 (mag x) - 1)%Z ->
+ generic_format beta fexp1 x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x + y).
+Proof.
+intros Hbeta fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Px Py Hly Fx.
+assert (Lxy : mag (x + y) = mag x :> Z);
+ [now apply (mag_plus_separated fexp1); [|apply Rlt_le| |omega]|].
+destruct Hexp as (_,(_,(_,Hexp4))).
+assert (Hf2 : (fexp2 (mag x) <= fexp1 (mag x))%Z);
+ [now apply Hexp4; omega|].
+assert (Bpow3 : bpow (- 1) <= / 3).
+{ unfold Raux.bpow, Z.pow_pos; simpl.
+ rewrite Zmult_1_r.
+ apply Rinv_le; [lra|].
+ now apply IZR_le. }
+assert (P1 : (0 < 1)%Z) by omega.
+unfold round_round_eq.
+apply round_round_lt_mid.
+- exact Vfexp1.
+- exact Vfexp2.
+- lra.
+- now rewrite Lxy.
+- rewrite Lxy.
+ assert (fexp1 (mag x) < mag x)%Z; [|omega].
+ now apply mag_generic_gt; [|apply Rgt_not_eq|].
+- unfold midp.
+ apply (Rplus_lt_reg_r (- round beta fexp1 Zfloor (x + y))).
+ apply (Rlt_le_trans _ _ _ (proj2 (round_round_plus_aux1_aux 1 P1 fexp1 x y Px
+ Py Hly Lxy Fx))).
+ ring_simplify.
+ rewrite ulp_neq_0;[idtac|now apply Rgt_not_eq, Rplus_lt_0_compat].
+ unfold cexp; rewrite Lxy.
+ apply (Rmult_le_reg_r (bpow (- fexp1 (mag x))));
+ [now apply bpow_gt_0|].
+ bpow_simplify.
+ apply (Rle_trans _ _ _ Bpow3); lra.
+- rewrite ulp_neq_0;[idtac|now apply Rgt_not_eq, Rplus_lt_0_compat].
+ unfold round, F2R, scaled_mantissa, cexp; simpl; rewrite Lxy.
+ intro Hf2'.
+ unfold midp.
+ apply (Rplus_lt_reg_r (- round beta fexp1 Zfloor (x + y))); ring_simplify.
+ rewrite <- Rmult_minus_distr_l.
+ apply (Rlt_le_trans _ _ _ (proj2 (round_round_plus_aux1_aux 1 P1 fexp1 x y Px
+ Py Hly Lxy Fx))).
+ rewrite ulp_neq_0;[idtac|now apply Rgt_not_eq, Rplus_lt_0_compat].
+ unfold cexp; rewrite Lxy.
+ apply (Rmult_le_reg_r (bpow (- fexp1 (mag x))));
+ [now apply bpow_gt_0|].
+ rewrite (Rmult_assoc (/ 2)).
+ rewrite Rmult_minus_distr_r.
+ bpow_simplify.
+ apply (Rle_trans _ _ _ Bpow3).
+ apply Rle_trans with (/ 2 * (2 / 3)); [lra|].
+ apply Rmult_le_compat_l; [lra|].
+ apply (Rplus_le_reg_r (- 1)); ring_simplify.
+ replace (_ - _) with (- (/ 3)) by lra.
+ apply Ropp_le_contravar.
+ now apply Rle_trans with (bpow (- 1)); [apply bpow_le; omega|].
+Qed.
+
+(* round_round_plus_radix_ge_3_aux{0,1} together *)
+Lemma round_round_plus_radix_ge_3_aux2 :
+ (3 <= beta)%Z ->
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_radix_ge_3_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < y -> y <= x ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x + y).
+Proof.
+intros Hbeta fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Py Hyx Fx Fy.
+assert (Px := Rlt_le_trans 0 y x Py Hyx).
+unfold round_round_eq.
+destruct (Zle_or_lt (mag y) (fexp1 (mag x) - 1)) as [Hly|Hly].
+- (* mag y <= fexp1 (mag x) - 1 *)
+ now apply round_round_plus_radix_ge_3_aux1.
+- (* fexp1 (mag x) - 1 < mag y *)
+ rewrite (round_generic beta fexp2).
+ + reflexivity.
+ + now apply valid_rnd_N.
+ + assert (Hf1 : (fexp1 (mag x) <= mag y)%Z); [omega|].
+ now apply (round_round_plus_radix_ge_3_aux0 fexp1).
+Qed.
+
+Lemma round_round_plus_radix_ge_3_aux :
+ (3 <= beta)%Z ->
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_radix_ge_3_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 <= x -> 0 <= y ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x + y).
+Proof.
+intros Hbeta fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Nnx Nny Fx Fy.
+unfold round_round_eq.
+destruct (Req_dec x 0) as [Zx|Nzx].
+- (* x = 0 *)
+ destruct Hexp as (_,(_,(_,Hexp4))).
+ rewrite Zx; rewrite Rplus_0_l.
+ rewrite (round_generic beta fexp2).
+ + reflexivity.
+ + now apply valid_rnd_N.
+ + apply (generic_inclusion_mag beta fexp1).
+ now intros _; apply Hexp4; omega.
+ exact Fy.
+- (* x <> 0 *)
+ destruct (Req_dec y 0) as [Zy|Nzy].
+ + (* y = 0 *)
+ destruct Hexp as (_,(_,(_,Hexp4))).
+ rewrite Zy; rewrite Rplus_0_r.
+ rewrite (round_generic beta fexp2).
+ * reflexivity.
+ * now apply valid_rnd_N.
+ * apply (generic_inclusion_mag beta fexp1).
+ now intros _; apply Hexp4; omega.
+ exact Fx.
+ + (* y <> 0 *)
+ assert (Px : 0 < x); [lra|].
+ assert (Py : 0 < y); [lra|].
+ destruct (Rlt_or_le x y) as [H|H].
+ * (* x < y *)
+ apply Rlt_le in H.
+ rewrite Rplus_comm.
+ now apply round_round_plus_radix_ge_3_aux2.
+ * now apply round_round_plus_radix_ge_3_aux2.
+Qed.
+
+(* fexp1 (mag x) <= mag y :
+ * substraction is exact in the largest precision (fexp2). *)
+Lemma round_round_minus_radix_ge_3_aux0 :
+ forall (fexp1 fexp2 : Z -> Z),
+ round_round_plus_radix_ge_3_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < y -> y < x ->
+ (fexp1 (mag x) <= mag y)%Z ->
+ generic_format beta fexp1 x -> generic_format beta fexp1 y ->
+ generic_format beta fexp2 (x - y).
+Proof.
+intros fexp1 fexp2 Hexp x y Py Hyx Hln Fx Fy.
+assert (Px := Rlt_trans 0 y x Py Hyx).
+destruct Hexp as (Hexp1,(_,(Hexp3,Hexp4))).
+assert (Lyx : (mag y <= mag x)%Z);
+ [now apply mag_le; [|apply Rlt_le]|].
+destruct (Z.lt_ge_cases (mag x - 2) (mag y)) as [Hlt|Hge].
+- (* mag x - 2 < mag y *)
+ assert (Hor : (mag y = mag x :> Z)
+ \/ (mag y = mag x - 1 :> Z)%Z); [omega|].
+ destruct Hor as [Heq|Heqm1].
+ + (* mag y = mag x *)
+ apply (round_round_minus_aux0_aux fexp1); [| |exact Fx|exact Fy].
+ * apply Hexp4.
+ apply Z.le_trans with (mag (x - y)); [omega|].
+ now apply mag_minus.
+ * rewrite Heq.
+ apply Hexp4.
+ apply Z.le_trans with (mag (x - y)); [omega|].
+ now apply mag_minus.
+ + (* mag y = mag x - 1 *)
+ apply (round_round_minus_aux0_aux fexp1); [| |exact Fx|exact Fy].
+ * apply Hexp4.
+ apply Z.le_trans with (mag (x - y)); [omega|].
+ now apply mag_minus.
+ * rewrite Heqm1.
+ apply Hexp4.
+ apply Zplus_le_compat_r.
+ now apply mag_minus.
+- (* mag y <= mag x - 2 *)
+ destruct (mag_minus_disj x y Px Py Hge) as [Lxmy|Lxmy].
+ + (* mag (x - y) = mag x *)
+ apply (round_round_minus_aux0_aux fexp1); [| |exact Fx|exact Fy].
+ * apply Hexp4.
+ omega.
+ * now rewrite Lxmy; apply Hexp3.
+ + (* mag (x - y) = mag x - 1 *)
+ apply (round_round_minus_aux0_aux fexp1); [| |exact Fx|exact Fy];
+ rewrite Lxmy.
+ * apply Hexp1.
+ replace (_ + _)%Z with (mag x : Z); [|ring].
+ now apply Z.le_trans with (mag y).
+ * apply Hexp1.
+ now replace (_ + _)%Z with (mag x : Z); [|ring].
+Qed.
+
+(* mag y <= fexp1 (mag x) - 1,
+ * fexp1 (mag (x - y)) <= mag y :
+ * substraction is exact in the largest precision (fexp2). *)
+Lemma round_round_minus_radix_ge_3_aux1 :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ round_round_plus_radix_ge_3_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < y -> y < x ->
+ (mag y <= fexp1 (mag x) - 1)%Z ->
+ (fexp1 (mag (x - y)) <= mag y)%Z ->
+ generic_format beta fexp1 x -> generic_format beta fexp1 y ->
+ generic_format beta fexp2 (x - y).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 Hexp x y Py Hyx Hln Hln' Fx Fy.
+assert (Px := Rlt_trans 0 y x Py Hyx).
+destruct Hexp as (Hexp1,(Hexp2,(Hexp3,Hexp4))).
+assert (Lyx : (mag y <= mag x)%Z);
+ [now apply mag_le; [|apply Rlt_le]|].
+assert (Hfx : (fexp1 (mag x) < mag x)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+assert (Hfy : (fexp1 (mag y) < mag y)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+apply (round_round_minus_aux0_aux fexp1); [| |exact Fx|exact Fy].
+- apply Z.le_trans with (fexp1 (mag (x - y))).
+ + apply Hexp4; omega.
+ + omega.
+- now apply Hexp3.
+Qed.
+
+(* mag y <= fexp1 (mag x) - 1 :
+ * mag y <= fexp1 (mag (x - y)) - 1 :
+ * round_round_gt_mid applies. *)
+Lemma round_round_minus_radix_ge_3_aux2 :
+ (3 <= beta)%Z ->
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_radix_ge_3_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < y -> y < x ->
+ (mag y <= fexp1 (mag x) - 1)%Z ->
+ (mag y <= fexp1 (mag (x - y)) - 1)%Z ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x - y).
+Proof.
+intros Hbeta fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Py Hxy Hly Hly' Fx Fy.
+assert (Px := Rlt_trans 0 y x Py Hxy).
+destruct Hexp as (_,(_,(_,Hexp4))).
+assert (Hf2 : (fexp2 (mag x) <= fexp1 (mag x))%Z);
+ [now apply Hexp4; omega|].
+assert (Hfx : (fexp1 (mag x) < mag x)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+assert (Bpow3 : bpow (- 1) <= / 3).
+{ unfold Raux.bpow, Z.pow_pos; simpl.
+ rewrite Zmult_1_r.
+ apply Rinv_le; [lra|].
+ now apply IZR_le. }
+assert (Ly : y < bpow (mag y)).
+{ apply Rabs_lt_inv.
+ apply bpow_mag_gt. }
+unfold round_round_eq.
+apply round_round_gt_mid.
+- exact Vfexp1.
+- exact Vfexp2.
+- lra.
+- apply Hexp4; omega.
+- assert (fexp1 (mag (x - y)) < mag (x - y))%Z; [|omega].
+ apply (valid_exp_large fexp1 (mag x - 1)).
+ + apply (valid_exp_large fexp1 (mag y)); [|omega].
+ now apply mag_generic_gt; [|apply Rgt_not_eq|].
+ + now apply mag_minus_lb; [| |omega].
+- unfold midp'.
+ apply (Rplus_lt_reg_r (/ 2 * ulp beta fexp1 (x - y) - (x - y))).
+ ring_simplify.
+ replace (_ + _) with (round beta fexp1 Zceil (x - y) - (x - y)) by ring.
+ apply Rlt_le_trans with (bpow (fexp1 (mag (x - y)) - 1)).
+ + apply Rle_lt_trans with y;
+ [now apply round_round_minus_aux2_aux|].
+ apply (Rlt_le_trans _ _ _ Ly).
+ now apply bpow_le.
+ + rewrite ulp_neq_0;[idtac|now apply Rgt_not_eq, Rgt_minus].
+ unfold cexp.
+ unfold Zminus at 1; rewrite bpow_plus.
+ rewrite Rmult_comm.
+ apply Rmult_le_compat_r; [now apply bpow_ge_0|].
+ unfold Raux.bpow, Z.pow_pos; simpl.
+ rewrite Zmult_1_r; apply Rinv_le; [lra|].
+ now apply IZR_le; omega.
+- intro Hf2'.
+ unfold midp'.
+ apply (Rplus_lt_reg_r (/ 2 * (ulp beta fexp1 (x - y)
+ - ulp beta fexp2 (x - y)) - (x - y))).
+ ring_simplify; rewrite <- Rmult_minus_distr_l.
+ replace (_ + _) with (round beta fexp1 Zceil (x - y) - (x - y)) by ring.
+ apply Rle_lt_trans with y;
+ [now apply round_round_minus_aux2_aux|].
+ apply (Rlt_le_trans _ _ _ Ly).
+ apply Rle_trans with (bpow (fexp1 (mag (x - y)) - 1));
+ [now apply bpow_le|].
+ rewrite 2!ulp_neq_0; try now apply Rgt_not_eq, Rgt_minus.
+ unfold cexp.
+ apply (Rmult_le_reg_r (bpow (- fexp1 (mag (x - y)))));
+ [now apply bpow_gt_0|].
+ rewrite Rmult_assoc.
+ rewrite Rmult_minus_distr_r.
+ bpow_simplify.
+ apply Rle_trans with (/ 3).
+ + unfold Raux.bpow, Z.pow_pos; simpl.
+ rewrite Zmult_1_r; apply Rinv_le; [lra|].
+ now apply IZR_le.
+ + replace (/ 3) with (/ 2 * (2 / 3)) by field.
+ apply Rmult_le_compat_l; [lra|].
+ apply (Rplus_le_reg_r (- 1)); ring_simplify.
+ replace (_ - _) with (- / 3) by field.
+ apply Ropp_le_contravar.
+ apply Rle_trans with (bpow (- 1)).
+ * apply bpow_le; omega.
+ * unfold Raux.bpow, Z.pow_pos; simpl.
+ rewrite Zmult_1_r; apply Rinv_le; [lra|].
+ now apply IZR_le.
+Qed.
+
+(* round_round_minus_aux{0,1,2} together *)
+Lemma round_round_minus_radix_ge_3_aux3 :
+ (3 <= beta)%Z ->
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_radix_ge_3_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < y -> y <= x ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x - y).
+Proof.
+intros Hbeta fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Py Hyx Fx Fy.
+assert (Px := Rlt_le_trans 0 y x Py Hyx).
+unfold round_round_eq.
+destruct (Req_dec y x) as [Hy|Hy].
+- (* y = x *)
+ rewrite Hy; replace (x - x) with 0 by ring.
+ rewrite round_0.
+ + reflexivity.
+ + now apply valid_rnd_N.
+- (* y < x *)
+ assert (Hyx' : y < x); [lra|].
+ destruct (Zle_or_lt (mag y) (fexp1 (mag x) - 1)) as [Hly|Hly].
+ + (* mag y <= fexp1 (mag x) - 1 *)
+ destruct (Zle_or_lt (mag y) (fexp1 (mag (x - y)) - 1))
+ as [Hly'|Hly'].
+ * (* mag y <= fexp1 (mag (x - y)) - 1 *)
+ now apply round_round_minus_radix_ge_3_aux2.
+ * (* fexp1 (mag (x - y)) - 1 < mag y *)
+ { rewrite (round_generic beta fexp2).
+ - reflexivity.
+ - now apply valid_rnd_N.
+ - assert (Hf1 : (fexp1 (mag (x - y)) <= mag y)%Z); [omega|].
+ now apply (round_round_minus_radix_ge_3_aux1 fexp1). }
+ + rewrite (round_generic beta fexp2).
+ * reflexivity.
+ * now apply valid_rnd_N.
+ * assert (Hf1 : (fexp1 (mag x) <= mag y)%Z); [omega|].
+ now apply (round_round_minus_radix_ge_3_aux0 fexp1).
+Qed.
+
+Lemma round_round_minus_radix_ge_3_aux :
+ (3 <= beta)%Z ->
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_radix_ge_3_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 <= x -> 0 <= y ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x - y).
+Proof.
+intros Hbeta fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Nnx Nny Fx Fy.
+unfold round_round_eq.
+destruct (Req_dec x 0) as [Zx|Nzx].
+- (* x = 0 *)
+ rewrite Zx; unfold Rminus; rewrite Rplus_0_l.
+ do 3 rewrite round_N_opp.
+ rewrite (round_generic beta fexp2).
+ * reflexivity.
+ * now apply valid_rnd_N.
+ * apply (generic_inclusion_mag beta fexp1).
+ destruct Hexp as (_,(_,(_,Hexp4))).
+ now intros _; apply Hexp4; omega.
+ exact Fy.
+- (* x <> 0 *)
+ destruct (Req_dec y 0) as [Zy|Nzy].
+ + (* y = 0 *)
+ rewrite Zy; unfold Rminus; rewrite Ropp_0; rewrite Rplus_0_r.
+ rewrite (round_generic beta fexp2).
+ * reflexivity.
+ * now apply valid_rnd_N.
+ * apply (generic_inclusion_mag beta fexp1).
+ destruct Hexp as (_,(_,(_,Hexp4))).
+ now intros _; apply Hexp4; omega.
+ exact Fx.
+ + (* y <> 0 *)
+ assert (Px : 0 < x); [lra|].
+ assert (Py : 0 < y); [lra|].
+ destruct (Rlt_or_le x y) as [H|H].
+ * (* x < y *)
+ apply Rlt_le in H.
+ replace (x - y) with (- (y - x)) by ring.
+ do 3 rewrite round_N_opp.
+ apply Ropp_eq_compat.
+ now apply round_round_minus_radix_ge_3_aux3.
+ * (* y <= x *)
+ now apply round_round_minus_radix_ge_3_aux3.
+Qed.
+
+Lemma round_round_plus_radix_ge_3 :
+ (3 <= beta)%Z ->
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_radix_ge_3_hyp fexp1 fexp2 ->
+ forall x y,
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x + y).
+Proof.
+intros Hbeta fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Fx Fy.
+unfold round_round_eq.
+destruct (Rlt_or_le x 0) as [Sx|Sx]; destruct (Rlt_or_le y 0) as [Sy|Sy].
+- (* x < 0, y < 0 *)
+ replace (x + y) with (- (- x - y)); [|ring].
+ do 3 rewrite round_N_opp.
+ apply Ropp_eq_compat.
+ assert (Px : 0 <= - x); [lra|].
+ assert (Py : 0 <= - y); [lra|].
+ apply generic_format_opp in Fx.
+ apply generic_format_opp in Fy.
+ now apply round_round_plus_radix_ge_3_aux.
+- (* x < 0, 0 <= y *)
+ replace (x + y) with (y - (- x)); [|ring].
+ assert (Px : 0 <= - x); [lra|].
+ apply generic_format_opp in Fx.
+ now apply round_round_minus_radix_ge_3_aux.
+- (* 0 <= x, y < 0 *)
+ replace (x + y) with (x - (- y)); [|ring].
+ assert (Py : 0 <= - y); [lra|].
+ apply generic_format_opp in Fy.
+ now apply round_round_minus_radix_ge_3_aux.
+- (* 0 <= x, 0 <= y *)
+ now apply round_round_plus_radix_ge_3_aux.
+Qed.
+
+Lemma round_round_minus_radix_ge_3 :
+ (3 <= beta)%Z ->
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_plus_radix_ge_3_hyp fexp1 fexp2 ->
+ forall x y,
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x - y).
+Proof.
+intros Hbeta fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Fx Fy.
+unfold Rminus.
+apply generic_format_opp in Fy.
+now apply round_round_plus_radix_ge_3.
+Qed.
+
+Section Double_round_plus_radix_ge_3_FLX.
+
+Variable prec : Z.
+Variable prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FLX_round_round_plus_radix_ge_3_hyp :
+ (2 * prec <= prec')%Z ->
+ round_round_plus_radix_ge_3_hyp (FLX_exp prec) (FLX_exp prec').
+Proof.
+intros Hprec.
+unfold FLX_exp.
+unfold round_round_plus_radix_ge_3_hyp; split; [|split; [|split]];
+intros ex ey; try omega.
+unfold Prec_gt_0 in prec_gt_0_.
+omega.
+Qed.
+
+Theorem round_round_plus_radix_ge_3_FLX :
+ (3 <= beta)%Z ->
+ forall choice1 choice2,
+ (2 * prec <= prec')%Z ->
+ forall x y,
+ FLX_format beta prec x -> FLX_format beta prec y ->
+ round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x + y).
+Proof.
+intros Hbeta choice1 choice2 Hprec x y Fx Fy.
+apply round_round_plus_radix_ge_3.
+- exact Hbeta.
+- now apply FLX_exp_valid.
+- now apply FLX_exp_valid.
+- now apply FLX_round_round_plus_radix_ge_3_hyp.
+- now apply generic_format_FLX.
+- now apply generic_format_FLX.
+Qed.
+
+Theorem round_round_minus_radix_ge_3_FLX :
+ (3 <= beta)%Z ->
+ forall choice1 choice2,
+ (2 * prec <= prec')%Z ->
+ forall x y,
+ FLX_format beta prec x -> FLX_format beta prec y ->
+ round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x - y).
+Proof.
+intros Hbeta choice1 choice2 Hprec x y Fx Fy.
+apply round_round_minus_radix_ge_3.
+- exact Hbeta.
+- now apply FLX_exp_valid.
+- now apply FLX_exp_valid.
+- now apply FLX_round_round_plus_radix_ge_3_hyp.
+- now apply generic_format_FLX.
+- now apply generic_format_FLX.
+Qed.
+
+End Double_round_plus_radix_ge_3_FLX.
+
+Section Double_round_plus_radix_ge_3_FLT.
+
+Variable emin prec : Z.
+Variable emin' prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FLT_round_round_plus_radix_ge_3_hyp :
+ (emin' <= emin)%Z -> (2 * prec <= prec')%Z ->
+ round_round_plus_radix_ge_3_hyp (FLT_exp emin prec) (FLT_exp emin' prec').
+Proof.
+intros Hemin Hprec.
+unfold FLT_exp.
+unfold round_round_plus_radix_ge_3_hyp; split; [|split; [|split]]; intros ex ey.
+- generalize (Zmax_spec (ex + 1 - prec) emin).
+ generalize (Zmax_spec (ex - prec') emin').
+ generalize (Zmax_spec (ey - prec) emin).
+ omega.
+- generalize (Zmax_spec (ex - 1 - prec) emin).
+ generalize (Zmax_spec (ex - prec') emin').
+ generalize (Zmax_spec (ey - prec) emin).
+ omega.
+- generalize (Zmax_spec (ex - prec) emin).
+ generalize (Zmax_spec (ex - prec') emin').
+ generalize (Zmax_spec (ey - prec) emin).
+ omega.
+- unfold Prec_gt_0 in prec_gt_0_.
+ generalize (Zmax_spec (ex - prec') emin').
+ generalize (Zmax_spec (ey - prec) emin).
+ omega.
+Qed.
+
+Theorem round_round_plus_radix_ge_3_FLT :
+ (3 <= beta)%Z ->
+ forall choice1 choice2,
+ (emin' <= emin)%Z -> (2 * prec <= prec')%Z ->
+ forall x y,
+ FLT_format beta emin prec x -> FLT_format beta emin prec y ->
+ round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec')
+ choice1 choice2 (x + y).
+Proof.
+intros Hbeta choice1 choice2 Hemin Hprec x y Fx Fy.
+apply round_round_plus_radix_ge_3.
+- exact Hbeta.
+- now apply FLT_exp_valid.
+- now apply FLT_exp_valid.
+- now apply FLT_round_round_plus_radix_ge_3_hyp.
+- now apply generic_format_FLT.
+- now apply generic_format_FLT.
+Qed.
+
+Theorem round_round_minus_radix_ge_3_FLT :
+ (3 <= beta)%Z ->
+ forall choice1 choice2,
+ (emin' <= emin)%Z -> (2 * prec <= prec')%Z ->
+ forall x y,
+ FLT_format beta emin prec x -> FLT_format beta emin prec y ->
+ round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec')
+ choice1 choice2 (x - y).
+Proof.
+intros Hbeta choice1 choice2 Hemin Hprec x y Fx Fy.
+apply round_round_minus_radix_ge_3.
+- exact Hbeta.
+- now apply FLT_exp_valid.
+- now apply FLT_exp_valid.
+- now apply FLT_round_round_plus_radix_ge_3_hyp.
+- now apply generic_format_FLT.
+- now apply generic_format_FLT.
+Qed.
+
+End Double_round_plus_radix_ge_3_FLT.
+
+Section Double_round_plus_radix_ge_3_FTZ.
+
+Variable emin prec : Z.
+Variable emin' prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FTZ_round_round_plus_radix_ge_3_hyp :
+ (emin' + prec' <= emin + 1)%Z -> (2 * prec <= prec')%Z ->
+ round_round_plus_radix_ge_3_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec').
+Proof.
+intros Hemin Hprec.
+unfold FTZ_exp.
+unfold Prec_gt_0 in *.
+unfold round_round_plus_radix_ge_3_hyp; split; [|split; [|split]]; intros ex ey.
+- destruct (Z.ltb_spec (ex + 1 - prec) emin);
+ destruct (Z.ltb_spec (ex - prec') emin');
+ destruct (Z.ltb_spec (ey - prec) emin);
+ omega.
+- destruct (Z.ltb_spec (ex - 1 - prec) emin);
+ destruct (Z.ltb_spec (ex - prec') emin');
+ destruct (Z.ltb_spec (ey - prec) emin);
+ omega.
+- destruct (Z.ltb_spec (ex - prec) emin);
+ destruct (Z.ltb_spec (ex - prec') emin');
+ destruct (Z.ltb_spec (ey - prec) emin);
+ omega.
+- destruct (Z.ltb_spec (ex - prec') emin');
+ destruct (Z.ltb_spec (ey - prec) emin);
+ omega.
+Qed.
+
+Theorem round_round_plus_radix_ge_3_FTZ :
+ (3 <= beta)%Z ->
+ forall choice1 choice2,
+ (emin' + prec' <= emin + 1)%Z -> (2 * prec <= prec')%Z ->
+ forall x y,
+ FTZ_format beta emin prec x -> FTZ_format beta emin prec y ->
+ round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec')
+ choice1 choice2 (x + y).
+Proof.
+intros Hbeta choice1 choice2 Hemin Hprec x y Fx Fy.
+apply round_round_plus_radix_ge_3.
+- exact Hbeta.
+- now apply FTZ_exp_valid.
+- now apply FTZ_exp_valid.
+- now apply FTZ_round_round_plus_radix_ge_3_hyp.
+- now apply generic_format_FTZ.
+- now apply generic_format_FTZ.
+Qed.
+
+Theorem round_round_minus_radix_ge_3_FTZ :
+ (3 <= beta)%Z ->
+ forall choice1 choice2,
+ (emin' + prec' <= emin + 1)%Z -> (2 * prec <= prec')%Z ->
+ forall x y,
+ FTZ_format beta emin prec x -> FTZ_format beta emin prec y ->
+ round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec')
+ choice1 choice2 (x - y).
+Proof.
+intros Hbeta choice1 choice2 Hemin Hprec x y Fx Fy.
+apply round_round_minus_radix_ge_3.
+- exact Hbeta.
+- now apply FTZ_exp_valid.
+- now apply FTZ_exp_valid.
+- now apply FTZ_round_round_plus_radix_ge_3_hyp.
+- now apply generic_format_FTZ.
+- now apply generic_format_FTZ.
+Qed.
+
+End Double_round_plus_radix_ge_3_FTZ.
+
+End Double_round_plus_radix_ge_3.
+
+End Double_round_plus.
+
+Lemma round_round_mid_cases :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ forall x,
+ 0 < x ->
+ (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z ->
+ (fexp1 (mag x) <= mag x)%Z ->
+ (Rabs (x - midp fexp1 x) <= / 2 * (ulp beta fexp2 x) ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x) ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x.
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 x Px Hf2f1 Hf1.
+unfold round_round_eq, midp.
+set (rd := round beta fexp1 Zfloor x).
+set (u1 := ulp beta fexp1 x).
+set (u2 := ulp beta fexp2 x).
+intros Cmid.
+destruct (generic_format_EM beta fexp1 x) as [Fx|Nfx].
+- (* generic_format beta fexp1 x *)
+ rewrite (round_generic beta fexp2); [reflexivity|now apply valid_rnd_N|].
+ now apply (generic_inclusion_mag beta fexp1); [omega|].
+- (* ~ generic_format beta fexp1 x *)
+ assert (Hceil : round beta fexp1 Zceil x = rd + u1);
+ [now apply round_UP_DN_ulp|].
+ assert (Hf2' : (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z); [omega|].
+ destruct (Rlt_or_le (x - rd) (/ 2 * (u1 - u2))).
+ + (* x - rd < / 2 * (u1 - u2) *)
+ apply round_round_lt_mid_further_place; try assumption.
+ unfold midp. fold rd; fold u1; fold u2.
+ apply (Rplus_lt_reg_r (- rd)); ring_simplify.
+ now rewrite <- Rmult_minus_distr_l.
+ + (* / 2 * (u1 - u2) <= x - rd *)
+ { destruct (Rlt_or_le (/ 2 * (u1 + u2)) (x - rd)).
+ - (* / 2 * (u1 + u2) < x - rd *)
+ assert (round beta fexp1 Zceil x - x
+ < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)).
+ { rewrite Hceil; fold u1; fold u2.
+ lra. }
+ apply round_round_gt_mid_further_place; try assumption.
+ unfold midp'; lra.
+ - (* x - rd <= / 2 * (u1 + u2) *)
+ apply Cmid, Rabs_le; split; lra. }
+Qed.
+
+Section Double_round_sqrt.
+
+Definition round_round_sqrt_hyp fexp1 fexp2 :=
+ (forall ex, (2 * fexp1 ex <= fexp1 (2 * ex))%Z)
+ /\ (forall ex, (2 * fexp1 ex <= fexp1 (2 * ex - 1))%Z)
+ /\ (forall ex, (fexp1 (2 * ex) < 2 * ex)%Z ->
+ (fexp2 ex + ex <= 2 * fexp1 ex - 2)%Z).
+
+Lemma mag_sqrt_disj :
+ forall x,
+ 0 < x ->
+ (mag x = 2 * mag (sqrt x) - 1 :> Z)%Z
+ \/ (mag x = 2 * mag (sqrt x) :> Z)%Z.
+Proof.
+intros x Px.
+rewrite (mag_sqrt beta x Px).
+generalize (Zdiv2_odd_eqn (mag x + 1)).
+destruct Z.odd ; intros ; omega.
+Qed.
+
+Lemma round_round_sqrt_aux :
+ forall fexp1 fexp2 : Z -> Z,
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ round_round_sqrt_hyp fexp1 fexp2 ->
+ forall x,
+ 0 < x ->
+ (fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z ->
+ generic_format beta fexp1 x ->
+ / 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x)).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 Hexp x Px Hf2 Fx.
+assert (Hbeta : (2 <= beta)%Z).
+{ destruct beta as (beta_val,beta_prop).
+ now apply Zle_bool_imp_le. }
+set (a := round beta fexp1 Zfloor (sqrt x)).
+set (u1 := bpow (fexp1 (mag (sqrt x)))).
+set (u2 := bpow (fexp2 (mag (sqrt x)))).
+set (b := / 2 * (u1 - u2)).
+set (b' := / 2 * (u1 + u2)).
+unfold midp; rewrite 2!ulp_neq_0; try now apply Rgt_not_eq, sqrt_lt_R0.
+apply Rnot_ge_lt; intro H; apply Rge_le in H.
+assert (Fa : generic_format beta fexp1 a).
+{ unfold a.
+ apply generic_format_round.
+ - exact Vfexp1.
+ - now apply valid_rnd_DN. }
+revert Fa; revert Fx.
+unfold generic_format, F2R, scaled_mantissa, cexp; simpl.
+set (mx := Ztrunc (x * bpow (- fexp1 (mag x)))).
+set (ma := Ztrunc (a * bpow (- fexp1 (mag a)))).
+intros Fx Fa.
+assert (Nna : 0 <= a).
+{ rewrite <- (round_0 beta fexp1 Zfloor).
+ unfold a; apply round_le.
+ - exact Vfexp1.
+ - now apply valid_rnd_DN.
+ - apply sqrt_pos. }
+assert (Phu1 : 0 < / 2 * u1).
+{ apply Rmult_lt_0_compat; [lra|apply bpow_gt_0]. }
+assert (Phu2 : 0 < / 2 * u2).
+{ apply Rmult_lt_0_compat; [lra|apply bpow_gt_0]. }
+assert (Pb : 0 < b).
+{ unfold b.
+ rewrite <- (Rmult_0_r (/ 2)).
+ apply Rmult_lt_compat_l; [lra|].
+ apply Rlt_Rminus.
+ unfold u2, u1.
+ apply bpow_lt.
+ omega. }
+assert (Pb' : 0 < b').
+{ now unfold b'; rewrite Rmult_plus_distr_l; apply Rplus_lt_0_compat. }
+assert (Hr : sqrt x <= a + b').
+{ unfold b'; apply (Rplus_le_reg_r (- / 2 * u1 - a)); ring_simplify.
+ replace (_ - _) with (sqrt x - (a + / 2 * u1)) by ring.
+ now apply Rabs_le_inv. }
+assert (Hl : a + b <= sqrt x).
+{ unfold b; apply (Rplus_le_reg_r (- / 2 * u1 - a)); ring_simplify.
+ replace (_ + sqrt _) with (sqrt x - (a + / 2 * u1)) by ring.
+ rewrite Ropp_mult_distr_l_reverse.
+ now apply Rabs_le_inv in H; destruct H. }
+assert (Hf1 : (2 * fexp1 (mag (sqrt x)) <= fexp1 (mag (x)))%Z);
+ [destruct (mag_sqrt_disj x Px) as [H'|H']; rewrite H'; apply Hexp|].
+assert (Hlx : (fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z).
+{ destruct (mag_sqrt_disj x Px) as [Hlx|Hlx].
+ - apply (valid_exp_large fexp1 (mag x)); [|omega].
+ now apply mag_generic_gt; [|apply Rgt_not_eq|].
+ - rewrite <- Hlx.
+ now apply mag_generic_gt; [|apply Rgt_not_eq|]. }
+assert (Hsl : a * a + u1 * a - u2 * a + b * b <= x).
+{ replace (_ + _) with ((a + b) * (a + b)); [|now unfold b; field].
+ rewrite <- sqrt_def; [|now apply Rlt_le].
+ assert (H' : 0 <= a + b); [now apply Rlt_le, Rplus_le_lt_0_compat|].
+ now apply Rmult_le_compat. }
+assert (Hsr : x <= a * a + u1 * a + u2 * a + b' * b').
+{ replace (_ + _) with ((a + b') * (a + b')); [|now unfold b'; field].
+ rewrite <- (sqrt_def x); [|now apply Rlt_le].
+ assert (H' : 0 <= sqrt x); [now apply sqrt_pos|].
+ now apply Rmult_le_compat. }
+destruct (Req_dec a 0) as [Za|Nza].
+- (* a = 0 *)
+ apply (Rlt_irrefl 0).
+ apply Rlt_le_trans with (b * b); [now apply Rmult_lt_0_compat|].
+ apply Rle_trans with x.
+ + revert Hsl; unfold Rminus; rewrite Za; do 3 rewrite Rmult_0_r.
+ now rewrite Ropp_0; do 3 rewrite Rplus_0_l.
+ + rewrite Fx.
+ apply (Rmult_le_reg_r (bpow (- fexp1 (mag x))));
+ [now apply bpow_gt_0|].
+ rewrite Rmult_0_l; bpow_simplify.
+ unfold mx.
+ rewrite Ztrunc_floor;
+ [|now apply Rmult_le_pos; [apply Rlt_le|apply bpow_ge_0]].
+ apply Req_le, IZR_eq.
+ apply Zfloor_imp.
+ split; [now apply Rmult_le_pos; [apply Rlt_le|apply bpow_ge_0]|simpl].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x)))); [now apply bpow_gt_0|].
+ rewrite Rmult_1_l; bpow_simplify.
+ apply Rlt_le_trans with (bpow (2 * fexp1 (mag (sqrt x))));
+ [|now apply bpow_le].
+ change 2%Z with (1 + 1)%Z; rewrite Zmult_plus_distr_l; rewrite Zmult_1_l.
+ rewrite bpow_plus.
+ rewrite <- (sqrt_def x) at 1; [|now apply Rlt_le].
+ assert (sqrt x < bpow (fexp1 (mag (sqrt x))));
+ [|now apply Rmult_lt_compat; [apply sqrt_pos|apply sqrt_pos| |]].
+ apply (Rle_lt_trans _ _ _ Hr); rewrite Za; rewrite Rplus_0_l.
+ unfold b'; change (bpow _) with u1.
+ apply Rlt_le_trans with (/ 2 * (u1 + u1)); [|lra].
+ apply Rmult_lt_compat_l; [lra|]; apply Rplus_lt_compat_l.
+ unfold u2, u1, ulp, cexp; apply bpow_lt; omega.
+- (* a <> 0 *)
+ assert (Pa : 0 < a); [lra|].
+ assert (Hla : (mag a = mag (sqrt x) :> Z)).
+ { unfold a; apply mag_DN.
+ - exact Vfexp1.
+ - now fold a. }
+ assert (Hl' : 0 < - (u2 * a) + b * b).
+ { apply (Rplus_lt_reg_r (u2 * a)); ring_simplify.
+ unfold b; ring_simplify.
+ apply (Rplus_lt_reg_r (/ 2 * u2 * u1)); field_simplify.
+ replace (_ / 2) with (u2 * (a + / 2 * u1)) by field.
+ replace (_ / 8) with (/ 4 * (u2 ^ 2 + u1 ^ 2)) by field.
+ apply Rlt_le_trans with (u2 * bpow (mag (sqrt x))).
+ - apply Rmult_lt_compat_l; [now unfold u2, ulp; apply bpow_gt_0|].
+ unfold u1; rewrite <- Hla.
+ apply Rlt_le_trans with (a + bpow (fexp1 (mag a))).
+ + apply Rplus_lt_compat_l.
+ rewrite <- (Rmult_1_l (bpow _)) at 2.
+ apply Rmult_lt_compat_r; [apply bpow_gt_0|lra].
+ + apply Rle_trans with (a+ ulp beta fexp1 a).
+ right; now rewrite ulp_neq_0.
+ apply (id_p_ulp_le_bpow _ _ _ _ Pa Fa).
+ apply Rabs_lt_inv, bpow_mag_gt.
+ - apply Rle_trans with (bpow (- 2) * u1 ^ 2).
+ + unfold pow; rewrite Rmult_1_r.
+ unfold u1, u2, ulp, cexp; bpow_simplify; apply bpow_le.
+ now apply Hexp.
+ + apply Rmult_le_compat.
+ * apply bpow_ge_0.
+ * apply pow2_ge_0.
+ * unfold Raux.bpow, Z.pow_pos; simpl; rewrite Zmult_1_r.
+ apply Rinv_le; [lra|].
+ change 4%Z with (2 * 2)%Z; apply IZR_le, Zmult_le_compat; omega.
+ * rewrite <- (Rplus_0_l (u1 ^ 2)) at 1; apply Rplus_le_compat_r.
+ apply pow2_ge_0. }
+ assert (Hr' : x <= a * a + u1 * a).
+ { rewrite Hla in Fa.
+ rewrite <- Rmult_plus_distr_r.
+ unfold u1, ulp, cexp.
+ rewrite <- (Rmult_1_l (bpow _)); rewrite Fa; rewrite <- Rmult_plus_distr_r.
+ rewrite <- Rmult_assoc; rewrite (Rmult_comm _ (IZR ma)).
+ rewrite <- (Rmult_assoc (IZR ma)); bpow_simplify.
+ apply (Rmult_le_reg_r (bpow (- 2 * fexp1 (mag (sqrt x)))));
+ [now apply bpow_gt_0|bpow_simplify].
+ rewrite Fx at 1; bpow_simplify.
+ rewrite <- IZR_Zpower; [|omega].
+ rewrite <- plus_IZR, <- 2!mult_IZR.
+ apply IZR_le, Zlt_succ_le, lt_IZR.
+ unfold Z.succ; rewrite plus_IZR; do 2 rewrite mult_IZR; rewrite plus_IZR.
+ rewrite IZR_Zpower; [|omega].
+ apply (Rmult_lt_reg_r (bpow (2 * fexp1 (mag (sqrt x)))));
+ [now apply bpow_gt_0|bpow_simplify].
+ rewrite <- Fx.
+ change 2%Z with (1 + 1)%Z; rewrite Zmult_plus_distr_l; rewrite Zmult_1_l.
+ rewrite bpow_plus; simpl.
+ replace (_ * _) with (a * a + u1 * a + u1 * u1);
+ [|unfold u1, ulp, cexp; rewrite Fa; ring].
+ apply (Rle_lt_trans _ _ _ Hsr).
+ rewrite Rplus_assoc; apply Rplus_lt_compat_l.
+ apply (Rplus_lt_reg_r (- b' * b' + / 2 * u1 * u2)); ring_simplify.
+ replace (_ + _) with ((a + / 2 * u1) * u2) by ring.
+ apply Rlt_le_trans with (bpow (mag (sqrt x)) * u2).
+ - apply Rmult_lt_compat_r; [now unfold u2, ulp; apply bpow_gt_0|].
+ apply Rlt_le_trans with (a + u1); [lra|].
+ unfold u1; fold (cexp beta fexp1 (sqrt x)).
+ rewrite <- cexp_DN; [|exact Vfexp1|exact Pa]; fold a.
+ rewrite <- ulp_neq_0; trivial.
+ apply id_p_ulp_le_bpow.
+ + exact Pa.
+ + now apply round_DN_pt.
+ + apply Rle_lt_trans with (sqrt x).
+ * now apply round_DN_pt.
+ * apply Rabs_lt_inv.
+ apply bpow_mag_gt.
+ - apply Rle_trans with (/ 2 * u1 ^ 2).
+ + apply Rle_trans with (bpow (- 2) * u1 ^ 2).
+ * unfold pow; rewrite Rmult_1_r.
+ unfold u2, u1, ulp, cexp.
+ bpow_simplify.
+ apply bpow_le.
+ rewrite Zplus_comm.
+ now apply Hexp.
+ * apply Rmult_le_compat_r; [now apply pow2_ge_0|].
+ unfold Raux.bpow; simpl; unfold Z.pow_pos; simpl.
+ rewrite Zmult_1_r.
+ apply Rinv_le; [lra|].
+ apply IZR_le.
+ rewrite <- (Zmult_1_l 2).
+ apply Zmult_le_compat; omega.
+ + assert (u2 ^ 2 < u1 ^ 2); [|unfold b'; lra].
+ unfold pow; do 2 rewrite Rmult_1_r.
+ assert (H' : 0 <= u2); [unfold u2, ulp; apply bpow_ge_0|].
+ assert (u2 < u1); [|now apply Rmult_lt_compat].
+ unfold u1, u2, ulp, cexp; apply bpow_lt; omega. }
+ apply (Rlt_irrefl (a * a + u1 * a)).
+ apply Rlt_le_trans with (a * a + u1 * a - u2 * a + b * b).
+ + rewrite <- (Rplus_0_r (a * a + _)) at 1.
+ unfold Rminus; rewrite (Rplus_assoc _ _ (b * b)).
+ now apply Rplus_lt_compat_l.
+ + now apply Rle_trans with x.
+Qed.
+
+
+Lemma round_round_sqrt :
+ forall fexp1 fexp2 : Z -> Z,
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_sqrt_hyp fexp1 fexp2 ->
+ forall x,
+ generic_format beta fexp1 x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (sqrt x).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x Fx.
+unfold round_round_eq.
+destruct (Rle_or_lt x 0) as [Npx|Px].
+- (* x <= 0 *)
+ rewrite (sqrt_neg _ Npx).
+ now rewrite round_0; [|apply valid_rnd_N].
+- (* 0 < x *)
+ assert (Hfx : (fexp1 (mag x) < mag x)%Z);
+ [now apply mag_generic_gt; try assumption; lra|].
+ assert (Hfsx : (fexp1 (mag (sqrt x)) < mag (sqrt x))%Z).
+ { destruct (Rle_or_lt x 1) as [Hx|Hx].
+ - (* x <= 1 *)
+ apply (valid_exp_large fexp1 (mag x)); [exact Hfx|].
+ apply mag_le; [exact Px|].
+ rewrite <- (sqrt_def x) at 1; [|lra].
+ rewrite <- Rmult_1_r.
+ apply Rmult_le_compat_l.
+ + apply sqrt_pos.
+ + rewrite <- sqrt_1.
+ now apply sqrt_le_1_alt.
+ - (* 1 < x *)
+ generalize ((proj1 (proj2 Hexp)) 1%Z).
+ replace (_ - 1)%Z with 1%Z by ring.
+ intro Hexp10.
+ assert (Hf0 : (fexp1 1 < 1)%Z); [omega|clear Hexp10].
+ apply (valid_exp_large fexp1 1); [exact Hf0|].
+ apply mag_ge_bpow.
+ rewrite Zeq_minus; [|reflexivity].
+ unfold Raux.bpow; simpl.
+ apply Rabs_ge; right.
+ rewrite <- sqrt_1.
+ apply sqrt_le_1_alt.
+ now apply Rlt_le. }
+ assert (Hf2 : (fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z).
+ { assert (H : (fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z).
+ { destruct (mag_sqrt_disj x Px) as [Hlx|Hlx].
+ - apply (valid_exp_large fexp1 (mag x)); [|omega].
+ now apply mag_generic_gt; [|apply Rgt_not_eq|].
+ - rewrite <- Hlx.
+ now apply mag_generic_gt; [|apply Rgt_not_eq|]. }
+ generalize ((proj2 (proj2 Hexp)) (mag (sqrt x)) H).
+ omega. }
+ apply round_round_mid_cases.
+ + exact Vfexp1.
+ + exact Vfexp2.
+ + now apply sqrt_lt_R0.
+ + omega.
+ + omega.
+ + intros Hmid; casetype False; apply (Rle_not_lt _ _ Hmid).
+ apply (round_round_sqrt_aux fexp1 fexp2 Vfexp1 Vfexp2 Hexp x Px Hf2 Fx).
+Qed.
+
+Section Double_round_sqrt_FLX.
+
+Variable prec : Z.
+Variable prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FLX_round_round_sqrt_hyp :
+ (2 * prec + 2 <= prec')%Z ->
+ round_round_sqrt_hyp (FLX_exp prec) (FLX_exp prec').
+Proof.
+intros Hprec.
+unfold FLX_exp.
+unfold Prec_gt_0 in prec_gt_0_.
+unfold round_round_sqrt_hyp; split; [|split]; intro ex; omega.
+Qed.
+
+Theorem round_round_sqrt_FLX :
+ forall choice1 choice2,
+ (2 * prec + 2 <= prec')%Z ->
+ forall x,
+ FLX_format beta prec x ->
+ round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (sqrt x).
+Proof.
+intros choice1 choice2 Hprec x Fx.
+apply round_round_sqrt.
+- now apply FLX_exp_valid.
+- now apply FLX_exp_valid.
+- now apply FLX_round_round_sqrt_hyp.
+- now apply generic_format_FLX.
+Qed.
+
+End Double_round_sqrt_FLX.
+
+Section Double_round_sqrt_FLT.
+
+Variable emin prec : Z.
+Variable emin' prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FLT_round_round_sqrt_hyp :
+ (emin <= 0)%Z ->
+ ((emin' <= emin - prec - 2)%Z
+ \/ (2 * emin' <= emin - 4 * prec - 2)%Z) ->
+ (2 * prec + 2 <= prec')%Z ->
+ round_round_sqrt_hyp (FLT_exp emin prec) (FLT_exp emin' prec').
+Proof.
+intros Hemin Heminprec Hprec.
+unfold FLT_exp.
+unfold Prec_gt_0 in prec_gt_0_.
+unfold round_round_sqrt_hyp; split; [|split]; intros ex.
+- generalize (Zmax_spec (ex - prec) emin).
+ generalize (Zmax_spec (2 * ex - prec) emin).
+ omega.
+- generalize (Zmax_spec (ex - prec) emin).
+ generalize (Zmax_spec (2 * ex - 1 - prec) emin).
+ omega.
+- generalize (Zmax_spec (2 * ex - prec) emin).
+ generalize (Zmax_spec (ex - prec') emin').
+ generalize (Zmax_spec (ex - prec) emin).
+ omega.
+Qed.
+
+Theorem round_round_sqrt_FLT :
+ forall choice1 choice2,
+ (emin <= 0)%Z ->
+ ((emin' <= emin - prec - 2)%Z
+ \/ (2 * emin' <= emin - 4 * prec - 2)%Z) ->
+ (2 * prec + 2 <= prec')%Z ->
+ forall x,
+ FLT_format beta emin prec x ->
+ round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec')
+ choice1 choice2 (sqrt x).
+Proof.
+intros choice1 choice2 Hemin Heminprec Hprec x Fx.
+apply round_round_sqrt.
+- now apply FLT_exp_valid.
+- now apply FLT_exp_valid.
+- now apply FLT_round_round_sqrt_hyp.
+- now apply generic_format_FLT.
+Qed.
+
+End Double_round_sqrt_FLT.
+
+Section Double_round_sqrt_FTZ.
+
+Variable emin prec : Z.
+Variable emin' prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FTZ_round_round_sqrt_hyp :
+ (2 * (emin' + prec') <= emin + prec <= 1)%Z ->
+ (2 * prec + 2 <= prec')%Z ->
+ round_round_sqrt_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec').
+Proof.
+intros Hemin Hprec.
+unfold FTZ_exp.
+unfold Prec_gt_0 in *.
+unfold round_round_sqrt_hyp; split; [|split]; intros ex.
+- destruct (Z.ltb_spec (ex - prec) emin);
+ destruct (Z.ltb_spec (2 * ex - prec) emin);
+ omega.
+- destruct (Z.ltb_spec (ex - prec) emin);
+ destruct (Z.ltb_spec (2 * ex - 1 - prec) emin);
+ omega.
+- intro H.
+ destruct (Zle_or_lt emin (2 * ex - prec)) as [H'|H'].
+ + destruct (Z.ltb_spec (ex - prec') emin');
+ destruct (Z.ltb_spec (ex - prec) emin);
+ omega.
+ + casetype False.
+ rewrite (Zlt_bool_true _ _ H') in H.
+ omega.
+Qed.
+
+Theorem round_round_sqrt_FTZ :
+ (4 <= beta)%Z ->
+ forall choice1 choice2,
+ (2 * (emin' + prec') <= emin + prec <= 1)%Z ->
+ (2 * prec + 2 <= prec')%Z ->
+ forall x,
+ FTZ_format beta emin prec x ->
+ round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec')
+ choice1 choice2 (sqrt x).
+Proof.
+intros Hbeta choice1 choice2 Hemin Hprec x Fx.
+apply round_round_sqrt.
+- now apply FTZ_exp_valid.
+- now apply FTZ_exp_valid.
+- now apply FTZ_round_round_sqrt_hyp.
+- now apply generic_format_FTZ.
+Qed.
+
+End Double_round_sqrt_FTZ.
+
+Section Double_round_sqrt_radix_ge_4.
+
+Definition round_round_sqrt_radix_ge_4_hyp fexp1 fexp2 :=
+ (forall ex, (2 * fexp1 ex <= fexp1 (2 * ex))%Z)
+ /\ (forall ex, (2 * fexp1 ex <= fexp1 (2 * ex - 1))%Z)
+ /\ (forall ex, (fexp1 (2 * ex) < 2 * ex)%Z ->
+ (fexp2 ex + ex <= 2 * fexp1 ex - 1)%Z).
+
+Lemma round_round_sqrt_radix_ge_4_aux :
+ (4 <= beta)%Z ->
+ forall fexp1 fexp2 : Z -> Z,
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ round_round_sqrt_radix_ge_4_hyp fexp1 fexp2 ->
+ forall x,
+ 0 < x ->
+ (fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z ->
+ generic_format beta fexp1 x ->
+ / 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x)).
+Proof.
+intros Hbeta fexp1 fexp2 Vfexp1 Vfexp2 Hexp x Px Hf2 Fx.
+set (a := round beta fexp1 Zfloor (sqrt x)).
+set (u1 := bpow (fexp1 (mag (sqrt x)))).
+set (u2 := bpow (fexp2 (mag (sqrt x)))).
+set (b := / 2 * (u1 - u2)).
+set (b' := / 2 * (u1 + u2)).
+unfold midp; rewrite 2!ulp_neq_0; try now apply Rgt_not_eq, sqrt_lt_R0.
+apply Rnot_ge_lt; intro H; apply Rge_le in H.
+assert (Fa : generic_format beta fexp1 a).
+{ unfold a.
+ apply generic_format_round.
+ - exact Vfexp1.
+ - now apply valid_rnd_DN. }
+revert Fa; revert Fx.
+unfold generic_format, F2R, scaled_mantissa, cexp; simpl.
+set (mx := Ztrunc (x * bpow (- fexp1 (mag x)))).
+set (ma := Ztrunc (a * bpow (- fexp1 (mag a)))).
+intros Fx Fa.
+assert (Nna : 0 <= a).
+{ rewrite <- (round_0 beta fexp1 Zfloor).
+ unfold a; apply round_le.
+ - exact Vfexp1.
+ - now apply valid_rnd_DN.
+ - apply sqrt_pos. }
+assert (Phu1 : 0 < / 2 * u1).
+{ apply Rmult_lt_0_compat; [lra|apply bpow_gt_0]. }
+assert (Phu2 : 0 < / 2 * u2).
+{ apply Rmult_lt_0_compat; [lra|apply bpow_gt_0]. }
+assert (Pb : 0 < b).
+{ unfold b.
+ rewrite <- (Rmult_0_r (/ 2)).
+ apply Rmult_lt_compat_l; [lra|].
+ apply Rlt_Rminus.
+ unfold u2, u1, ulp, cexp.
+ apply bpow_lt.
+ omega. }
+assert (Pb' : 0 < b').
+{ now unfold b'; rewrite Rmult_plus_distr_l; apply Rplus_lt_0_compat. }
+assert (Hr : sqrt x <= a + b').
+{ unfold b'; apply (Rplus_le_reg_r (- / 2 * u1 - a)); ring_simplify.
+ replace (_ - _) with (sqrt x - (a + / 2 * u1)) by ring.
+ now apply Rabs_le_inv. }
+assert (Hl : a + b <= sqrt x).
+{ unfold b; apply (Rplus_le_reg_r (- / 2 * u1 - a)); ring_simplify.
+ replace (_ + sqrt _) with (sqrt x - (a + / 2 * u1)) by ring.
+ rewrite Ropp_mult_distr_l_reverse.
+ now apply Rabs_le_inv in H; destruct H. }
+assert (Hf1 : (2 * fexp1 (mag (sqrt x)) <= fexp1 (mag (x)))%Z);
+ [destruct (mag_sqrt_disj x Px) as [H'|H']; rewrite H'; apply Hexp|].
+assert (Hlx : (fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z).
+{ destruct (mag_sqrt_disj x Px) as [Hlx|Hlx].
+ - apply (valid_exp_large fexp1 (mag x)); [|omega].
+ now apply mag_generic_gt; [|apply Rgt_not_eq|].
+ - rewrite <- Hlx.
+ now apply mag_generic_gt; [|apply Rgt_not_eq|]. }
+assert (Hsl : a * a + u1 * a - u2 * a + b * b <= x).
+{ replace (_ + _) with ((a + b) * (a + b)); [|now unfold b; field].
+ rewrite <- sqrt_def; [|now apply Rlt_le].
+ assert (H' : 0 <= a + b); [now apply Rlt_le, Rplus_le_lt_0_compat|].
+ now apply Rmult_le_compat. }
+assert (Hsr : x <= a * a + u1 * a + u2 * a + b' * b').
+{ replace (_ + _) with ((a + b') * (a + b')); [|now unfold b'; field].
+ rewrite <- (sqrt_def x); [|now apply Rlt_le].
+ assert (H' : 0 <= sqrt x); [now apply sqrt_pos|].
+ now apply Rmult_le_compat. }
+destruct (Req_dec a 0) as [Za|Nza].
+- (* a = 0 *)
+ apply (Rlt_irrefl 0).
+ apply Rlt_le_trans with (b * b); [now apply Rmult_lt_0_compat|].
+ apply Rle_trans with x.
+ + revert Hsl; unfold Rminus; rewrite Za; do 3 rewrite Rmult_0_r.
+ now rewrite Ropp_0; do 3 rewrite Rplus_0_l.
+ + rewrite Fx.
+ apply (Rmult_le_reg_r (bpow (- fexp1 (mag x))));
+ [now apply bpow_gt_0|].
+ rewrite Rmult_0_l; bpow_simplify.
+ unfold mx.
+ rewrite Ztrunc_floor;
+ [|now apply Rmult_le_pos; [apply Rlt_le|apply bpow_ge_0]].
+ apply Req_le, IZR_eq.
+ apply Zfloor_imp.
+ split; [now apply Rmult_le_pos; [apply Rlt_le|apply bpow_ge_0]|simpl].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x)))); [now apply bpow_gt_0|].
+ rewrite Rmult_1_l; bpow_simplify.
+ apply Rlt_le_trans with (bpow (2 * fexp1 (mag (sqrt x))));
+ [|now apply bpow_le].
+ change 2%Z with (1 + 1)%Z; rewrite Zmult_plus_distr_l; rewrite Zmult_1_l.
+ rewrite bpow_plus.
+ rewrite <- (sqrt_def x) at 1; [|now apply Rlt_le].
+ assert (sqrt x < bpow (fexp1 (mag (sqrt x))));
+ [|now apply Rmult_lt_compat; [apply sqrt_pos|apply sqrt_pos| |]].
+ apply (Rle_lt_trans _ _ _ Hr); rewrite Za; rewrite Rplus_0_l.
+ unfold b'; change (bpow _) with u1.
+ apply Rlt_le_trans with (/ 2 * (u1 + u1)); [|lra].
+ apply Rmult_lt_compat_l; [lra|]; apply Rplus_lt_compat_l.
+ unfold u2, u1, ulp, cexp; apply bpow_lt; omega.
+- (* a <> 0 *)
+ assert (Pa : 0 < a); [lra|].
+ assert (Hla : (mag a = mag (sqrt x) :> Z)).
+ { unfold a; apply mag_DN.
+ - exact Vfexp1.
+ - now fold a. }
+ assert (Hl' : 0 < - (u2 * a) + b * b).
+ { apply (Rplus_lt_reg_r (u2 * a)); ring_simplify.
+ unfold b; ring_simplify.
+ apply (Rplus_lt_reg_r (/ 2 * u2 * u1)); field_simplify.
+ replace (_ / 2) with (u2 * (a + / 2 * u1)) by field.
+ replace (_ / 8) with (/ 4 * (u2 ^ 2 + u1 ^ 2)) by field.
+ apply Rlt_le_trans with (u2 * bpow (mag (sqrt x))).
+ - apply Rmult_lt_compat_l; [now unfold u2, ulp; apply bpow_gt_0|].
+ unfold u1; rewrite <- Hla.
+ apply Rlt_le_trans with (a + ulp beta fexp1 a).
+ + apply Rplus_lt_compat_l.
+ rewrite <- (Rmult_1_l (ulp _ _ _)).
+ rewrite ulp_neq_0; trivial.
+ apply Rmult_lt_compat_r; [apply bpow_gt_0|lra].
+ + apply (id_p_ulp_le_bpow _ _ _ _ Pa Fa).
+ apply Rabs_lt_inv, bpow_mag_gt.
+ - apply Rle_trans with (bpow (- 1) * u1 ^ 2).
+ + unfold pow; rewrite Rmult_1_r.
+ unfold u1, u2, ulp, cexp; bpow_simplify; apply bpow_le.
+ now apply Hexp.
+ + apply Rmult_le_compat.
+ * apply bpow_ge_0.
+ * apply pow2_ge_0.
+ * unfold Raux.bpow, Z.pow_pos; simpl; rewrite Zmult_1_r.
+ apply Rinv_le; [lra|].
+ now apply IZR_le.
+ * rewrite <- (Rplus_0_l (u1 ^ 2)) at 1; apply Rplus_le_compat_r.
+ apply pow2_ge_0. }
+ assert (Hr' : x <= a * a + u1 * a).
+ { rewrite Hla in Fa.
+ rewrite <- Rmult_plus_distr_r.
+ unfold u1, ulp, cexp.
+ rewrite <- (Rmult_1_l (bpow _)); rewrite Fa; rewrite <- Rmult_plus_distr_r.
+ rewrite <- Rmult_assoc; rewrite (Rmult_comm _ (IZR ma)).
+ rewrite <- (Rmult_assoc (IZR ma)); bpow_simplify.
+ apply (Rmult_le_reg_r (bpow (- 2 * fexp1 (mag (sqrt x)))));
+ [now apply bpow_gt_0|bpow_simplify].
+ rewrite Fx at 1; bpow_simplify.
+ rewrite <- IZR_Zpower; [|omega].
+ rewrite <- plus_IZR, <- 2!mult_IZR.
+ apply IZR_le, Zlt_succ_le, lt_IZR.
+ unfold Z.succ; rewrite plus_IZR; do 2 rewrite mult_IZR; rewrite plus_IZR.
+ rewrite IZR_Zpower; [|omega].
+ apply (Rmult_lt_reg_r (bpow (2 * fexp1 (mag (sqrt x)))));
+ [now apply bpow_gt_0|bpow_simplify].
+ rewrite <- Fx.
+ change 2%Z with (1 + 1)%Z; rewrite Zmult_plus_distr_l; rewrite Zmult_1_l.
+ rewrite bpow_plus; simpl.
+ replace (_ * _) with (a * a + u1 * a + u1 * u1);
+ [|unfold u1, ulp, cexp; rewrite Fa; ring].
+ apply (Rle_lt_trans _ _ _ Hsr).
+ rewrite Rplus_assoc; apply Rplus_lt_compat_l.
+ apply (Rplus_lt_reg_r (- b' * b' + / 2 * u1 * u2)); ring_simplify.
+ replace (_ + _) with ((a + / 2 * u1) * u2) by ring.
+ apply Rlt_le_trans with (bpow (mag (sqrt x)) * u2).
+ - apply Rmult_lt_compat_r; [now unfold u2, ulp; apply bpow_gt_0|].
+ apply Rlt_le_trans with (a + u1); [lra|].
+ unfold u1; fold (cexp beta fexp1 (sqrt x)).
+ rewrite <- cexp_DN; [|exact Vfexp1|exact Pa]; fold a.
+ rewrite <- ulp_neq_0; trivial.
+ apply id_p_ulp_le_bpow.
+ + exact Pa.
+ + now apply round_DN_pt.
+ + apply Rle_lt_trans with (sqrt x).
+ * now apply round_DN_pt.
+ * apply Rabs_lt_inv.
+ apply bpow_mag_gt.
+ - apply Rle_trans with (/ 2 * u1 ^ 2).
+ + apply Rle_trans with (bpow (- 1) * u1 ^ 2).
+ * unfold pow; rewrite Rmult_1_r.
+ unfold u2, u1, ulp, cexp.
+ bpow_simplify.
+ apply bpow_le.
+ rewrite Zplus_comm.
+ now apply Hexp.
+ * apply Rmult_le_compat_r; [now apply pow2_ge_0|].
+ unfold Raux.bpow; simpl; unfold Z.pow_pos; simpl.
+ rewrite Zmult_1_r.
+ apply Rinv_le; [lra|].
+ apply IZR_le; omega.
+ + assert (u2 ^ 2 < u1 ^ 2); [|unfold b'; lra].
+ unfold pow; do 2 rewrite Rmult_1_r.
+ assert (H' : 0 <= u2); [unfold u2, ulp; apply bpow_ge_0|].
+ assert (u2 < u1); [|now apply Rmult_lt_compat].
+ unfold u1, u2, ulp, cexp; apply bpow_lt; omega. }
+ apply (Rlt_irrefl (a * a + u1 * a)).
+ apply Rlt_le_trans with (a * a + u1 * a - u2 * a + b * b).
+ + rewrite <- (Rplus_0_r (a * a + _)) at 1.
+ unfold Rminus; rewrite (Rplus_assoc _ _ (b * b)).
+ now apply Rplus_lt_compat_l.
+ + now apply Rle_trans with x.
+Qed.
+
+Lemma round_round_sqrt_radix_ge_4 :
+ (4 <= beta)%Z ->
+ forall fexp1 fexp2 : Z -> Z,
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_sqrt_radix_ge_4_hyp fexp1 fexp2 ->
+ forall x,
+ generic_format beta fexp1 x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (sqrt x).
+Proof.
+intros Hbeta fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x Fx.
+unfold round_round_eq.
+destruct (Rle_or_lt x 0) as [Npx|Px].
+- (* x <= 0 *)
+ assert (Hs : sqrt x = 0).
+ { destruct (Req_dec x 0) as [Zx|Nzx].
+ - (* x = 0 *)
+ rewrite Zx.
+ exact sqrt_0.
+ - (* x < 0 *)
+ unfold sqrt.
+ destruct Rcase_abs.
+ + reflexivity.
+ + casetype False; lra. }
+ rewrite Hs.
+ rewrite round_0.
+ + reflexivity.
+ + now apply valid_rnd_N.
+- (* 0 < x *)
+ assert (Hfx : (fexp1 (mag x) < mag x)%Z);
+ [now apply mag_generic_gt; try assumption; lra|].
+ assert (Hfsx : (fexp1 (mag (sqrt x)) < mag (sqrt x))%Z).
+ { destruct (Rle_or_lt x 1) as [Hx|Hx].
+ - (* x <= 1 *)
+ apply (valid_exp_large fexp1 (mag x)); [exact Hfx|].
+ apply mag_le; [exact Px|].
+ rewrite <- (sqrt_def x) at 1; [|lra].
+ rewrite <- Rmult_1_r.
+ apply Rmult_le_compat_l.
+ + apply sqrt_pos.
+ + rewrite <- sqrt_1.
+ now apply sqrt_le_1_alt.
+ - (* 1 < x *)
+ generalize ((proj1 (proj2 Hexp)) 1%Z).
+ replace (_ - 1)%Z with 1%Z by ring.
+ intro Hexp10.
+ assert (Hf0 : (fexp1 1 < 1)%Z); [omega|clear Hexp10].
+ apply (valid_exp_large fexp1 1); [exact Hf0|].
+ apply mag_ge_bpow.
+ rewrite Zeq_minus; [|reflexivity].
+ unfold Raux.bpow; simpl.
+ apply Rabs_ge; right.
+ rewrite <- sqrt_1.
+ apply sqrt_le_1_alt.
+ now apply Rlt_le. }
+ assert (Hf2 : (fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z).
+ { assert (H : (fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z).
+ { destruct (mag_sqrt_disj x Px) as [Hlx|Hlx].
+ - apply (valid_exp_large fexp1 (mag x)); [|omega].
+ now apply mag_generic_gt; [|apply Rgt_not_eq|].
+ - rewrite <- Hlx.
+ now apply mag_generic_gt; [|apply Rgt_not_eq|]. }
+ generalize ((proj2 (proj2 Hexp)) (mag (sqrt x)) H).
+ omega. }
+ apply round_round_mid_cases.
+ + exact Vfexp1.
+ + exact Vfexp2.
+ + now apply sqrt_lt_R0.
+ + omega.
+ + omega.
+ + intros Hmid; casetype False; apply (Rle_not_lt _ _ Hmid).
+ apply (round_round_sqrt_radix_ge_4_aux Hbeta fexp1 fexp2 Vfexp1 Vfexp2
+ Hexp x Px Hf2 Fx).
+Qed.
+
+Section Double_round_sqrt_radix_ge_4_FLX.
+
+Variable prec : Z.
+Variable prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FLX_round_round_sqrt_radix_ge_4_hyp :
+ (2 * prec + 1 <= prec')%Z ->
+ round_round_sqrt_radix_ge_4_hyp (FLX_exp prec) (FLX_exp prec').
+Proof.
+intros Hprec.
+unfold FLX_exp.
+unfold Prec_gt_0 in prec_gt_0_.
+unfold round_round_sqrt_radix_ge_4_hyp; split; [|split]; intro ex; omega.
+Qed.
+
+Theorem round_round_sqrt_radix_ge_4_FLX :
+ (4 <= beta)%Z ->
+ forall choice1 choice2,
+ (2 * prec + 1 <= prec')%Z ->
+ forall x,
+ FLX_format beta prec x ->
+ round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (sqrt x).
+Proof.
+intros Hbeta choice1 choice2 Hprec x Fx.
+apply round_round_sqrt_radix_ge_4.
+- exact Hbeta.
+- now apply FLX_exp_valid.
+- now apply FLX_exp_valid.
+- now apply FLX_round_round_sqrt_radix_ge_4_hyp.
+- now apply generic_format_FLX.
+Qed.
+
+End Double_round_sqrt_radix_ge_4_FLX.
+
+Section Double_round_sqrt_radix_ge_4_FLT.
+
+Variable emin prec : Z.
+Variable emin' prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FLT_round_round_sqrt_radix_ge_4_hyp :
+ (emin <= 0)%Z ->
+ ((emin' <= emin - prec - 1)%Z
+ \/ (2 * emin' <= emin - 4 * prec)%Z) ->
+ (2 * prec + 1 <= prec')%Z ->
+ round_round_sqrt_radix_ge_4_hyp (FLT_exp emin prec) (FLT_exp emin' prec').
+Proof.
+intros Hemin Heminprec Hprec.
+unfold FLT_exp.
+unfold Prec_gt_0 in prec_gt_0_.
+unfold round_round_sqrt_radix_ge_4_hyp; split; [|split]; intros ex.
+- generalize (Zmax_spec (ex - prec) emin).
+ generalize (Zmax_spec (2 * ex - prec) emin).
+ omega.
+- generalize (Zmax_spec (ex - prec) emin).
+ generalize (Zmax_spec (2 * ex - 1 - prec) emin).
+ omega.
+- generalize (Zmax_spec (2 * ex - prec) emin).
+ generalize (Zmax_spec (ex - prec') emin').
+ generalize (Zmax_spec (ex - prec) emin).
+ omega.
+Qed.
+
+Theorem round_round_sqrt_radix_ge_4_FLT :
+ (4 <= beta)%Z ->
+ forall choice1 choice2,
+ (emin <= 0)%Z ->
+ ((emin' <= emin - prec - 1)%Z
+ \/ (2 * emin' <= emin - 4 * prec)%Z) ->
+ (2 * prec + 1 <= prec')%Z ->
+ forall x,
+ FLT_format beta emin prec x ->
+ round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec')
+ choice1 choice2 (sqrt x).
+Proof.
+intros Hbeta choice1 choice2 Hemin Heminprec Hprec x Fx.
+apply round_round_sqrt_radix_ge_4.
+- exact Hbeta.
+- now apply FLT_exp_valid.
+- now apply FLT_exp_valid.
+- now apply FLT_round_round_sqrt_radix_ge_4_hyp.
+- now apply generic_format_FLT.
+Qed.
+
+End Double_round_sqrt_radix_ge_4_FLT.
+
+Section Double_round_sqrt_radix_ge_4_FTZ.
+
+Variable emin prec : Z.
+Variable emin' prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FTZ_round_round_sqrt_radix_ge_4_hyp :
+ (2 * (emin' + prec') <= emin + prec <= 1)%Z ->
+ (2 * prec + 1 <= prec')%Z ->
+ round_round_sqrt_radix_ge_4_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec').
+Proof.
+intros Hemin Hprec.
+unfold FTZ_exp.
+unfold Prec_gt_0 in *.
+unfold round_round_sqrt_radix_ge_4_hyp; split; [|split]; intros ex.
+- destruct (Z.ltb_spec (ex - prec) emin);
+ destruct (Z.ltb_spec (2 * ex - prec) emin);
+ omega.
+- destruct (Z.ltb_spec (ex - prec) emin);
+ destruct (Z.ltb_spec (2 * ex - 1 - prec) emin);
+ omega.
+- intro H.
+ destruct (Zle_or_lt emin (2 * ex - prec)) as [H'|H'].
+ + destruct (Z.ltb_spec (ex - prec') emin');
+ destruct (Z.ltb_spec (ex - prec) emin);
+ omega.
+ + casetype False.
+ rewrite (Zlt_bool_true _ _ H') in H.
+ omega.
+Qed.
+
+Theorem round_round_sqrt_radix_ge_4_FTZ :
+ (4 <= beta)%Z ->
+ forall choice1 choice2,
+ (2 * (emin' + prec') <= emin + prec <= 1)%Z ->
+ (2 * prec + 1 <= prec')%Z ->
+ forall x,
+ FTZ_format beta emin prec x ->
+ round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec')
+ choice1 choice2 (sqrt x).
+Proof.
+intros Hbeta choice1 choice2 Hemin Hprec x Fx.
+apply round_round_sqrt_radix_ge_4.
+- exact Hbeta.
+- now apply FTZ_exp_valid.
+- now apply FTZ_exp_valid.
+- now apply FTZ_round_round_sqrt_radix_ge_4_hyp.
+- now apply generic_format_FTZ.
+Qed.
+
+End Double_round_sqrt_radix_ge_4_FTZ.
+
+End Double_round_sqrt_radix_ge_4.
+
+End Double_round_sqrt.
+
+Section Double_round_div.
+
+Lemma round_round_eq_mid_beta_even :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ (exists n, (beta = 2 * n :> Z)%Z) ->
+ forall x,
+ 0 < x ->
+ (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z ->
+ (fexp1 (mag x) <= mag x)%Z ->
+ x = midp fexp1 x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x.
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Ebeta x Px Hf2 Hf1.
+unfold round_round_eq.
+unfold midp.
+set (rd := round beta fexp1 Zfloor x).
+set (u := ulp beta fexp1 x).
+intro H; apply (Rplus_eq_compat_l (- rd)) in H.
+ring_simplify in H; revert H.
+rewrite Rplus_comm; fold (Rminus x rd).
+intro Xmid.
+destruct Ebeta as (n,Ebeta).
+assert (Hbeta : (2 <= beta)%Z).
+{ destruct beta as (beta_val,beta_prop).
+ now apply Zle_bool_imp_le. }
+apply (Rplus_eq_compat_l rd) in Xmid; ring_simplify in Xmid.
+rewrite (round_generic beta fexp2); [reflexivity|now apply valid_rnd_N|].
+set (f := Float beta (Zfloor (scaled_mantissa beta fexp2 rd)
+ + n * beta ^ (fexp1 (mag x) - 1
+ - fexp2 (mag x)))
+ (cexp beta fexp2 x)).
+assert (Hf : F2R f = x).
+{ unfold f, F2R; simpl.
+ rewrite plus_IZR.
+ rewrite Rmult_plus_distr_r.
+ rewrite mult_IZR.
+ rewrite IZR_Zpower; [|omega].
+ unfold cexp at 2; bpow_simplify.
+ unfold Zminus; rewrite bpow_plus.
+ rewrite (Rmult_comm _ (bpow (- 1))).
+ rewrite <- (Rmult_assoc (IZR n)).
+ change (bpow (- 1)) with (/ IZR (beta * 1)).
+ rewrite Zmult_1_r.
+ rewrite Ebeta.
+ rewrite (mult_IZR 2).
+ rewrite Rinv_mult_distr;
+ [|simpl; lra | apply IZR_neq; omega].
+ rewrite <- Rmult_assoc; rewrite (Rmult_comm (IZR n));
+ rewrite (Rmult_assoc _ (IZR n)).
+ rewrite Rinv_r;
+ [rewrite Rmult_1_r | apply IZR_neq; omega].
+ simpl; fold (cexp beta fexp1 x).
+ rewrite <- 2!ulp_neq_0; try now apply Rgt_not_eq.
+ fold u; rewrite Xmid at 2.
+ apply f_equal2; [|reflexivity].
+ rewrite ulp_neq_0; try now apply Rgt_not_eq.
+ destruct (Req_dec rd 0) as [Zrd|Nzrd].
+ - (* rd = 0 *)
+ rewrite Zrd.
+ rewrite scaled_mantissa_0.
+ rewrite Zfloor_IZR.
+ now rewrite Rmult_0_l.
+ - (* rd <> 0 *)
+ assert (Nnrd : 0 <= rd).
+ { apply round_DN_pt.
+ - exact Vfexp1.
+ - apply generic_format_0.
+ - now apply Rlt_le. }
+ assert (Prd : 0 < rd); [lra|].
+ assert (Lrd : (mag rd = mag x :> Z)).
+ { apply Zle_antisym.
+ - apply mag_le; [exact Prd|].
+ now apply round_DN_pt.
+ - apply mag_round_ge.
+ + exact Vfexp1.
+ + now apply valid_rnd_DN.
+ + exact Nzrd. }
+ unfold scaled_mantissa.
+ unfold rd at 1.
+ unfold round, F2R, scaled_mantissa, cexp; simpl.
+ bpow_simplify.
+ rewrite Lrd.
+ rewrite <- (IZR_Zpower _ (_ - _)); [|omega].
+ rewrite <- mult_IZR.
+ rewrite (Zfloor_imp (Zfloor (x * bpow (- fexp1 (mag x))) *
+ beta ^ (fexp1 (mag x) - fexp2 (mag x)))).
+ + rewrite mult_IZR.
+ rewrite IZR_Zpower; [|omega].
+ bpow_simplify.
+ now unfold rd.
+ + split; [now apply Rle_refl|].
+ rewrite plus_IZR.
+ simpl; lra. }
+apply (generic_format_F2R' _ _ x f Hf).
+intros _.
+apply Z.le_refl.
+Qed.
+
+Lemma round_round_really_zero :
+ forall (fexp1 fexp2 : Z -> Z),
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ forall x,
+ 0 < x ->
+ (mag x <= fexp1 (mag x) - 2)%Z ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x.
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 x Px Hf1.
+assert (Hlx : bpow (mag x - 1) <= x < bpow (mag x)).
+{ destruct (mag x) as (ex,Hex); simpl.
+ rewrite <- (Rabs_right x); [|now apply Rle_ge; apply Rlt_le].
+ apply Hex.
+ now apply Rgt_not_eq. }
+unfold round_round_eq.
+rewrite (round_N_small_pos beta fexp1 _ x (mag x)); [|exact Hlx|omega].
+set (x'' := round beta fexp2 (Znearest choice2) x).
+destruct (Req_dec x'' 0) as [Zx''|Nzx''];
+ [now rewrite Zx''; rewrite round_0; [|apply valid_rnd_N]|].
+destruct (Zle_or_lt (fexp2 (mag x)) (mag x)).
+- (* fexp2 (mag x) <= mag x *)
+ destruct (Rlt_or_le x'' (bpow (mag x))).
+ + (* x'' < bpow (mag x) *)
+ rewrite (round_N_small_pos beta fexp1 _ _ (mag x));
+ [reflexivity|split; [|exact H0]|omega].
+ apply round_large_pos_ge_bpow; [now apply valid_rnd_N| |now apply Hlx].
+ fold x''; assert (0 <= x''); [|lra]; unfold x''.
+ rewrite <- (round_0 beta fexp2 (Znearest choice2)).
+ now apply round_le; [|apply valid_rnd_N|apply Rlt_le].
+ + (* bpow (mag x) <= x'' *)
+ assert (Hx'' : x'' = bpow (mag x)).
+ { apply Rle_antisym; [|exact H0].
+ rewrite <- (round_generic beta fexp2 (Znearest choice2) (bpow _)).
+ - now apply round_le; [|apply valid_rnd_N|apply Rlt_le].
+ - now apply generic_format_bpow'. }
+ rewrite Hx''.
+ unfold round, F2R, scaled_mantissa, cexp; simpl.
+ rewrite mag_bpow.
+ assert (Hf11 : (fexp1 (mag x + 1) = fexp1 (mag x) :> Z)%Z);
+ [apply Vfexp1; omega|].
+ rewrite Hf11.
+ apply (Rmult_eq_reg_r (bpow (- fexp1 (mag x))));
+ [|now apply Rgt_not_eq; apply bpow_gt_0].
+ rewrite Rmult_0_l; bpow_simplify.
+ apply IZR_eq.
+ apply Znearest_imp.
+ simpl; unfold Rminus; rewrite Ropp_0; rewrite Rplus_0_r.
+ rewrite Rabs_right; [|now apply Rle_ge; apply bpow_ge_0].
+ apply Rle_lt_trans with (bpow (- 2)); [now apply bpow_le; omega|].
+ unfold Raux.bpow, Z.pow_pos; simpl; rewrite Zmult_1_r.
+ assert (Hbeta : (2 <= beta)%Z).
+ { destruct beta as (beta_val,beta_prop); simpl.
+ now apply Zle_bool_imp_le. }
+ apply Rinv_lt_contravar.
+ * apply Rmult_lt_0_compat; [lra|].
+ rewrite mult_IZR; apply Rmult_lt_0_compat;
+ apply IZR_lt; omega.
+ * apply IZR_lt.
+ apply (Z.le_lt_trans _ _ _ Hbeta).
+ rewrite <- (Zmult_1_r beta) at 1.
+ apply Zmult_lt_compat_l; omega.
+- (* mag x < fexp2 (mag x) *)
+ casetype False; apply Nzx''.
+ now apply (round_N_small_pos beta _ _ _ (mag x)).
+Qed.
+
+Lemma round_round_zero :
+ forall fexp1 fexp2 : Z -> Z,
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ forall x,
+ 0 < x ->
+ (fexp1 (mag x) = mag x + 1 :> Z)%Z ->
+ x < bpow (mag x) - / 2 * ulp beta fexp2 x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x.
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 x Px Hf1.
+unfold round_round_eq.
+set (x'' := round beta fexp2 (Znearest choice2) x).
+set (u1 := ulp beta fexp1 x).
+set (u2 := ulp beta fexp2 x).
+intro Hx.
+assert (Hlx : bpow (mag x - 1) <= x < bpow (mag x)).
+{ destruct (mag x) as (ex,Hex); simpl.
+ rewrite <- (Rabs_right x); [|now apply Rle_ge; apply Rlt_le].
+ apply Hex.
+ now apply Rgt_not_eq. }
+rewrite (round_N_small_pos beta fexp1 choice1 x (mag x));
+ [|exact Hlx|omega].
+destruct (Req_dec x'' 0) as [Zx''|Nzx''];
+ [now rewrite Zx''; rewrite round_0; [reflexivity|apply valid_rnd_N]|].
+rewrite (round_N_small_pos beta _ _ x'' (mag x));
+ [reflexivity| |omega].
+split.
+- apply round_large_pos_ge_bpow.
+ + now apply valid_rnd_N.
+ + assert (0 <= x''); [|now fold x''; lra].
+ rewrite <- (round_0 beta fexp2 (Znearest choice2)).
+ now apply round_le; [|apply valid_rnd_N|apply Rlt_le].
+ + apply Rle_trans with (Rabs x);
+ [|now rewrite Rabs_right; [apply Rle_refl|apply Rle_ge; apply Rlt_le]].
+ destruct (mag x) as (ex,Hex); simpl; apply Hex.
+ now apply Rgt_not_eq.
+- replace x'' with (x + (x'' - x)) by ring.
+ replace (bpow _) with (bpow (mag x) - / 2 * u2 + / 2 * u2) by ring.
+ apply Rplus_lt_le_compat; [exact Hx|].
+ apply Rabs_le_inv.
+ now apply error_le_half_ulp.
+Qed.
+
+Lemma round_round_all_mid_cases :
+ forall fexp1 fexp2 : Z -> Z,
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ forall x,
+ 0 < x ->
+ (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z ->
+ ((fexp1 (mag x) = mag x + 1 :> Z)%Z ->
+ bpow (mag x) - / 2 * ulp beta fexp2 x <= x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x) ->
+ ((fexp1 (mag x) <= mag x)%Z ->
+ midp fexp1 x - / 2 * ulp beta fexp2 x <= x < midp fexp1 x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x) ->
+ ((fexp1 (mag x) <= mag x)%Z ->
+ x = midp fexp1 x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x) ->
+ ((fexp1 (mag x) <= mag x)%Z ->
+ midp fexp1 x < x <= midp fexp1 x + / 2 * ulp beta fexp2 x ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x) ->
+ round_round_eq fexp1 fexp2 choice1 choice2 x.
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 x Px Hf2.
+set (x' := round beta fexp1 Zfloor x).
+set (u1 := ulp beta fexp1 x).
+set (u2 := ulp beta fexp2 x).
+intros Cz Clt Ceq Cgt.
+destruct (Ztrichotomy (mag x) (fexp1 (mag x) - 1)) as [Hlt|[Heq|Hgt]].
+- (* mag x < fexp1 (mag x) - 1 *)
+ assert (H : (mag x <= fexp1 (mag x) - 2)%Z) by omega.
+ now apply round_round_really_zero.
+- (* mag x = fexp1 (mag x) - 1 *)
+ assert (H : (fexp1 (mag x) = (mag x + 1))%Z) by omega.
+ destruct (Rlt_or_le x (bpow (mag x) - / 2 * u2)) as [Hlt'|Hge'].
+ + now apply round_round_zero.
+ + now apply Cz.
+- (* mag x > fexp1 (mag x) - 1 *)
+ assert (H : (fexp1 (mag x) <= mag x)%Z) by omega.
+ destruct (Rtotal_order x (midp fexp1 x)) as [Hlt'|[Heq'|Hgt']].
+ + (* x < midp fexp1 x *)
+ destruct (Rlt_or_le x (midp fexp1 x - / 2 * u2)) as [Hlt''|Hle''].
+ * now apply round_round_lt_mid_further_place; [| | |omega| |].
+ * now apply Clt; [|split].
+ + (* x = midp fexp1 x *)
+ now apply Ceq.
+ + (* x > midp fexp1 x *)
+ destruct (Rle_or_lt x (midp fexp1 x + / 2 * u2)) as [Hlt''|Hle''].
+ * now apply Cgt; [|split].
+ * { destruct (generic_format_EM beta fexp1 x) as [Fx|Nfx].
+ - (* generic_format beta fexp1 x *)
+ unfold round_round_eq; rewrite (round_generic beta fexp2);
+ [reflexivity|now apply valid_rnd_N|].
+ now apply (generic_inclusion_mag beta fexp1); [omega|].
+ - (* ~ generic_format beta fexp1 x *)
+ assert (Hceil : round beta fexp1 Zceil x = x' + u1);
+ [now apply round_UP_DN_ulp|].
+ assert (Hf2' : (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z);
+ [omega|].
+ assert (midp' fexp1 x + / 2 * ulp beta fexp2 x < x);
+ [|now apply round_round_gt_mid_further_place].
+ revert Hle''; unfold midp, midp'; fold x'.
+ rewrite Hceil; fold u1; fold u2.
+ lra. }
+Qed.
+
+Lemma mag_div_disj :
+ forall x y : R,
+ 0 < x -> 0 < y ->
+ ((mag (x / y) = mag x - mag y :> Z)%Z
+ \/ (mag (x / y) = mag x - mag y + 1 :> Z)%Z).
+Proof.
+intros x y Px Py.
+generalize (mag_div beta x y (Rgt_not_eq _ _ Px) (Rgt_not_eq _ _ Py)).
+omega.
+Qed.
+
+Definition round_round_div_hyp fexp1 fexp2 :=
+ (forall ex, (fexp2 ex <= fexp1 ex - 1)%Z)
+ /\ (forall ex ey, (fexp1 ex < ex)%Z -> (fexp1 ey < ey)%Z ->
+ (fexp1 (ex - ey) <= ex - ey + 1)%Z ->
+ (fexp2 (ex - ey) <= fexp1 ex - ey)%Z)
+ /\ (forall ex ey, (fexp1 ex < ex)%Z -> (fexp1 ey < ey)%Z ->
+ (fexp1 (ex - ey + 1) <= ex - ey + 1 + 1)%Z ->
+ (fexp2 (ex - ey + 1) <= fexp1 ex - ey)%Z)
+ /\ (forall ex ey, (fexp1 ex < ex)%Z -> (fexp1 ey < ey)%Z ->
+ (fexp1 (ex - ey) <= ex - ey)%Z ->
+ (fexp2 (ex - ey) <= fexp1 (ex - ey)
+ + fexp1 ey - ey)%Z)
+ /\ (forall ex ey, (fexp1 ex < ex)%Z -> (fexp1 ey < ey)%Z ->
+ (fexp1 (ex - ey) = ex - ey + 1)%Z ->
+ (fexp2 (ex - ey) <= ex - ey - ey + fexp1 ey)%Z).
+
+Lemma round_round_div_aux0 :
+ forall fexp1 fexp2 : Z -> Z,
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_div_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < x -> 0 < y ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z ->
+ ~ (bpow (mag (x / y)) - / 2 * ulp beta fexp2 (x / y) <= x / y).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Px Py Fx Fy Hf1.
+assert (Hfx : (fexp1 (mag x) < mag x)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+assert (Hfy : (fexp1 (mag y) < mag y)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+set (p := bpow (mag (x / y))).
+set (u2 := bpow (fexp2 (mag (x / y)))).
+revert Fx Fy.
+unfold generic_format, F2R, scaled_mantissa, cexp; simpl.
+set (mx := Ztrunc (x * bpow (- fexp1 (mag x)))).
+set (my := Ztrunc (y * bpow (- fexp1 (mag y)))).
+intros Fx Fy.
+rewrite ulp_neq_0.
+2: apply Rmult_integral_contrapositive_currified; [now apply Rgt_not_eq|idtac].
+2: now apply Rinv_neq_0_compat, Rgt_not_eq.
+intro Hl.
+assert (Hr : x / y < p);
+ [now apply Rabs_lt_inv; apply bpow_mag_gt|].
+apply (Rlt_irrefl (p - / 2 * u2)).
+apply (Rle_lt_trans _ _ _ Hl).
+apply (Rmult_lt_reg_r y _ _ Py).
+unfold Rdiv; rewrite Rmult_assoc.
+rewrite Rinv_l; [|now apply Rgt_not_eq]; rewrite Rmult_1_r.
+destruct (Zle_or_lt Z0 (fexp1 (mag x) - mag (x / y)
+ - fexp1 (mag y))%Z) as [He|He].
+- (* mag (x / y) + fexp1 (mag y) <= fexp1 (mag x) *)
+ apply Rle_lt_trans with (p * y - p * bpow (fexp1 (mag y))).
+ + rewrite Fx; rewrite Fy at 1.
+ rewrite <- Rmult_assoc.
+ rewrite (Rmult_comm p).
+ unfold p; bpow_simplify.
+ apply (Rmult_le_reg_r (bpow (- mag (x / y) - fexp1 (mag y))));
+ [now apply bpow_gt_0|].
+ rewrite Rmult_minus_distr_r.
+ bpow_simplify.
+ rewrite <- IZR_Zpower; [|exact He].
+ rewrite <- mult_IZR.
+ rewrite <- minus_IZR.
+ apply IZR_le.
+ apply (Zplus_le_reg_r _ _ 1); ring_simplify.
+ apply Zlt_le_succ.
+ apply lt_IZR.
+ rewrite mult_IZR.
+ rewrite IZR_Zpower; [|exact He].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag y) + mag (x / y))));
+ [now apply bpow_gt_0|].
+ bpow_simplify.
+ rewrite <- Fx.
+ rewrite bpow_plus.
+ rewrite <- Rmult_assoc; rewrite <- Fy.
+ fold p.
+ apply (Rmult_lt_reg_r (/ y)); [now apply Rinv_0_lt_compat|].
+ field_simplify; lra.
+ + rewrite Rmult_minus_distr_r.
+ unfold Rminus; apply Rplus_lt_compat_l.
+ apply Ropp_lt_contravar.
+ apply Rlt_le_trans with (u2 * bpow (mag y)).
+ * rewrite <- (Rmult_1_l (u2 * _)).
+ rewrite Rmult_assoc.
+ { apply Rmult_lt_compat.
+ - lra.
+ - now apply Rmult_le_pos; [apply bpow_ge_0|apply Rlt_le].
+ - lra.
+ - apply Rmult_lt_compat_l; [now apply bpow_gt_0|].
+ apply Rabs_lt_inv.
+ apply bpow_mag_gt. }
+ * unfold u2, p, ulp, cexp; bpow_simplify; apply bpow_le.
+ apply (Zplus_le_reg_r _ _ (- mag y)); ring_simplify.
+ rewrite (Zplus_comm (- _)); fold (Zminus (mag (x / y)) (mag y)).
+ destruct (mag_div_disj x y Px Py) as [Hxy|Hxy]; rewrite Hxy;
+ [now apply Hexp; [| |rewrite <- Hxy]|].
+ replace (_ - _ + 1)%Z with ((mag x + 1) - mag y)%Z by ring.
+ apply Hexp.
+ { now assert (fexp1 (mag x + 1) <= mag x)%Z;
+ [apply valid_exp|omega]. }
+ { assumption. }
+ replace (_ + 1 - _)%Z with (mag x - mag y + 1)%Z by ring.
+ now rewrite <- Hxy.
+- (* fexp1 (mag x) < mag (x / y) + fexp1 (mag y) *)
+ apply Rle_lt_trans with (p * y - bpow (fexp1 (mag x))).
+ + rewrite Fx at 1; rewrite Fy at 1.
+ apply (Rmult_le_reg_r (bpow (- fexp1 (mag x))));
+ [now apply bpow_gt_0|].
+ rewrite Rmult_minus_distr_r.
+ bpow_simplify.
+ rewrite (Rmult_comm p).
+ unfold p; bpow_simplify.
+ rewrite <- IZR_Zpower; [|omega].
+ rewrite <- mult_IZR.
+ rewrite <- minus_IZR.
+ apply IZR_le.
+ apply (Zplus_le_reg_r _ _ 1); ring_simplify.
+ apply Zlt_le_succ.
+ apply lt_IZR.
+ rewrite mult_IZR.
+ rewrite IZR_Zpower; [|omega].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x))));
+ [now apply bpow_gt_0|bpow_simplify].
+ rewrite <- Fx.
+ rewrite Zplus_comm; rewrite bpow_plus.
+ rewrite <- Rmult_assoc; rewrite <- Fy.
+ fold p.
+ apply (Rmult_lt_reg_r (/ y)); [now apply Rinv_0_lt_compat|].
+ field_simplify; lra.
+ + rewrite Rmult_minus_distr_r.
+ unfold Rminus; apply Rplus_lt_compat_l.
+ apply Ropp_lt_contravar.
+ apply Rlt_le_trans with (u2 * bpow (mag y)).
+ * rewrite <- (Rmult_1_l (u2 * _)).
+ rewrite Rmult_assoc.
+ { apply Rmult_lt_compat.
+ - lra.
+ - now apply Rmult_le_pos; [apply bpow_ge_0|apply Rlt_le].
+ - lra.
+ - apply Rmult_lt_compat_l; [now apply bpow_gt_0|].
+ apply Rabs_lt_inv.
+ apply bpow_mag_gt. }
+ * unfold u2, p, ulp, cexp; bpow_simplify; apply bpow_le.
+ apply (Zplus_le_reg_r _ _ (- mag y)); ring_simplify.
+ rewrite (Zplus_comm (- _)); fold (Zminus (mag (x / y)) (mag y)).
+ destruct (mag_div_disj x y Px Py) as [Hxy|Hxy]; rewrite Hxy;
+ apply Hexp; try assumption; rewrite <- Hxy; rewrite Hf1; apply Z.le_refl.
+Qed.
+
+Lemma round_round_div_aux1 :
+ forall fexp1 fexp2 : Z -> Z,
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_div_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < x -> 0 < y ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ (fexp1 (mag (x / y)) <= mag (x / y))%Z ->
+ ~ (midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y)
+ <= x / y
+ < midp fexp1 (x / y)).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Px Py Fx Fy Hf1.
+assert (Hfx : (fexp1 (mag x) < mag x)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+assert (Hfy : (fexp1 (mag y) < mag y)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+assert (S : (x / y <> 0)%R).
+apply Rmult_integral_contrapositive_currified; [now apply Rgt_not_eq|idtac].
+now apply Rinv_neq_0_compat, Rgt_not_eq.
+cut (~ (/ 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y))
+ <= x / y - round beta fexp1 Zfloor (x / y)
+ < / 2 * ulp beta fexp1 (x / y))).
+{ intro H; intro H'; apply H; split.
+ - apply (Rplus_le_reg_l (round beta fexp1 Zfloor (x / y))).
+ ring_simplify.
+ apply H'.
+ - apply (Rplus_lt_reg_l (round beta fexp1 Zfloor (x / y))).
+ ring_simplify.
+ apply H'. }
+set (u1 := bpow (fexp1 (mag (x / y)))).
+set (u2 := bpow (fexp2 (mag (x / y)))).
+set (x' := round beta fexp1 Zfloor (x / y)).
+rewrite 2!ulp_neq_0; trivial.
+revert Fx Fy.
+unfold generic_format, F2R, scaled_mantissa, cexp; simpl.
+set (mx := Ztrunc (x * bpow (- fexp1 (mag x)))).
+set (my := Ztrunc (y * bpow (- fexp1 (mag y)))).
+intros Fx Fy.
+intro Hlr.
+apply (Rlt_irrefl (/ 2 * (u1 - u2))).
+apply (Rle_lt_trans _ _ _ (proj1 Hlr)).
+apply (Rplus_lt_reg_r x'); ring_simplify.
+apply (Rmult_lt_reg_r y _ _ Py).
+unfold Rdiv; rewrite Rmult_assoc.
+rewrite Rinv_l; [|now apply Rgt_not_eq]; rewrite Rmult_1_r.
+rewrite Rmult_minus_distr_r; rewrite Rmult_plus_distr_r.
+apply (Rmult_lt_reg_l 2); [lra|].
+rewrite Rmult_minus_distr_l; rewrite Rmult_plus_distr_l.
+do 5 rewrite <- Rmult_assoc.
+rewrite Rinv_r; [|lra]; do 2 rewrite Rmult_1_l.
+destruct (Zle_or_lt Z0 (fexp1 (mag x) - fexp1 (mag (x / y))
+ - fexp1 (mag y))%Z) as [He|He].
+- (* fexp1 (mag (x / y)) + fexp1 (mag y)) <= fexp1 (mag x) *)
+ apply Rle_lt_trans with (2 * x' * y + u1 * y
+ - bpow (fexp1 (mag (x / y))
+ + fexp1 (mag y))).
+ + rewrite Fx at 1; rewrite Fy at 1 2.
+ apply (Rmult_le_reg_r (bpow (- fexp1 (mag (x / y))
+ - fexp1 (mag y))));
+ [now apply bpow_gt_0|].
+ rewrite Rmult_minus_distr_r; rewrite (Rmult_plus_distr_r (_ * _ * _)).
+ bpow_simplify.
+ replace (2 * x' * _ * _)
+ with (2 * IZR my * x' * bpow (- fexp1 (mag (x / y)))) by ring.
+ rewrite (Rmult_comm u1).
+ unfold x', u1, round, F2R, ulp, scaled_mantissa, cexp; simpl.
+ bpow_simplify.
+ rewrite <- IZR_Zpower; [|exact He].
+ do 4 rewrite <- mult_IZR.
+ rewrite <- plus_IZR.
+ rewrite <- minus_IZR.
+ apply IZR_le.
+ apply (Zplus_le_reg_r _ _ 1); ring_simplify.
+ apply Zlt_le_succ.
+ apply lt_IZR.
+ rewrite plus_IZR.
+ do 4 rewrite mult_IZR; simpl.
+ rewrite IZR_Zpower; [|exact He].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag (x / y))
+ + fexp1 (mag y))));
+ [now apply bpow_gt_0|bpow_simplify].
+ rewrite Rmult_assoc.
+ rewrite <- Fx.
+ rewrite (Rmult_plus_distr_r _ _ (Raux.bpow _ _)).
+ rewrite Rmult_assoc.
+ rewrite bpow_plus.
+ rewrite <- (Rmult_assoc (IZR (Zfloor _))).
+ change (IZR (Zfloor _) * _) with x'.
+ do 2 rewrite (Rmult_comm _ (bpow (fexp1 (mag y)))).
+ rewrite Rmult_assoc.
+ do 2 rewrite <- (Rmult_assoc (IZR my)).
+ rewrite <- Fy.
+ change (bpow _) with u1.
+ apply (Rmult_lt_reg_l (/ 2)); [lra|].
+ rewrite Rmult_plus_distr_l.
+ do 4 rewrite <- Rmult_assoc.
+ rewrite Rinv_l; [|lra]; do 2 rewrite Rmult_1_l.
+ apply (Rplus_lt_reg_r (- y * x')); ring_simplify.
+ apply (Rmult_lt_reg_l (/ y)); [now apply Rinv_0_lt_compat|].
+ rewrite Rmult_minus_distr_l.
+ do 3 rewrite <- Rmult_assoc.
+ rewrite Rinv_l; [|now apply Rgt_not_eq]; do 2 rewrite Rmult_1_l.
+ now rewrite Rmult_comm.
+ + apply Rplus_lt_compat_l.
+ apply Ropp_lt_contravar.
+ apply Rlt_le_trans with (u2 * bpow (mag y)).
+ * { apply Rmult_lt_compat_l.
+ - apply bpow_gt_0.
+ - apply Rabs_lt_inv.
+ apply bpow_mag_gt. }
+ * unfold u2, ulp, cexp; bpow_simplify; apply bpow_le.
+ apply (Zplus_le_reg_r _ _ (- mag y)); ring_simplify.
+ rewrite <- Zplus_assoc; rewrite (Zplus_comm (- _)).
+ destruct (mag_div_disj x y Px Py) as [Hxy|Hxy]; rewrite Hxy;
+ [now apply Hexp; [| |rewrite <- Hxy]|].
+ replace (_ - _ + 1)%Z with ((mag x + 1) - mag y)%Z by ring.
+ apply Hexp.
+ { now assert (fexp1 (mag x + 1) <= mag x)%Z;
+ [apply valid_exp|omega]. }
+ { assumption. }
+ replace (_ + 1 - _)%Z with (mag x - mag y + 1)%Z by ring.
+ now rewrite <- Hxy.
+- (* fexp1 (mag x) < fexp1 (mag (x / y)) + fexp1 (mag y) *)
+ apply Rle_lt_trans with (2 * x' * y + u1 * y - bpow (fexp1 (mag x))).
+ + rewrite Fx at 1; rewrite Fy at 1 2.
+ apply (Rmult_le_reg_r (bpow (- fexp1 (mag x))));
+ [now apply bpow_gt_0|].
+ rewrite Rmult_minus_distr_r; rewrite (Rmult_plus_distr_r (_ * _ * _)).
+ bpow_simplify.
+ replace (2 * x' * _ * _)
+ with (2 * IZR my * x' * bpow (fexp1 (mag y) - fexp1 (mag x))) by ring.
+ rewrite (Rmult_comm u1).
+ unfold x', u1, round, F2R, ulp, scaled_mantissa, cexp; simpl.
+ bpow_simplify.
+ rewrite <- (IZR_Zpower _ (_ - _)%Z); [|omega].
+ do 5 rewrite <- mult_IZR.
+ rewrite <- plus_IZR.
+ rewrite <- minus_IZR.
+ apply IZR_le.
+ apply (Zplus_le_reg_r _ _ 1); ring_simplify.
+ apply Zlt_le_succ.
+ apply lt_IZR.
+ rewrite plus_IZR.
+ do 5 rewrite mult_IZR; simpl.
+ rewrite IZR_Zpower; [|omega].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x))));
+ [now apply bpow_gt_0|].
+ rewrite Rmult_assoc.
+ rewrite <- Fx.
+ rewrite (Rmult_plus_distr_r _ _ (Raux.bpow _ _)).
+ bpow_simplify.
+ rewrite Rmult_assoc.
+ rewrite bpow_plus.
+ rewrite <- (Rmult_assoc (IZR (Zfloor _))).
+ change (IZR (Zfloor _) * _) with x'.
+ do 2 rewrite (Rmult_comm _ (bpow (fexp1 (mag y)))).
+ rewrite Rmult_assoc.
+ do 2 rewrite <- (Rmult_assoc (IZR my)).
+ rewrite <- Fy.
+ change (bpow _) with u1.
+ apply (Rmult_lt_reg_l (/ 2)); [lra|].
+ rewrite Rmult_plus_distr_l.
+ do 4 rewrite <- Rmult_assoc.
+ rewrite Rinv_l; [|lra]; do 2 rewrite Rmult_1_l.
+ apply (Rplus_lt_reg_r (- y * x')); ring_simplify.
+ apply (Rmult_lt_reg_l (/ y)); [now apply Rinv_0_lt_compat|].
+ rewrite Rmult_minus_distr_l.
+ do 3 rewrite <- Rmult_assoc.
+ rewrite Rinv_l; [|now apply Rgt_not_eq]; do 2 rewrite Rmult_1_l.
+ now rewrite Rmult_comm.
+ + apply Rplus_lt_compat_l.
+ apply Ropp_lt_contravar.
+ apply Rlt_le_trans with (u2 * bpow (mag y)).
+ * { apply Rmult_lt_compat_l.
+ - apply bpow_gt_0.
+ - apply Rabs_lt_inv.
+ apply bpow_mag_gt. }
+ * unfold u2, ulp, cexp; bpow_simplify; apply bpow_le.
+ apply (Zplus_le_reg_r _ _ (- mag y)); ring_simplify.
+ rewrite (Zplus_comm (- _)).
+ destruct (mag_div_disj x y Px Py) as [Hxy|Hxy]; rewrite Hxy;
+ apply Hexp; try assumption; rewrite <- Hxy; omega.
+Qed.
+
+Lemma round_round_div_aux2 :
+ forall fexp1 fexp2 : Z -> Z,
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ round_round_div_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < x -> 0 < y ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ (fexp1 (mag (x / y)) <= mag (x / y))%Z ->
+ ~ (midp fexp1 (x / y)
+ < x / y
+ <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Hexp x y Px Py Fx Fy Hf1.
+assert (Hfx : (fexp1 (mag x) < mag x)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+assert (Hfy : (fexp1 (mag y) < mag y)%Z);
+ [now apply mag_generic_gt; [|apply Rgt_not_eq|]|].
+cut (~ (/ 2 * ulp beta fexp1 (x / y)
+ < x / y - round beta fexp1 Zfloor (x / y)
+ <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y)))).
+{ intro H; intro H'; apply H; split.
+ - apply (Rplus_lt_reg_l (round beta fexp1 Zfloor (x / y))).
+ ring_simplify.
+ apply H'.
+ - apply (Rplus_le_reg_l (round beta fexp1 Zfloor (x / y))).
+ ring_simplify.
+ apply H'. }
+set (u1 := bpow (fexp1 (mag (x / y)))).
+set (u2 := bpow (fexp2 (mag (x / y)))).
+set (x' := round beta fexp1 Zfloor (x / y)).
+assert (S : (x / y <> 0)%R).
+apply Rmult_integral_contrapositive_currified; [now apply Rgt_not_eq|idtac].
+now apply Rinv_neq_0_compat, Rgt_not_eq.
+rewrite 2!ulp_neq_0; trivial.
+revert Fx Fy.
+unfold generic_format, F2R, scaled_mantissa, cexp; simpl.
+set (mx := Ztrunc (x * bpow (- fexp1 (mag x)))).
+set (my := Ztrunc (y * bpow (- fexp1 (mag y)))).
+intros Fx Fy.
+intro Hlr.
+apply (Rlt_irrefl (/ 2 * (u1 + u2))).
+apply Rlt_le_trans with (x / y - x'); [|now apply Hlr].
+apply (Rplus_lt_reg_r x'); ring_simplify.
+apply (Rmult_lt_reg_r y _ _ Py).
+unfold Rdiv; rewrite Rmult_assoc.
+rewrite Rinv_l; [|now apply Rgt_not_eq]; rewrite Rmult_1_r.
+do 2 rewrite Rmult_plus_distr_r.
+apply (Rmult_lt_reg_l 2); [lra|].
+do 2 rewrite Rmult_plus_distr_l.
+do 5 rewrite <- Rmult_assoc.
+rewrite Rinv_r; [|lra]; do 2 rewrite Rmult_1_l.
+destruct (Zle_or_lt Z0 (fexp1 (mag x) - fexp1 (mag (x / y))
+ - fexp1 (mag y))%Z) as [He|He].
+- (* fexp1 (mag (x / y)) + fexp1 (mag y) <= fexp1 (mag x) *)
+ apply Rlt_le_trans with (u1 * y + bpow (fexp1 (mag (x / y))
+ + fexp1 (mag y))
+ + 2 * x' * y).
+ + apply Rplus_lt_compat_r, Rplus_lt_compat_l.
+ apply Rlt_le_trans with (u2 * bpow (mag y)).
+ * { apply Rmult_lt_compat_l.
+ - apply bpow_gt_0.
+ - apply Rabs_lt_inv.
+ apply bpow_mag_gt. }
+ * unfold u2, ulp, cexp; bpow_simplify; apply bpow_le.
+ apply (Zplus_le_reg_r _ _ (- mag y)); ring_simplify.
+ rewrite <- Zplus_assoc; rewrite (Zplus_comm (- _)).
+ destruct (mag_div_disj x y Px Py) as [Hxy|Hxy]; rewrite Hxy;
+ [now apply Hexp; [| |rewrite <- Hxy]|].
+ replace (_ - _ + 1)%Z with ((mag x + 1) - mag y)%Z by ring.
+ apply Hexp.
+ { now assert (fexp1 (mag x + 1) <= mag x)%Z;
+ [apply valid_exp|omega]. }
+ { assumption. }
+ replace (_ + 1 - _)%Z with (mag x - mag y + 1)%Z by ring.
+ now rewrite <- Hxy.
+ + apply Rge_le; rewrite Fx at 1; apply Rle_ge.
+ replace (u1 * y) with (u1 * (IZR my * bpow (fexp1 (mag y))));
+ [|now apply eq_sym; rewrite Fy at 1].
+ replace (2 * x' * y) with (2 * x' * (IZR my * bpow (fexp1 (mag y))));
+ [|now apply eq_sym; rewrite Fy at 1].
+ apply (Rmult_le_reg_r (bpow (- fexp1 (mag (x / y))
+ - fexp1 (mag y))));
+ [now apply bpow_gt_0|].
+ do 2 rewrite Rmult_plus_distr_r.
+ bpow_simplify.
+ rewrite (Rmult_comm u1).
+ unfold u1, ulp, cexp; bpow_simplify.
+ rewrite (Rmult_assoc 2).
+ rewrite (Rmult_comm x').
+ rewrite (Rmult_assoc 2).
+ unfold x', round, F2R, scaled_mantissa, cexp; simpl.
+ bpow_simplify.
+ rewrite <- (IZR_Zpower _ (_ - _)%Z); [|exact He].
+ do 4 rewrite <- mult_IZR.
+ do 2 rewrite <- plus_IZR.
+ apply IZR_le.
+ rewrite Zplus_comm, Zplus_assoc.
+ apply Zlt_le_succ.
+ apply lt_IZR.
+ rewrite plus_IZR.
+ do 4 rewrite mult_IZR; simpl.
+ rewrite IZR_Zpower; [|exact He].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag y))));
+ [now apply bpow_gt_0|].
+ rewrite Rmult_plus_distr_r.
+ rewrite (Rmult_comm _ (IZR _)).
+ do 2 rewrite Rmult_assoc.
+ rewrite <- Fy.
+ bpow_simplify.
+ unfold Zminus; rewrite bpow_plus.
+ rewrite (Rmult_assoc _ (IZR mx)).
+ rewrite <- (Rmult_assoc (IZR mx)).
+ rewrite <- Fx.
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag (x / y)))));
+ [now apply bpow_gt_0|].
+ rewrite Rmult_plus_distr_r.
+ bpow_simplify.
+ rewrite (Rmult_comm _ y).
+ do 2 rewrite Rmult_assoc.
+ change (IZR (Zfloor _) * _) with x'.
+ change (bpow _) with u1.
+ apply (Rmult_lt_reg_l (/ 2)); [lra|].
+ rewrite Rmult_plus_distr_l.
+ do 4 rewrite <- Rmult_assoc.
+ rewrite Rinv_l; [|lra]; do 2 rewrite Rmult_1_l.
+ apply (Rplus_lt_reg_r (- y * x')); ring_simplify.
+ apply (Rmult_lt_reg_l (/ y)); [now apply Rinv_0_lt_compat|].
+ rewrite Rmult_plus_distr_l.
+ do 3 rewrite <- Rmult_assoc.
+ rewrite Ropp_mult_distr_r_reverse.
+ rewrite Ropp_mult_distr_l_reverse.
+ rewrite Rinv_l; [|now apply Rgt_not_eq]; do 2 rewrite Rmult_1_l.
+ rewrite (Rmult_comm (/ y)).
+ now rewrite (Rplus_comm (- x')).
+- (* fexp1 (mag x) < fexp1 (mag (x / y)) + fexp1 (mag y) *)
+ apply Rlt_le_trans with (2 * x' * y + u1 * y + bpow (fexp1 (mag x))).
+ + rewrite Rplus_comm, Rplus_assoc; do 2 apply Rplus_lt_compat_l.
+ apply Rlt_le_trans with (u2 * bpow (mag y)).
+ * apply Rmult_lt_compat_l.
+ now apply bpow_gt_0.
+ now apply Rabs_lt_inv; apply bpow_mag_gt.
+ * unfold u2, ulp, cexp; bpow_simplify; apply bpow_le.
+ apply (Zplus_le_reg_r _ _ (- mag y)); ring_simplify.
+ rewrite (Zplus_comm (- _)).
+ destruct (mag_div_disj x y Px Py) as [Hxy|Hxy]; rewrite Hxy;
+ apply Hexp; try assumption; rewrite <- Hxy; omega.
+ + apply Rge_le; rewrite Fx at 1; apply Rle_ge.
+ rewrite Fy at 1 2.
+ apply (Rmult_le_reg_r (bpow (- fexp1 (mag x))));
+ [now apply bpow_gt_0|].
+ do 2 rewrite Rmult_plus_distr_r.
+ bpow_simplify.
+ replace (2 * x' * _ * _)
+ with (2 * IZR my * x' * bpow (fexp1 (mag y) - fexp1 (mag x))) by ring.
+ rewrite (Rmult_comm u1).
+ unfold x', u1, round, F2R, ulp, scaled_mantissa, cexp; simpl.
+ bpow_simplify.
+ rewrite <- (IZR_Zpower _ (_ - _)%Z); [|omega].
+ do 5 rewrite <- mult_IZR.
+ do 2 rewrite <- plus_IZR.
+ apply IZR_le.
+ apply Zlt_le_succ.
+ apply lt_IZR.
+ rewrite plus_IZR.
+ do 5 rewrite mult_IZR; simpl.
+ rewrite IZR_Zpower; [|omega].
+ apply (Rmult_lt_reg_r (bpow (fexp1 (mag x))));
+ [now apply bpow_gt_0|].
+ rewrite (Rmult_assoc _ (IZR mx)).
+ rewrite <- Fx.
+ rewrite Rmult_plus_distr_r.
+ bpow_simplify.
+ rewrite bpow_plus.
+ rewrite Rmult_assoc.
+ rewrite <- (Rmult_assoc (IZR _)).
+ change (IZR _ * bpow _) with x'.
+ do 2 rewrite (Rmult_comm _ (bpow (fexp1 (mag y)))).
+ rewrite Rmult_assoc.
+ do 2 rewrite <- (Rmult_assoc (IZR my)).
+ rewrite <- Fy.
+ change (bpow _) with u1.
+ apply (Rmult_lt_reg_l (/ 2)); [lra|].
+ rewrite Rmult_plus_distr_l.
+ do 4 rewrite <- Rmult_assoc.
+ rewrite Rinv_l; [|lra]; do 2 rewrite Rmult_1_l.
+ apply (Rplus_lt_reg_r (- y * x')); ring_simplify.
+ apply (Rmult_lt_reg_l (/ y)); [now apply Rinv_0_lt_compat|].
+ rewrite Rmult_plus_distr_l.
+ do 3 rewrite <- Rmult_assoc.
+ rewrite Ropp_mult_distr_r_reverse.
+ rewrite Ropp_mult_distr_l_reverse.
+ rewrite Rinv_l; [|now apply Rgt_not_eq]; do 2 rewrite Rmult_1_l.
+ rewrite (Rmult_comm (/ y)).
+ now rewrite (Rplus_comm (- x')).
+Qed.
+
+Lemma round_round_div_aux :
+ forall fexp1 fexp2 : Z -> Z,
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ (exists n, (beta = 2 * n :> Z)%Z) ->
+ round_round_div_hyp fexp1 fexp2 ->
+ forall x y,
+ 0 < x -> 0 < y ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x / y).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Ebeta Hexp x y Px Py Fx Fy.
+assert (Pxy : 0 < x / y).
+{ apply Rmult_lt_0_compat; [exact Px|].
+ now apply Rinv_0_lt_compat. }
+apply round_round_all_mid_cases.
+- exact Vfexp1.
+- exact Vfexp2.
+- exact Pxy.
+- apply Hexp.
+- intros Hf1 Hlxy.
+ casetype False.
+ now apply (round_round_div_aux0 fexp1 fexp2 _ _ choice1 choice2 Hexp x y).
+- intros Hf1 Hlxy.
+ casetype False.
+ now apply (round_round_div_aux1 fexp1 fexp2 _ _ choice1 choice2 Hexp x y).
+- intro H.
+ apply round_round_eq_mid_beta_even; try assumption.
+ apply Hexp.
+- intros Hf1 Hlxy.
+ casetype False.
+ now apply (round_round_div_aux2 fexp1 fexp2 _ _ choice1 choice2 Hexp x y).
+Qed.
+
+Lemma round_round_div :
+ forall fexp1 fexp2 : Z -> Z,
+ Valid_exp fexp1 -> Valid_exp fexp2 ->
+ forall (choice1 choice2 : Z -> bool),
+ (exists n, (beta = 2 * n :> Z)%Z) ->
+ round_round_div_hyp fexp1 fexp2 ->
+ forall x y,
+ y <> 0 ->
+ generic_format beta fexp1 x ->
+ generic_format beta fexp1 y ->
+ round_round_eq fexp1 fexp2 choice1 choice2 (x / y).
+Proof.
+intros fexp1 fexp2 Vfexp1 Vfexp2 choice1 choice2 Ebeta Hexp x y Nzy Fx Fy.
+unfold round_round_eq.
+destruct (Rtotal_order x 0) as [Nx|[Zx|Px]].
+- (* x < 0 *)
+ destruct (Rtotal_order y 0) as [Ny|[Zy|Py]].
+ + (* y < 0 *)
+ rewrite <- (Ropp_involutive x).
+ rewrite <- (Ropp_involutive y).
+ rewrite Ropp_div.
+ unfold Rdiv; rewrite <- Ropp_inv_permute; [|lra].
+ rewrite Ropp_mult_distr_r_reverse.
+ rewrite Ropp_involutive.
+ fold ((- x) / (- y)).
+ apply Ropp_lt_contravar in Nx.
+ apply Ropp_lt_contravar in Ny.
+ rewrite Ropp_0 in Nx, Ny.
+ apply generic_format_opp in Fx.
+ apply generic_format_opp in Fy.
+ now apply round_round_div_aux.
+ + (* y = 0 *)
+ now casetype False; apply Nzy.
+ + (* y > 0 *)
+ rewrite <- (Ropp_involutive x).
+ rewrite Ropp_div.
+ do 3 rewrite round_N_opp.
+ apply Ropp_eq_compat.
+ apply Ropp_lt_contravar in Nx.
+ rewrite Ropp_0 in Nx.
+ apply generic_format_opp in Fx.
+ now apply round_round_div_aux.
+- (* x = 0 *)
+ rewrite Zx.
+ unfold Rdiv; rewrite Rmult_0_l.
+ now rewrite round_0; [|apply valid_rnd_N].
+- (* x > 0 *)
+ destruct (Rtotal_order y 0) as [Ny|[Zy|Py]].
+ + (* y < 0 *)
+ rewrite <- (Ropp_involutive y).
+ unfold Rdiv; rewrite <- Ropp_inv_permute; [|lra].
+ rewrite Ropp_mult_distr_r_reverse.
+ do 3 rewrite round_N_opp.
+ apply Ropp_eq_compat.
+ apply Ropp_lt_contravar in Ny.
+ rewrite Ropp_0 in Ny.
+ apply generic_format_opp in Fy.
+ now apply round_round_div_aux.
+ + (* y = 0 *)
+ now casetype False; apply Nzy.
+ + (* y > 0 *)
+ now apply round_round_div_aux.
+Qed.
+
+Section Double_round_div_FLX.
+
+Variable prec : Z.
+Variable prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FLX_round_round_div_hyp :
+ (2 * prec <= prec')%Z ->
+ round_round_div_hyp (FLX_exp prec) (FLX_exp prec').
+Proof.
+intros Hprec.
+unfold Prec_gt_0 in prec_gt_0_.
+unfold FLX_exp.
+unfold round_round_div_hyp.
+split; [now intro ex; omega|].
+split; [|split; [|split]]; intros ex ey; omega.
+Qed.
+
+Theorem round_round_div_FLX :
+ forall choice1 choice2,
+ (exists n, (beta = 2 * n :> Z)%Z) ->
+ (2 * prec <= prec')%Z ->
+ forall x y,
+ y <> 0 ->
+ FLX_format beta prec x -> FLX_format beta prec y ->
+ round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x / y).
+Proof.
+intros choice1 choice2 Ebeta Hprec x y Nzy Fx Fy.
+apply round_round_div.
+- now apply FLX_exp_valid.
+- now apply FLX_exp_valid.
+- exact Ebeta.
+- now apply FLX_round_round_div_hyp.
+- exact Nzy.
+- now apply generic_format_FLX.
+- now apply generic_format_FLX.
+Qed.
+
+End Double_round_div_FLX.
+
+Section Double_round_div_FLT.
+
+Variable emin prec : Z.
+Variable emin' prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FLT_round_round_div_hyp :
+ (emin' <= emin - prec - 2)%Z ->
+ (2 * prec <= prec')%Z ->
+ round_round_div_hyp (FLT_exp emin prec) (FLT_exp emin' prec').
+Proof.
+intros Hemin Hprec.
+unfold FLT_exp.
+unfold Prec_gt_0 in prec_gt_0_.
+unfold round_round_div_hyp.
+split; [intro ex|split; [|split; [|split]]; intros ex ey].
+- generalize (Zmax_spec (ex - prec') emin').
+ generalize (Zmax_spec (ex - prec) emin).
+ omega.
+- generalize (Zmax_spec (ex - prec) emin).
+ generalize (Zmax_spec (ey - prec) emin).
+ generalize (Zmax_spec (ex - ey - prec) emin).
+ generalize (Zmax_spec (ex - ey - prec') emin').
+ omega.
+- generalize (Zmax_spec (ex - prec) emin).
+ generalize (Zmax_spec (ey - prec) emin).
+ generalize (Zmax_spec (ex - ey + 1 - prec) emin).
+ generalize (Zmax_spec (ex - ey + 1 - prec') emin').
+ omega.
+- generalize (Zmax_spec (ex - prec) emin).
+ generalize (Zmax_spec (ey - prec) emin).
+ generalize (Zmax_spec (ex - ey - prec) emin).
+ generalize (Zmax_spec (ex - ey - prec') emin').
+ omega.
+- generalize (Zmax_spec (ex - prec) emin).
+ generalize (Zmax_spec (ey - prec) emin).
+ generalize (Zmax_spec (ex - ey - prec) emin).
+ generalize (Zmax_spec (ex - ey - prec') emin').
+ omega.
+Qed.
+
+Theorem round_round_div_FLT :
+ forall choice1 choice2,
+ (exists n, (beta = 2 * n :> Z)%Z) ->
+ (emin' <= emin - prec - 2)%Z ->
+ (2 * prec <= prec')%Z ->
+ forall x y,
+ y <> 0 ->
+ FLT_format beta emin prec x -> FLT_format beta emin prec y ->
+ round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec')
+ choice1 choice2 (x / y).
+Proof.
+intros choice1 choice2 Ebeta Hemin Hprec x y Nzy Fx Fy.
+apply round_round_div.
+- now apply FLT_exp_valid.
+- now apply FLT_exp_valid.
+- exact Ebeta.
+- now apply FLT_round_round_div_hyp.
+- exact Nzy.
+- now apply generic_format_FLT.
+- now apply generic_format_FLT.
+Qed.
+
+End Double_round_div_FLT.
+
+Section Double_round_div_FTZ.
+
+Variable emin prec : Z.
+Variable emin' prec' : Z.
+
+Context { prec_gt_0_ : Prec_gt_0 prec }.
+Context { prec_gt_0_' : Prec_gt_0 prec' }.
+
+Lemma FTZ_round_round_div_hyp :
+ (emin' + prec' <= emin - 1)%Z ->
+ (2 * prec <= prec')%Z ->
+ round_round_div_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec').
+Proof.
+intros Hemin Hprec.
+unfold FTZ_exp.
+unfold Prec_gt_0 in prec_gt_0_.
+unfold Prec_gt_0 in prec_gt_0_.
+unfold round_round_div_hyp.
+split; [intro ex|split; [|split; [|split]]; intros ex ey].
+- destruct (Z.ltb_spec (ex - prec') emin');
+ destruct (Z.ltb_spec (ex - prec) emin);
+ omega.
+- destruct (Z.ltb_spec (ex - prec) emin);
+ destruct (Z.ltb_spec (ey - prec) emin);
+ destruct (Z.ltb_spec (ex - ey - prec) emin);
+ destruct (Z.ltb_spec (ex - ey - prec') emin');
+ omega.
+- destruct (Z.ltb_spec (ex - prec) emin);
+ destruct (Z.ltb_spec (ey - prec) emin);
+ destruct (Z.ltb_spec (ex - ey + 1 - prec) emin);
+ destruct (Z.ltb_spec (ex - ey + 1 - prec') emin');
+ omega.
+- destruct (Z.ltb_spec (ex - prec) emin);
+ destruct (Z.ltb_spec (ey - prec) emin);
+ destruct (Z.ltb_spec (ex - ey - prec) emin);
+ destruct (Z.ltb_spec (ex - ey - prec') emin');
+ omega.
+- destruct (Z.ltb_spec (ex - prec) emin);
+ destruct (Z.ltb_spec (ey - prec) emin);
+ destruct (Z.ltb_spec (ex - ey - prec) emin);
+ destruct (Z.ltb_spec (ex - ey - prec') emin');
+ omega.
+Qed.
+
+Theorem round_round_div_FTZ :
+ forall choice1 choice2,
+ (exists n, (beta = 2 * n :> Z)%Z) ->
+ (emin' + prec' <= emin - 1)%Z ->
+ (2 * prec <= prec')%Z ->
+ forall x y,
+ y <> 0 ->
+ FTZ_format beta emin prec x -> FTZ_format beta emin prec y ->
+ round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec')
+ choice1 choice2 (x / y).
+Proof.
+intros choice1 choice2 Ebeta Hemin Hprec x y Nzy Fx Fy.
+apply round_round_div.
+- now apply FTZ_exp_valid.
+- now apply FTZ_exp_valid.
+- exact Ebeta.
+- now apply FTZ_round_round_div_hyp.
+- exact Nzy.
+- now apply generic_format_FTZ.
+- now apply generic_format_FTZ.
+Qed.
+
+End Double_round_div_FTZ.
+
+End Double_round_div.
+
+End Double_round.
diff --git a/flocq/Prop/Fprop_div_sqrt_error.v b/flocq/Prop/Fprop_div_sqrt_error.v
deleted file mode 100644
index 422b6b64..00000000
--- a/flocq/Prop/Fprop_div_sqrt_error.v
+++ /dev/null
@@ -1,300 +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.
-*)
-
-(** * Remainder of the division and square root are in the FLX format *)
-Require Import Fcore.
-Require Import Fcalc_ops.
-Require Import Fprop_relative.
-
-Section Fprop_divsqrt_error.
-
-Variable beta : radix.
-Notation bpow e := (bpow beta e).
-
-Variable prec : Z.
-
-Theorem generic_format_plus_prec:
- forall fexp, (forall e, (fexp e <= e - prec)%Z) ->
- forall x y (fx fy: float beta),
- (x = F2R fx)%R -> (y = F2R fy)%R -> (Rabs (x+y) < bpow (prec+Fexp fx))%R -> (Rabs (x+y) < bpow (prec+Fexp fy))%R
- -> generic_format beta fexp (x+y)%R.
-intros fexp Hfexp x y fx fy Hx Hy H1 H2.
-case (Req_dec (x+y) 0); intros H.
-rewrite H; apply generic_format_0.
-rewrite Hx, Hy, <- F2R_plus.
-apply generic_format_F2R.
-intros _.
-case_eq (Fplus beta fx fy).
-intros mz ez Hz.
-rewrite <- Hz.
-apply Zle_trans with (Zmin (Fexp fx) (Fexp fy)).
-rewrite F2R_plus, <- Hx, <- Hy.
-unfold canonic_exp.
-apply Zle_trans with (1:=Hfexp _).
-apply Zplus_le_reg_l with prec; ring_simplify.
-apply ln_beta_le_bpow with (1 := H).
-now apply Zmin_case.
-rewrite <- Fexp_Fplus, Hz.
-apply Zle_refl.
-Qed.
-
-Theorem ex_Fexp_canonic: forall fexp, forall x, generic_format beta fexp x
- -> exists fx:float beta, (x=F2R fx)%R /\ Fexp fx = canonic_exp beta fexp x.
-intros fexp x; unfold generic_format.
-exists (Float beta (Ztrunc (scaled_mantissa beta fexp x)) (canonic_exp beta fexp x)).
-split; auto.
-Qed.
-
-
-Context { prec_gt_0_ : Prec_gt_0 prec }.
-
-Notation format := (generic_format beta (FLX_exp prec)).
-Notation cexp := (canonic_exp beta (FLX_exp prec)).
-
-Variable choice : Z -> bool.
-
-
-(** Remainder of the division in FLX *)
-Theorem div_error_FLX :
- forall rnd { Zrnd : Valid_rnd rnd } x y,
- format x -> format y ->
- format (x - round beta (FLX_exp prec) rnd (x/y) * y)%R.
-Proof with auto with typeclass_instances.
-intros rnd Zrnd x y Hx Hy.
-destruct (Req_dec y 0) as [Zy|Zy].
-now rewrite Zy, Rmult_0_r, Rminus_0_r.
-destruct (Req_dec (round beta (FLX_exp prec) rnd (x/y)) 0) as [Hr|Hr].
-rewrite Hr; ring_simplify (x-0*y)%R; assumption.
-assert (Zx: x <> R0).
-contradict Hr.
-rewrite Hr.
-unfold Rdiv.
-now rewrite Rmult_0_l, round_0.
-destruct (ex_Fexp_canonic _ x Hx) as (fx,(Hx1,Hx2)).
-destruct (ex_Fexp_canonic _ y Hy) as (fy,(Hy1,Hy2)).
-destruct (ex_Fexp_canonic (FLX_exp prec) (round beta (FLX_exp prec) rnd (x / y))) as (fr,(Hr1,Hr2)).
-apply generic_format_round...
-unfold Rminus; apply generic_format_plus_prec with fx (Fopp beta (Fmult beta fr fy)); trivial.
-intros e; apply Zle_refl.
-now rewrite F2R_opp, F2R_mult, <- Hr1, <- Hy1.
-(* *)
-destruct (relative_error_FLX_ex beta prec (prec_gt_0 prec) rnd (x / y)%R) as (eps,(Heps1,Heps2)).
-rewrite Heps2.
-rewrite <- Rabs_Ropp.
-replace (-(x + - (x / y * (1 + eps) * y)))%R with (x * eps)%R by now field.
-rewrite Rabs_mult.
-apply Rlt_le_trans with (Rabs x * 1)%R.
-apply Rmult_lt_compat_l.
-now apply Rabs_pos_lt.
-apply Rlt_le_trans with (1 := Heps1).
-change 1%R with (bpow 0).
-apply bpow_le.
-generalize (prec_gt_0 prec).
-clear ; omega.
-rewrite Rmult_1_r.
-rewrite Hx2.
-unfold canonic_exp.
-destruct (ln_beta beta x) as (ex, Hex).
-simpl.
-specialize (Hex Zx).
-apply Rlt_le.
-apply Rlt_le_trans with (1 := proj2 Hex).
-apply bpow_le.
-unfold FLX_exp.
-ring_simplify.
-apply Zle_refl.
-(* *)
-replace (Fexp (Fopp beta (Fmult beta fr fy))) with (Fexp fr + Fexp fy)%Z.
-2: unfold Fopp, Fmult; destruct fr; destruct fy; now simpl.
-replace (x + - (round beta (FLX_exp prec) rnd (x / y) * y))%R with
- (y * (-(round beta (FLX_exp prec) rnd (x / y) - x/y)))%R.
-2: field; assumption.
-rewrite Rabs_mult.
-apply Rlt_le_trans with (Rabs y * bpow (Fexp fr))%R.
-apply Rmult_lt_compat_l.
-now apply Rabs_pos_lt.
-rewrite Rabs_Ropp.
-replace (bpow (Fexp fr)) with (ulp beta (FLX_exp prec) (F2R fr)).
-rewrite <- Hr1.
-apply error_lt_ulp_round...
-apply Rmult_integral_contrapositive_currified; try apply Rinv_neq_0_compat; assumption.
-rewrite ulp_neq_0.
-2: now rewrite <- Hr1.
-apply f_equal.
-now rewrite Hr2, <- Hr1.
-replace (prec+(Fexp fr+Fexp fy))%Z with ((prec+Fexp fy)+Fexp fr)%Z by ring.
-rewrite bpow_plus.
-apply Rmult_le_compat_r.
-apply bpow_ge_0.
-rewrite Hy2; unfold canonic_exp, FLX_exp.
-ring_simplify (prec + (ln_beta beta y - prec))%Z.
-destruct (ln_beta beta y); simpl.
-left; now apply a.
-Qed.
-
-(** Remainder of the square in FLX (with p>1) and rounding to nearest *)
-Variable Hp1 : Zlt 1 prec.
-
-Theorem sqrt_error_FLX_N :
- forall x, format x ->
- format (x - Rsqr (round beta (FLX_exp prec) (Znearest choice) (sqrt x)))%R.
-Proof with auto with typeclass_instances.
-intros x Hx.
-destruct (total_order_T x 0) as [[Hxz|Hxz]|Hxz].
-unfold sqrt.
-destruct (Rcase_abs x).
-rewrite round_0...
-unfold Rsqr.
-now rewrite Rmult_0_l, Rminus_0_r.
-elim (Rlt_irrefl 0).
-now apply Rgt_ge_trans with x.
-rewrite Hxz, sqrt_0, round_0...
-unfold Rsqr.
-rewrite Rmult_0_l, Rminus_0_r.
-apply generic_format_0.
-case (Req_dec (round beta (FLX_exp prec) (Znearest choice) (sqrt x)) 0); intros Hr.
-rewrite Hr; unfold Rsqr; ring_simplify (x-0*0)%R; assumption.
-destruct (ex_Fexp_canonic _ x Hx) as (fx,(Hx1,Hx2)).
-destruct (ex_Fexp_canonic (FLX_exp prec) (round beta (FLX_exp prec) (Znearest choice) (sqrt x))) as (fr,(Hr1,Hr2)).
-apply generic_format_round...
-unfold Rminus; apply generic_format_plus_prec with fx (Fopp beta (Fmult beta fr fr)); trivial.
-intros e; apply Zle_refl.
-unfold Rsqr; now rewrite F2R_opp,F2R_mult, <- Hr1.
-(* *)
-apply Rle_lt_trans with x.
-apply Rabs_minus_le.
-apply Rle_0_sqr.
-destruct (relative_error_N_FLX_ex beta prec (prec_gt_0 prec) choice (sqrt x)) as (eps,(Heps1,Heps2)).
-rewrite Heps2.
-rewrite Rsqr_mult, Rsqr_sqrt, Rmult_comm. 2: now apply Rlt_le.
-apply Rmult_le_compat_r.
-now apply Rlt_le.
-apply Rle_trans with (5²/4²)%R.
-rewrite <- Rsqr_div.
-apply Rsqr_le_abs_1.
-apply Rle_trans with (1 := Rabs_triang _ _).
-rewrite Rabs_R1.
-apply Rplus_le_reg_l with (-1)%R.
-replace (-1 + (1 + Rabs eps))%R with (Rabs eps) by ring.
-apply Rle_trans with (1 := Heps1).
-rewrite Rabs_pos_eq.
-apply Rmult_le_reg_l with 2%R.
-now apply (Z2R_lt 0 2).
-rewrite <- Rmult_assoc, Rinv_r, Rmult_1_l.
-apply Rle_trans with (bpow (-1)).
-apply bpow_le.
-omega.
-replace (2 * (-1 + 5 / 4))%R with (/2)%R by field.
-apply Rinv_le.
-now apply (Z2R_lt 0 2).
-apply (Z2R_le 2).
-unfold Zpower_pos. simpl.
-rewrite Zmult_1_r.
-apply Zle_bool_imp_le.
-apply beta.
-apply Rgt_not_eq.
-now apply (Z2R_lt 0 2).
-unfold Rdiv.
-apply Rmult_le_pos.
-now apply (Z2R_le 0 5).
-apply Rlt_le.
-apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 4).
-apply Rgt_not_eq.
-now apply (Z2R_lt 0 4).
-unfold Rsqr.
-replace (5 * 5 / (4 * 4))%R with (25 * /16)%R by field.
-apply Rmult_le_reg_r with 16%R.
-now apply (Z2R_lt 0 16).
-rewrite Rmult_assoc, Rinv_l, Rmult_1_r.
-now apply (Z2R_le 25 32).
-apply Rgt_not_eq.
-now apply (Z2R_lt 0 16).
-rewrite Hx2; unfold canonic_exp, FLX_exp.
-ring_simplify (prec + (ln_beta beta x - prec))%Z.
-destruct (ln_beta beta x); simpl.
-rewrite <- (Rabs_right x).
-apply a.
-now apply Rgt_not_eq.
-now apply Rgt_ge.
-(* *)
-replace (Fexp (Fopp beta (Fmult beta fr fr))) with (Fexp fr + Fexp fr)%Z.
-2: unfold Fopp, Fmult; destruct fr; now simpl.
-rewrite Hr1.
-replace (x + - Rsqr (F2R fr))%R with (-((F2R fr - sqrt x)*(F2R fr + sqrt x)))%R.
-2: rewrite <- (sqrt_sqrt x) at 3; auto.
-2: unfold Rsqr; ring.
-rewrite Rabs_Ropp, Rabs_mult.
-apply Rle_lt_trans with ((/2*bpow (Fexp fr))* Rabs (F2R fr + sqrt x))%R.
-apply Rmult_le_compat_r.
-apply Rabs_pos.
-apply Rle_trans with (/2*ulp beta (FLX_exp prec) (F2R fr))%R.
-rewrite <- Hr1.
-apply error_le_half_ulp_round...
-right; rewrite ulp_neq_0.
-2: now rewrite <- Hr1.
-apply f_equal.
-rewrite Hr2, <- Hr1; trivial.
-rewrite Rmult_assoc, Rmult_comm.
-replace (prec+(Fexp fr+Fexp fr))%Z with (Fexp fr + (prec+Fexp fr))%Z by ring.
-rewrite bpow_plus, Rmult_assoc.
-apply Rmult_lt_compat_l.
-apply bpow_gt_0.
-apply Rmult_lt_reg_l with (1 := Rlt_0_2).
-apply Rle_lt_trans with (Rabs (F2R fr + sqrt x)).
-right; field.
-apply Rle_lt_trans with (1:=Rabs_triang _ _).
-(* . *)
-assert (Rabs (F2R fr) < bpow (prec + Fexp fr))%R.
-rewrite Hr2; unfold canonic_exp; rewrite Hr1.
-unfold FLX_exp.
-ring_simplify (prec + (ln_beta beta (F2R fr) - prec))%Z.
-destruct (ln_beta beta (F2R fr)); simpl.
-apply a.
-rewrite <- Hr1; auto.
-(* . *)
-apply Rlt_le_trans with (bpow (prec + Fexp fr)+ Rabs (sqrt x))%R.
-now apply Rplus_lt_compat_r.
-(* . *)
-replace (2 * bpow (prec + Fexp fr))%R with (bpow (prec + Fexp fr) + bpow (prec + Fexp fr))%R by ring.
-apply Rplus_le_compat_l.
-assert (sqrt x <> 0)%R.
-apply Rgt_not_eq.
-now apply sqrt_lt_R0.
-destruct (ln_beta beta (sqrt x)) as (es,Es).
-specialize (Es H0).
-apply Rle_trans with (bpow es).
-now apply Rlt_le.
-apply bpow_le.
-case (Zle_or_lt es (prec + Fexp fr)) ; trivial.
-intros H1.
-absurd (Rabs (F2R fr) < bpow (es - 1))%R.
-apply Rle_not_lt.
-rewrite <- Hr1.
-apply abs_round_ge_generic...
-apply generic_format_bpow.
-unfold FLX_exp; omega.
-apply Es.
-apply Rlt_le_trans with (1:=H).
-apply bpow_le.
-omega.
-now apply Rlt_le.
-Qed.
-
-End Fprop_divsqrt_error.
diff --git a/flocq/Prop/Fprop_mult_error.v b/flocq/Prop/Mult_error.v
index 44448cd6..57a3856f 100644
--- a/flocq/Prop/Fprop_mult_error.v
+++ b/flocq/Prop/Mult_error.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,8 +18,7 @@ COPYING file for more details.
*)
(** * Error of the multiplication is in the FLX/FLT format *)
-Require Import Fcore.
-Require Import Fcalc_ops.
+Require Import Core Operations Plus_error.
Section Fprop_mult_error.
@@ -30,7 +29,7 @@ Variable prec : Z.
Context { prec_gt_0_ : Prec_gt_0 prec }.
Notation format := (generic_format beta (FLX_exp prec)).
-Notation cexp := (canonic_exp beta (FLX_exp prec)).
+Notation cexp := (cexp beta (FLX_exp prec)).
Variable rnd : R -> Z.
Context { valid_rnd : Valid_rnd rnd }.
@@ -41,9 +40,9 @@ Lemma mult_error_FLX_aux:
format x -> format y ->
(round beta (FLX_exp prec) rnd (x * y) - (x * y) <> 0)%R ->
exists f:float beta,
- (F2R f = round beta (FLX_exp prec) rnd (x * y) - (x * y))%R
- /\ (canonic_exp beta (FLX_exp prec) (F2R f) <= Fexp f)%Z
- /\ (Fexp f = cexp x + cexp y)%Z.
+ (F2R f = round beta (FLX_exp prec) rnd (x * y) - (x * y))%R
+ /\ (cexp (F2R f) <= Fexp f)%Z
+ /\ (Fexp f = cexp x + cexp y)%Z.
Proof with auto with typeclass_instances.
intros x y Hx Hy Hz.
set (f := (round beta (FLX_exp prec) rnd (x * y))).
@@ -52,26 +51,26 @@ contradict Hz.
rewrite Hxy0.
rewrite round_0...
ring.
-destruct (ln_beta beta (x * y)) as (exy, Hexy).
+destruct (mag beta (x * y)) as (exy, Hexy).
specialize (Hexy Hxy0).
-destruct (ln_beta beta (f - x * y)) as (er, Her).
+destruct (mag beta (f - x * y)) as (er, Her).
specialize (Her Hz).
-destruct (ln_beta beta x) as (ex, Hex).
+destruct (mag beta x) as (ex, Hex).
assert (Hx0: (x <> 0)%R).
contradict Hxy0.
now rewrite Hxy0, Rmult_0_l.
specialize (Hex Hx0).
-destruct (ln_beta beta y) as (ey, Hey).
+destruct (mag beta y) as (ey, Hey).
assert (Hy0: (y <> 0)%R).
contradict Hxy0.
now rewrite Hxy0, Rmult_0_r.
specialize (Hey Hy0).
(* *)
assert (Hc1: (cexp (x * y)%R - prec <= cexp x + cexp y)%Z).
-unfold canonic_exp, FLX_exp.
-rewrite ln_beta_unique with (1 := Hex).
-rewrite ln_beta_unique with (1 := Hey).
-rewrite ln_beta_unique with (1 := Hexy).
+unfold cexp, FLX_exp.
+rewrite mag_unique with (1 := Hex).
+rewrite mag_unique with (1 := Hey).
+rewrite mag_unique with (1 := Hexy).
cut (exy - 1 < ex + ey)%Z. omega.
apply (lt_bpow beta).
apply Rle_lt_trans with (1 := proj1 Hexy).
@@ -84,10 +83,10 @@ apply Hex.
apply Hey.
(* *)
assert (Hc2: (cexp x + cexp y <= cexp (x * y)%R)%Z).
-unfold canonic_exp, FLX_exp.
-rewrite ln_beta_unique with (1 := Hex).
-rewrite ln_beta_unique with (1 := Hey).
-rewrite ln_beta_unique with (1 := Hexy).
+unfold cexp, FLX_exp.
+rewrite mag_unique with (1 := Hex).
+rewrite mag_unique with (1 := Hey).
+rewrite mag_unique with (1 := Hexy).
cut ((ex - 1) + (ey - 1) < exy)%Z.
generalize (prec_gt_0 prec).
clear ; omega.
@@ -120,16 +119,16 @@ split;[assumption|split].
rewrite Hr.
simpl.
clear Hr.
-apply Zle_trans with (cexp (x * y)%R - prec)%Z.
-unfold canonic_exp, FLX_exp.
+apply Z.le_trans with (cexp (x * y)%R - prec)%Z.
+unfold cexp, FLX_exp.
apply Zplus_le_compat_r.
-rewrite ln_beta_unique with (1 := Hexy).
-apply ln_beta_le_bpow with (1 := Hz).
+rewrite mag_unique with (1 := Hexy).
+apply mag_le_bpow with (1 := Hz).
replace (bpow (exy - prec)) with (ulp beta (FLX_exp prec) (x * y)).
apply error_lt_ulp...
rewrite ulp_neq_0; trivial.
-unfold canonic_exp.
-now rewrite ln_beta_unique with (1 := Hexy).
+unfold cexp.
+now rewrite mag_unique with (1 := Hexy).
apply Hc1.
reflexivity.
Qed.
@@ -149,6 +148,24 @@ rewrite <- H1.
now apply generic_format_F2R.
Qed.
+Lemma mult_bpow_exact_FLX :
+ forall x e,
+ format x ->
+ format (x * bpow e)%R.
+Proof.
+intros x e Fx.
+destruct (Req_dec x 0) as [Zx|Nzx].
+{ rewrite Zx, Rmult_0_l; apply generic_format_0. }
+rewrite Fx.
+set (mx := Ztrunc _); set (ex := cexp _).
+pose (f := {| Fnum := mx; Fexp := ex + e |} : float beta).
+apply (generic_format_F2R' _ _ _ f).
+{ now unfold F2R; simpl; rewrite bpow_plus, Rmult_assoc. }
+intro Nzmx; unfold mx, ex; rewrite <- Fx.
+unfold f, ex; simpl; unfold cexp; rewrite (mag_mult_bpow _ _ _ Nzx).
+unfold FLX_exp; omega.
+Qed.
+
End Fprop_mult_error.
Section Fprop_mult_error_FLT.
@@ -160,7 +177,7 @@ Variable emin prec : Z.
Context { prec_gt_0_ : Prec_gt_0 prec }.
Notation format := (generic_format beta (FLT_exp emin prec)).
-Notation cexp := (canonic_exp beta (FLT_exp emin prec)).
+Notation cexp := (cexp beta (FLT_exp emin prec)).
Variable rnd : R -> Z.
Context { valid_rnd : Valid_rnd rnd }.
@@ -169,7 +186,7 @@ Context { valid_rnd : Valid_rnd rnd }.
Theorem mult_error_FLT :
forall x y,
format x -> format y ->
- (x*y = 0)%R \/ (bpow (emin + 2*prec - 1) <= Rabs (x * y))%R ->
+ (x * y <> 0 -> bpow (emin + 2*prec - 1) <= Rabs (x * y))%R ->
format (round beta (FLT_exp emin prec) rnd (x * y) - (x * y))%R.
Proof with auto with typeclass_instances.
intros x y Hx Hy Hxy.
@@ -177,12 +194,13 @@ set (f := (round beta (FLT_exp emin prec) rnd (x * y))).
destruct (Req_dec (f - x * y) 0) as [Hr0|Hr0].
rewrite Hr0.
apply generic_format_0.
-destruct Hxy as [Hxy|Hxy].
+destruct (Req_dec (x * y) 0) as [Hxy'|Hxy'].
unfold f.
-rewrite Hxy.
+rewrite Hxy'.
rewrite round_0...
ring_simplify (0 - 0)%R.
apply generic_format_0.
+specialize (Hxy Hxy').
destruct (mult_error_FLX_aux beta prec rnd x y) as ((m,e),(H1,(H2,H3))).
now apply generic_format_FLX_FLT with emin.
now apply generic_format_FLX_FLT with emin.
@@ -199,14 +217,14 @@ unfold f; rewrite <- H1.
apply generic_format_F2R.
intros _.
simpl in H2, H3.
-unfold canonic_exp, FLT_exp.
-case (Zmax_spec (ln_beta beta (F2R (Float beta m e)) - prec) emin);
+unfold cexp, FLT_exp.
+case (Zmax_spec (mag beta (F2R (Float beta m e)) - prec) emin);
intros (M1,M2); rewrite M2.
-apply Zle_trans with (2:=H2).
-unfold canonic_exp, FLX_exp.
-apply Zle_refl.
+apply Z.le_trans with (2:=H2).
+unfold cexp, FLX_exp.
+apply Z.le_refl.
rewrite H3.
-unfold canonic_exp, FLX_exp.
+unfold cexp, FLX_exp.
assert (Hxy0:(x*y <> 0)%R).
contradict Hr0.
unfold f.
@@ -219,9 +237,9 @@ now rewrite Hxy0, Rmult_0_l.
assert (Hy0: (y <> 0)%R).
contradict Hxy0.
now rewrite Hxy0, Rmult_0_r.
-destruct (ln_beta beta x) as (ex,Ex) ; simpl.
+destruct (mag beta x) as (ex,Ex) ; simpl.
specialize (Ex Hx0).
-destruct (ln_beta beta y) as (ey,Ey) ; simpl.
+destruct (mag beta y) as (ey,Ey) ; simpl.
specialize (Ey Hy0).
assert (emin + 2 * prec -1 < ex + ey)%Z.
2: omega.
@@ -233,4 +251,85 @@ apply Ex.
apply Ey.
Qed.
+Lemma F2R_ge: forall (y:float beta),
+ (F2R y <> 0)%R -> (bpow (Fexp y) <= Rabs (F2R y))%R.
+Proof.
+intros (ny,ey).
+rewrite <- F2R_Zabs; unfold F2R; simpl.
+case (Zle_lt_or_eq 0 (Z.abs ny)).
+apply Z.abs_nonneg.
+intros Hy _.
+rewrite <- (Rmult_1_l (bpow _)) at 1.
+apply Rmult_le_compat_r.
+apply bpow_ge_0.
+apply IZR_le; omega.
+intros H1 H2; contradict H2.
+replace ny with 0%Z.
+simpl; ring.
+now apply sym_eq, Z.abs_0_iff, sym_eq.
+Qed.
+
+Theorem mult_error_FLT_ge_bpow :
+ forall x y e,
+ format x -> format y ->
+ (bpow (e+2*prec-1) <= Rabs (x * y))%R ->
+ (round beta (FLT_exp emin prec) rnd (x * y) - (x * y) <> 0)%R ->
+ (bpow e <= Rabs (round beta (FLT_exp emin prec) rnd (x * y) - (x * y)))%R.
+Proof with auto with typeclass_instances.
+intros x y e.
+set (f := (round beta (FLT_exp emin prec) rnd (x * y))).
+intros Fx Fy H1.
+unfold f; rewrite Fx, Fy, <- F2R_mult.
+simpl Fmult.
+destruct (round_repr_same_exp beta (FLT_exp emin prec)
+ rnd (Ztrunc (scaled_mantissa beta (FLT_exp emin prec) x) *
+ Ztrunc (scaled_mantissa beta (FLT_exp emin prec) y))
+ (cexp x + cexp y)) as (n,Hn).
+rewrite Hn; clear Hn.
+rewrite <- F2R_minus, Fminus_same_exp.
+intros K.
+eapply Rle_trans with (2:=F2R_ge _ K).
+simpl (Fexp _).
+apply bpow_le.
+unfold cexp, FLT_exp.
+destruct (mag beta x) as (ex,Hx).
+destruct (mag beta y) as (ey,Hy).
+simpl; apply Z.le_trans with ((ex-prec)+(ey-prec))%Z.
+2: apply Zplus_le_compat; apply Z.le_max_l.
+assert (e + 2*prec -1< ex+ey)%Z;[idtac|omega].
+apply lt_bpow with beta.
+apply Rle_lt_trans with (1:=H1).
+rewrite Rabs_mult, bpow_plus.
+apply Rmult_lt_compat.
+apply Rabs_pos.
+apply Rabs_pos.
+apply Hx.
+intros K'; contradict H1; apply Rlt_not_le.
+rewrite K', Rmult_0_l, Rabs_R0; apply bpow_gt_0.
+apply Hy.
+intros K'; contradict H1; apply Rlt_not_le.
+rewrite K', Rmult_0_r, Rabs_R0; apply bpow_gt_0.
+Qed.
+
+Lemma mult_bpow_exact_FLT :
+ forall x e,
+ format x ->
+ (emin + prec - mag beta x <= e)%Z ->
+ format (x * bpow e)%R.
+Proof.
+intros x e Fx He.
+destruct (Req_dec x 0) as [Zx|Nzx].
+{ rewrite Zx, Rmult_0_l; apply generic_format_0. }
+rewrite Fx.
+set (mx := Ztrunc _); set (ex := cexp _).
+pose (f := {| Fnum := mx; Fexp := ex + e |} : float beta).
+apply (generic_format_F2R' _ _ _ f).
+{ now unfold F2R; simpl; rewrite bpow_plus, Rmult_assoc. }
+intro Nzmx; unfold mx, ex; rewrite <- Fx.
+unfold f, ex; simpl; unfold cexp; rewrite (mag_mult_bpow _ _ _ Nzx).
+unfold FLT_exp; rewrite Z.max_l; [|omega]; rewrite <- Z.add_max_distr_r.
+set (n := (_ - _ + _)%Z); apply (Z.le_trans _ n); [unfold n; omega|].
+apply Z.le_max_l.
+Qed.
+
End Fprop_mult_error_FLT.
diff --git a/flocq/Prop/Fprop_plus_error.v b/flocq/Prop/Plus_error.v
index 9bb5aee8..42f80093 100644
--- a/flocq/Prop/Fprop_plus_error.v
+++ b/flocq/Prop/Plus_error.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
@@ -20,15 +20,9 @@ COPYING file for more details.
(** * Error of the rounded-to-nearest addition is representable. *)
Require Import Psatz.
-Require Import Fcore_Raux.
-Require Import Fcore_defs.
-Require Import Fcore_float_prop.
-Require Import Fcore_generic_fmt.
-Require Import Fcore_FIX.
-Require Import Fcore_FLX.
-Require Import Fcore_FLT.
-Require Import Fcore_ulp.
-Require Import Fcalc_ops.
+Require Import Raux Defs Float_prop Generic_fmt.
+Require Import FIX FLX FLT Ulp Operations.
+Require Import Relative.
Section Fprop_plus_error.
@@ -44,31 +38,31 @@ Section round_repr_same_exp.
Variable rnd : R -> Z.
Context { valid_rnd : Valid_rnd rnd }.
-Theorem round_repr_same_exp :
+Lemma round_repr_same_exp :
forall m e,
exists m',
round beta fexp rnd (F2R (Float beta m e)) = F2R (Float beta m' e).
Proof with auto with typeclass_instances.
intros m e.
-set (e' := canonic_exp beta fexp (F2R (Float beta m e))).
+set (e' := cexp beta fexp (F2R (Float beta m e))).
unfold round, scaled_mantissa. fold e'.
destruct (Zle_or_lt e' e) as [He|He].
exists m.
unfold F2R at 2. simpl.
rewrite Rmult_assoc, <- bpow_plus.
-rewrite <- Z2R_Zpower. 2: omega.
-rewrite <- Z2R_mult, Zrnd_Z2R...
+rewrite <- IZR_Zpower. 2: omega.
+rewrite <- mult_IZR, Zrnd_IZR...
unfold F2R. simpl.
-rewrite Z2R_mult.
+rewrite mult_IZR.
rewrite Rmult_assoc.
-rewrite Z2R_Zpower. 2: omega.
+rewrite IZR_Zpower. 2: omega.
rewrite <- bpow_plus.
-apply (f_equal (fun v => Z2R m * bpow v)%R).
+apply (f_equal (fun v => IZR m * bpow v)%R).
ring.
-exists ((rnd (Z2R m * bpow (e - e'))) * Zpower beta (e' - e))%Z.
+exists ((rnd (IZR m * bpow (e - e'))) * Zpower beta (e' - e))%Z.
unfold F2R. simpl.
-rewrite Z2R_mult.
-rewrite Z2R_Zpower. 2: omega.
+rewrite mult_IZR.
+rewrite IZR_Zpower. 2: omega.
rewrite 2!Rmult_assoc.
rewrite <- 2!bpow_plus.
apply (f_equal (fun v => _ * bpow v)%R).
@@ -84,13 +78,13 @@ Variable choice : Z -> bool.
Lemma plus_error_aux :
forall x y,
- (canonic_exp beta fexp x <= canonic_exp beta fexp y)%Z ->
+ (cexp beta fexp x <= cexp beta fexp y)%Z ->
format x -> format y ->
format (round beta fexp (Znearest choice) (x + y) - (x + y))%R.
Proof.
intros x y.
-set (ex := canonic_exp beta fexp x).
-set (ey := canonic_exp beta fexp y).
+set (ex := cexp beta fexp x).
+set (ey := cexp beta fexp y).
intros He Hx Hy.
destruct (Req_dec (round beta fexp (Znearest choice) (x + y) - (x + y)) R0) as [H0|H0].
rewrite H0.
@@ -116,7 +110,7 @@ apply generic_format_F2R.
intros _.
apply monotone_exp.
rewrite <- H, <- Hxy', <- Hxy.
-apply ln_beta_le_abs.
+apply mag_le_abs.
exact H0.
pattern x at 3 ; replace x with (-(y - (x + y)))%R by ring.
rewrite Rabs_Ropp.
@@ -130,7 +124,7 @@ Theorem plus_error :
format (round beta fexp (Znearest choice) (x + y) - (x + y))%R.
Proof.
intros x y Hx Hy.
-destruct (Zle_or_lt (canonic_exp beta fexp x) (canonic_exp beta fexp y)).
+destruct (Zle_or_lt (cexp beta fexp x) (cexp beta fexp y)).
now apply plus_error_aux.
rewrite Rplus_comm.
apply plus_error_aux ; try easy.
@@ -154,20 +148,17 @@ Section round_plus_eq_zero_aux.
Variable rnd : R -> Z.
Context { valid_rnd : Valid_rnd rnd }.
-Lemma round_plus_eq_zero_aux :
+Lemma round_plus_neq_0_aux :
forall x y,
- (canonic_exp beta fexp x <= canonic_exp beta fexp y)%Z ->
+ (cexp beta fexp x <= cexp beta fexp y)%Z ->
format x -> format y ->
- (0 <= x + y)%R ->
- round beta fexp rnd (x + y) = 0%R ->
- (x + y = 0)%R.
+ (0 < x + y)%R ->
+ round beta fexp rnd (x + y) <> 0%R.
Proof with auto with typeclass_instances.
-intros x y He Hx Hy Hp Hxy.
-destruct (Req_dec (x + y) 0) as [H0|H0].
-exact H0.
-destruct (ln_beta beta (x + y)) as (exy, Hexy).
+intros x y He Hx Hy Hxy.
+destruct (mag beta (x + y)) as (exy, Hexy).
simpl.
-specialize (Hexy H0).
+specialize (Hexy (Rgt_not_eq _ _ Hxy)).
destruct (Zle_or_lt exy (fexp exy)) as [He'|He'].
(* . *)
assert (H: (x + y)%R = F2R (Float beta (Ztrunc (x * bpow (- fexp exy)) +
@@ -175,19 +166,21 @@ assert (H: (x + y)%R = F2R (Float beta (Ztrunc (x * bpow (- fexp exy)) +
rewrite (subnormal_exponent beta fexp exy x He' Hx) at 1.
rewrite (subnormal_exponent beta fexp exy y He' Hy) at 1.
now rewrite <- F2R_plus, Fplus_same_exp.
-rewrite H in Hxy.
-rewrite round_generic in Hxy...
-now rewrite <- H in Hxy.
+rewrite H.
+rewrite round_generic...
+rewrite <- H.
+now apply Rgt_not_eq.
apply generic_format_F2R.
intros _.
rewrite <- H.
-unfold canonic_exp.
-rewrite ln_beta_unique with (1 := Hexy).
-apply Zle_refl.
+unfold cexp.
+rewrite mag_unique with (1 := Hexy).
+apply Z.le_refl.
(* . *)
+intros H.
elim Rle_not_lt with (1 := round_le beta _ rnd _ _ (proj1 Hexy)).
-rewrite (Rabs_pos_eq _ Hp).
-rewrite Hxy.
+rewrite (Rabs_pos_eq _ (Rlt_le _ _ Hxy)).
+rewrite H.
rewrite round_generic...
apply bpow_gt_0.
apply generic_format_bpow.
@@ -201,40 +194,46 @@ Variable rnd : R -> Z.
Context { valid_rnd : Valid_rnd rnd }.
(** rnd(x+y)=0 -> x+y = 0 provided this is not a FTZ format *)
-Theorem round_plus_eq_zero :
+Theorem round_plus_neq_0 :
forall x y,
format x -> format y ->
- round beta fexp rnd (x + y) = 0%R ->
- (x + y = 0)%R.
+ (x + y <> 0)%R ->
+ round beta fexp rnd (x + y) <> 0%R.
Proof with auto with typeclass_instances.
-intros x y Hx Hy.
+intros x y Hx Hy Hxy.
destruct (Rle_or_lt 0 (x + y)) as [H1|H1].
(* . *)
-revert H1.
-destruct (Zle_or_lt (canonic_exp beta fexp x) (canonic_exp beta fexp y)) as [H2|H2].
-now apply round_plus_eq_zero_aux.
+destruct (Zle_or_lt (cexp beta fexp x) (cexp beta fexp y)) as [H2|H2].
+apply round_plus_neq_0_aux...
+lra.
rewrite Rplus_comm.
-apply round_plus_eq_zero_aux ; try easy.
+apply round_plus_neq_0_aux ; try easy.
now apply Zlt_le_weak.
+lra.
(* . *)
-revert H1.
-rewrite <- (Ropp_involutive (x + y)), Ropp_plus_distr, <- Ropp_0.
-intros H1.
+rewrite <- (Ropp_involutive (x + y)), Ropp_plus_distr.
rewrite round_opp.
-intros Hxy.
-apply f_equal.
-cut (round beta fexp (Zrnd_opp rnd) (- x + - y) = 0)%R.
-cut (0 <= -x + -y)%R.
-destruct (Zle_or_lt (canonic_exp beta fexp (-x)) (canonic_exp beta fexp (-y))) as [H2|H2].
-apply round_plus_eq_zero_aux ; try apply generic_format_opp...
+apply Ropp_neq_0_compat.
+destruct (Zle_or_lt (cexp beta fexp (-x)) (cexp beta fexp (-y))) as [H2|H2].
+apply round_plus_neq_0_aux; try apply generic_format_opp...
+lra.
rewrite Rplus_comm.
-apply round_plus_eq_zero_aux ; try apply generic_format_opp...
+apply round_plus_neq_0_aux; try apply generic_format_opp...
now apply Zlt_le_weak.
-apply Rlt_le.
-now apply Ropp_lt_cancel.
-rewrite <- (Ropp_involutive (round _ _ _ _)).
-rewrite Hxy.
-apply Ropp_involutive.
+lra.
+Qed.
+
+Theorem round_plus_eq_0 :
+ forall x y,
+ format x -> format y ->
+ round beta fexp rnd (x + y) = 0%R ->
+ (x + y = 0)%R.
+Proof with auto with typeclass_instances.
+intros x y Fx Fy H.
+destruct (Req_dec (x + y) 0) as [H'|H'].
+exact H'.
+contradict H.
+now apply round_plus_neq_0.
Qed.
End Fprop_plus_zero.
@@ -258,14 +257,48 @@ apply generic_format_FLT_FIX...
rewrite Zplus_comm; assumption.
apply generic_format_FIX_FLT, FIX_format_generic in Fx.
apply generic_format_FIX_FLT, FIX_format_generic in Fy.
-destruct Fx as (nx,(H1x,H2x)).
-destruct Fy as (ny,(H1y,H2y)).
+destruct Fx as [nx H1x H2x].
+destruct Fy as [ny H1y H2y].
apply generic_format_FIX.
exists (Float beta (Fnum nx+Fnum ny)%Z emin).
-split;[idtac|reflexivity].
rewrite H1x,H1y; unfold F2R; simpl.
rewrite H2x, H2y.
-rewrite Z2R_plus; ring.
+rewrite plus_IZR; ring.
+easy.
+Qed.
+
+Variable choice : Z -> bool.
+
+Lemma FLT_plus_error_N_ex : forall x y,
+ generic_format beta (FLT_exp emin prec) x ->
+ generic_format beta (FLT_exp emin prec) y ->
+ exists eps,
+ (Rabs eps <= u_ro beta prec / (1 + u_ro beta prec))%R /\
+ round beta (FLT_exp emin prec) (Znearest choice) (x + y)
+ = ((x + y) * (1 + eps))%R.
+Proof.
+intros x y Fx Fy.
+assert (Pb := u_rod1pu_ro_pos beta prec).
+destruct (Rle_or_lt (bpow (emin + prec - 1)) (Rabs (x + y))) as [M|M].
+{ destruct (relative_error_N_FLX'_ex beta prec prec_gt_0_ choice (x + y))
+ as (d, (Bd, Hd)).
+ now exists d; split; [exact Bd|]; rewrite <- Hd; apply round_FLT_FLX. }
+exists 0%R; rewrite Rabs_R0; split; [exact Pb|]; rewrite Rplus_0_r, Rmult_1_r.
+apply round_generic; [apply valid_rnd_N|].
+apply FLT_format_plus_small; [exact Fx|exact Fy|].
+apply Rlt_le, (Rlt_le_trans _ _ _ M), bpow_le; lia.
+Qed.
+
+Lemma FLT_plus_error_N_round_ex : forall x y,
+ generic_format beta (FLT_exp emin prec) x ->
+ generic_format beta (FLT_exp emin prec) y ->
+ exists eps,
+ (Rabs eps <= u_ro beta prec)%R /\
+ (x + y
+ = round beta (FLT_exp emin prec) (Znearest choice) (x + y) * (1 + eps))%R.
+Proof.
+intros x y Fx Fy.
+now apply relative_error_N_round_ex_derive, FLT_plus_error_N_ex.
Qed.
End Fprop_plus_FLT.
@@ -282,62 +315,58 @@ Variable rnd : R -> Z.
Context { valid_rnd : Valid_rnd rnd }.
Notation format := (generic_format beta fexp).
-Notation cexp := (canonic_exp beta fexp).
+Notation cexp := (cexp beta fexp).
Lemma ex_shift :
forall x e, format x -> (e <= cexp x)%Z ->
- exists m, (x = Z2R m * bpow e)%R.
+ exists m, (x = IZR m * bpow e)%R.
Proof with auto with typeclass_instances.
intros x e Fx He.
exists (Ztrunc (scaled_mantissa beta fexp x)*Zpower beta (cexp x -e))%Z.
rewrite Fx at 1; unfold F2R; simpl.
-rewrite Z2R_mult, Rmult_assoc.
+rewrite mult_IZR, Rmult_assoc.
f_equal.
-rewrite Z2R_Zpower.
+rewrite IZR_Zpower.
2: omega.
rewrite <- bpow_plus; f_equal; ring.
Qed.
-Lemma ln_beta_minus1 :
+Lemma mag_minus1 :
forall z, z <> 0%R ->
- (ln_beta beta z - 1)%Z = ln_beta beta (z / Z2R beta).
+ (mag beta z - 1)%Z = mag beta (z / IZR beta).
Proof.
intros z Hz.
unfold Zminus.
-rewrite <- ln_beta_mult_bpow with (1 := Hz).
+rewrite <- mag_mult_bpow by easy.
now rewrite bpow_opp, bpow_1.
Qed.
-Theorem round_plus_mult_ulp :
+Theorem round_plus_F2R :
forall x y, format x -> format y -> (x <> 0)%R ->
- exists m, (round beta fexp rnd (x+y) = Z2R m * ulp beta fexp (x/Z2R beta))%R.
+ exists m,
+ round beta fexp rnd (x+y) = F2R (Float beta m (cexp (x / IZR beta))).
Proof with auto with typeclass_instances.
intros x y Fx Fy Zx.
-case (Zle_or_lt (ln_beta beta (x/Z2R beta)) (ln_beta beta y)); intros H1.
-pose (e:=cexp (x / Z2R beta)).
+case (Zle_or_lt (mag beta (x/IZR beta)) (mag beta y)); intros H1.
+pose (e:=cexp (x / IZR beta)).
destruct (ex_shift x e) as (nx, Hnx); try exact Fx.
apply monotone_exp.
-rewrite <- (ln_beta_minus1 x Zx); omega.
+rewrite <- (mag_minus1 x Zx); omega.
destruct (ex_shift y e) as (ny, Hny); try assumption.
apply monotone_exp...
destruct (round_repr_same_exp beta fexp rnd (nx+ny) e) as (n,Hn).
exists n.
-apply trans_eq with (F2R (Float beta n e)).
+fold e.
rewrite <- Hn; f_equal.
-rewrite Hnx, Hny; unfold F2R; simpl; rewrite Z2R_plus; ring.
+rewrite Hnx, Hny; unfold F2R; simpl; rewrite plus_IZR; ring.
unfold F2R; simpl.
-rewrite ulp_neq_0; try easy.
-apply Rmult_integral_contrapositive_currified; try assumption.
-apply Rinv_neq_0_compat.
-apply Rgt_not_eq.
-apply radix_pos.
(* *)
-destruct (ex_shift (round beta fexp rnd (x + y)) (cexp (x/Z2R beta))) as (n,Hn).
+destruct (ex_shift (round beta fexp rnd (x + y)) (cexp (x/IZR beta))) as (n,Hn).
apply generic_format_round...
-apply Zle_trans with (cexp (x+y)).
+apply Z.le_trans with (cexp (x+y)).
apply monotone_exp.
-rewrite <- ln_beta_minus1 by easy.
-rewrite <- (ln_beta_abs beta (x+y)).
+rewrite <- mag_minus1 by easy.
+rewrite <- (mag_abs beta (x+y)).
(* . *)
assert (U: (Rabs (x+y) = Rabs x + Rabs y)%R \/ (y <> 0 /\ Rabs (x+y) = Rabs x - Rabs y)%R).
assert (V: forall x y, (Rabs y <= Rabs x)%R ->
@@ -374,94 +403,89 @@ rewrite Rabs_left1.
ring.
lra.
apply V; left.
-apply ln_beta_lt_pos with beta.
+apply lt_mag with beta.
now apply Rabs_pos_lt.
-rewrite <- ln_beta_minus1 in H1; try assumption.
-rewrite 2!ln_beta_abs; omega.
+rewrite <- mag_minus1 in H1; try assumption.
+rewrite 2!mag_abs; omega.
(* . *)
destruct U as [U|U].
-rewrite U; apply Zle_trans with (ln_beta beta x).
+rewrite U; apply Z.le_trans with (mag beta x).
omega.
-rewrite <- ln_beta_abs.
-apply ln_beta_le.
+rewrite <- mag_abs.
+apply mag_le.
now apply Rabs_pos_lt.
apply Rplus_le_reg_l with (-Rabs x)%R; ring_simplify.
apply Rabs_pos.
destruct U as (U',U); rewrite U.
-rewrite <- ln_beta_abs.
-apply ln_beta_minus_lb.
+rewrite <- mag_abs.
+apply mag_minus_lb.
now apply Rabs_pos_lt.
now apply Rabs_pos_lt.
-rewrite 2!ln_beta_abs.
-assert (ln_beta beta y < ln_beta beta x - 1)%Z.
-now rewrite (ln_beta_minus1 x Zx).
+rewrite 2!mag_abs.
+assert (mag beta y < mag beta x - 1)%Z.
+now rewrite (mag_minus1 x Zx).
omega.
-apply canonic_exp_round_ge...
-intros K.
-apply round_plus_eq_zero in K...
+apply cexp_round_ge...
+apply round_plus_neq_0...
contradict H1; apply Zle_not_lt.
-rewrite <- (ln_beta_minus1 x Zx).
+rewrite <- (mag_minus1 x Zx).
replace y with (-x)%R.
-rewrite ln_beta_opp; omega.
+rewrite mag_opp; omega.
lra.
-exists n.
-rewrite ulp_neq_0.
-assumption.
-apply Rmult_integral_contrapositive_currified; try assumption.
-apply Rinv_neq_0_compat.
-apply Rgt_not_eq.
-apply radix_pos.
+now exists n.
Qed.
Context {exp_not_FTZ : Exp_not_FTZ fexp}.
Theorem round_plus_ge_ulp :
forall x y, format x -> format y ->
- round beta fexp rnd (x+y) = 0%R \/
- (ulp beta fexp (x/Z2R beta) <= Rabs (round beta fexp rnd (x+y)))%R.
+ round beta fexp rnd (x+y) <> 0%R ->
+ (ulp beta fexp (x/IZR beta) <= Rabs (round beta fexp rnd (x+y)))%R.
Proof with auto with typeclass_instances.
-intros x y Fx Fy.
+intros x y Fx Fy KK.
case (Req_dec x 0); intros Zx.
(* *)
rewrite Zx, Rplus_0_l.
rewrite round_generic...
unfold Rdiv; rewrite Rmult_0_l.
-rewrite Fy at 2.
+rewrite Fy.
unfold F2R; simpl; rewrite Rabs_mult.
rewrite (Rabs_pos_eq (bpow _)) by apply bpow_ge_0.
case (Z.eq_dec (Ztrunc (scaled_mantissa beta fexp y)) 0); intros Hm.
-left.
-rewrite Fy, Hm; unfold F2R; simpl; ring.
-right.
+contradict KK.
+rewrite Zx, Fy, Hm; unfold F2R; simpl.
+rewrite Rplus_0_l, Rmult_0_l.
+apply round_0...
apply Rle_trans with (1*bpow (cexp y))%R.
rewrite Rmult_1_l.
rewrite <- ulp_neq_0.
apply ulp_ge_ulp_0...
intros K; apply Hm.
rewrite K, scaled_mantissa_0.
-apply (Ztrunc_Z2R 0).
+apply Ztrunc_IZR.
apply Rmult_le_compat_r.
apply bpow_ge_0.
-rewrite <- Z2R_abs.
-apply (Z2R_le 1).
+rewrite <- abs_IZR.
+apply IZR_le.
apply (Zlt_le_succ 0).
now apply Z.abs_pos.
(* *)
-destruct (round_plus_mult_ulp x y Fx Fy Zx) as (m,Hm).
+destruct (round_plus_F2R x y Fx Fy Zx) as (m,Hm).
case (Z.eq_dec m 0); intros Zm.
-left.
-rewrite Hm, Zm; simpl; ring.
-right.
-rewrite Hm, Rabs_mult.
-rewrite (Rabs_pos_eq (ulp _ _ _)) by apply ulp_ge_0.
-apply Rle_trans with (1*ulp beta fexp (x/Z2R beta))%R.
-right; ring.
+contradict KK.
+rewrite Hm, Zm.
+apply F2R_0.
+rewrite Hm, <- F2R_Zabs.
+rewrite ulp_neq_0.
+rewrite <- (Rmult_1_l (bpow _)).
apply Rmult_le_compat_r.
-apply ulp_ge_0.
-rewrite <- Z2R_abs.
-apply (Z2R_le 1).
+apply bpow_ge_0.
+apply IZR_le.
apply (Zlt_le_succ 0).
now apply Z.abs_pos.
+apply Rmult_integral_contrapositive_currified with (1 := Zx).
+apply Rinv_neq_0_compat.
+apply Rgt_not_eq, radix_pos.
Qed.
End Fprop_plus_mult_ulp.
@@ -476,27 +500,27 @@ Context { valid_rnd : Valid_rnd rnd }.
Variable emin prec : Z.
Context { prec_gt_0_ : Prec_gt_0 prec }.
-Theorem round_plus_ge_ulp_FLT : forall x y e,
+Theorem round_FLT_plus_ge :
+ forall x y e,
generic_format beta (FLT_exp emin prec) x -> generic_format beta (FLT_exp emin prec) y ->
- (bpow e <= Rabs x)%R ->
- round beta (FLT_exp emin prec) rnd (x+y) = 0%R \/
- (bpow (e - prec) <= Rabs (round beta (FLT_exp emin prec) rnd (x+y)))%R.
+ (bpow (e + prec) <= Rabs x)%R ->
+ round beta (FLT_exp emin prec) rnd (x + y) <> 0%R ->
+ (bpow e <= Rabs (round beta (FLT_exp emin prec) rnd (x + y)))%R.
Proof with auto with typeclass_instances.
-intros x y e Fx Fy He.
+intros x y e Fx Fy He KK.
assert (Zx: x <> 0%R).
contradict He.
apply Rlt_not_le; rewrite He, Rabs_R0.
apply bpow_gt_0.
-case round_plus_ge_ulp with beta (FLT_exp emin prec) rnd x y...
-intros H; right.
-apply Rle_trans with (2:=H).
+apply Rle_trans with (ulp beta (FLT_exp emin prec) (x/IZR beta)).
+2: apply round_plus_ge_ulp...
rewrite ulp_neq_0.
-unfold canonic_exp.
-rewrite <- ln_beta_minus1 by easy.
+unfold cexp.
+rewrite <- mag_minus1; try assumption.
unfold FLT_exp; apply bpow_le.
-apply Zle_trans with (2:=Z.le_max_l _ _).
-destruct (ln_beta beta x) as (n,Hn); simpl.
-assert (e < n)%Z; try omega.
+apply Z.le_trans with (2:=Z.le_max_l _ _).
+destruct (mag beta x) as (n,Hn); simpl.
+assert (e + prec < n)%Z; try omega.
apply lt_bpow with beta.
apply Rle_lt_trans with (1:=He).
now apply Hn.
@@ -506,26 +530,45 @@ apply Rgt_not_eq.
apply radix_pos.
Qed.
-Theorem round_plus_ge_ulp_FLX : forall x y e,
+Lemma round_FLT_plus_ge' :
+ forall x y e,
+ generic_format beta (FLT_exp emin prec) x -> generic_format beta (FLT_exp emin prec) y ->
+ (x <> 0%R -> (bpow (e+prec) <= Rabs x)%R) ->
+ (x = 0%R -> y <> 0%R -> (bpow e <= Rabs y)%R) ->
+ round beta (FLT_exp emin prec) rnd (x+y) <> 0%R ->
+ (bpow e <= Rabs (round beta (FLT_exp emin prec) rnd (x+y)))%R.
+Proof with auto with typeclass_instances.
+intros x y e Fx Fy H1 H2 H3.
+case (Req_dec x 0); intros H4.
+case (Req_dec y 0); intros H5.
+contradict H3.
+rewrite H4, H5, Rplus_0_l; apply round_0...
+rewrite H4, Rplus_0_l.
+rewrite round_generic...
+apply round_FLT_plus_ge; try easy.
+now apply H1.
+Qed.
+
+Theorem round_FLX_plus_ge :
+ forall x y e,
generic_format beta (FLX_exp prec) x -> generic_format beta (FLX_exp prec) y ->
- (bpow e <= Rabs x)%R ->
- round beta (FLX_exp prec) rnd (x+y) = 0%R \/
- (bpow (e - prec) <= Rabs (round beta (FLX_exp prec) rnd (x+y)))%R.
+ (bpow (e+prec) <= Rabs x)%R ->
+ (round beta (FLX_exp prec) rnd (x+y) <> 0)%R ->
+ (bpow e <= Rabs (round beta (FLX_exp prec) rnd (x+y)))%R.
Proof with auto with typeclass_instances.
-intros x y e Fx Fy He.
+intros x y e Fx Fy He KK.
assert (Zx: x <> 0%R).
contradict He.
apply Rlt_not_le; rewrite He, Rabs_R0.
apply bpow_gt_0.
-case round_plus_ge_ulp with beta (FLX_exp prec) rnd x y...
-intros H; right.
-apply Rle_trans with (2:=H).
+apply Rle_trans with (ulp beta (FLX_exp prec) (x/IZR beta)).
+2: apply round_plus_ge_ulp...
rewrite ulp_neq_0.
-unfold canonic_exp.
-rewrite <- ln_beta_minus1 by easy.
+unfold cexp.
+rewrite <- mag_minus1 by easy.
unfold FLX_exp; apply bpow_le.
-destruct (ln_beta beta x) as (n,Hn); simpl.
-assert (e < n)%Z; try omega.
+destruct (mag beta x) as (n,Hn); simpl.
+assert (e + prec < n)%Z; try omega.
apply lt_bpow with beta.
apply Rle_lt_trans with (1:=He).
now apply Hn.
@@ -536,3 +579,28 @@ apply radix_pos.
Qed.
End Fprop_plus_ge_ulp.
+
+Section Fprop_plus_le_ops.
+
+Variable beta : radix.
+Variable fexp : Z -> Z.
+Context { valid_exp : Valid_exp fexp }.
+Variable choice : Z -> bool.
+
+Lemma plus_error_le_l :
+ forall x y,
+ generic_format beta fexp x -> generic_format beta fexp y ->
+ (Rabs (round beta fexp (Znearest choice) (x + y) - (x + y)) <= Rabs x)%R.
+Proof.
+intros x y Fx Fy.
+apply (Rle_trans _ (Rabs (y - (x + y)))); [now apply round_N_pt|].
+rewrite Rabs_minus_sym; right; f_equal; ring.
+Qed.
+
+Lemma plus_error_le_r :
+ forall x y,
+ generic_format beta fexp x -> generic_format beta fexp y ->
+ (Rabs (round beta fexp (Znearest choice) (x + y) - (x + y)) <= Rabs y)%R.
+Proof. now intros x y Fx Fy; rewrite Rplus_comm; apply plus_error_le_l. Qed.
+
+End Fprop_plus_le_ops.
diff --git a/flocq/Prop/Fprop_relative.v b/flocq/Prop/Relative.v
index 276ccd3b..b936f2f7 100644
--- a/flocq/Prop/Fprop_relative.v
+++ b/flocq/Prop/Relative.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,7 +18,8 @@ COPYING file for more details.
*)
(** * Relative error of the roundings *)
-Require Import Fcore.
+Require Import Core.
+Require Import Psatz. (* for lra *)
Section Fprop_relative.
@@ -88,6 +89,32 @@ rewrite Rinv_l with (1 := Hx0).
now rewrite Rabs_R1, Rmult_1_r.
Qed.
+Lemma relative_error_le_conversion_inv :
+ forall x b,
+ (exists eps,
+ (Rabs eps <= b)%R /\ round beta fexp rnd x = (x * (1 + eps))%R) ->
+ (Rabs (round beta fexp rnd x - x) <= b * Rabs x)%R.
+Proof with auto with typeclass_instances.
+intros x b (eps, (Beps, Heps)).
+assert (Pb : (0 <= b)%R); [now revert Beps; apply Rle_trans, Rabs_pos|].
+rewrite Heps; replace (_ - _)%R with (eps * x)%R; [|ring].
+now rewrite Rabs_mult; apply Rmult_le_compat_r; [apply Rabs_pos|].
+Qed.
+
+Lemma relative_error_le_conversion_round_inv :
+ forall x b,
+ (exists eps,
+ (Rabs eps <= b)%R /\ x = (round beta fexp rnd x * (1 + eps))%R) ->
+ (Rabs (round beta fexp rnd x - x) <= b * Rabs (round beta fexp rnd x))%R.
+Proof with auto with typeclass_instances.
+intros x b.
+set (rx := round _ _ _ _).
+intros (eps, (Beps, Heps)).
+assert (Pb : (0 <= b)%R); [now revert Beps; apply Rle_trans, Rabs_pos|].
+rewrite Heps; replace (_ - _)%R with (- (eps * rx))%R; [|ring].
+now rewrite Rabs_Ropp, Rabs_mult; apply Rmult_le_compat_r; [apply Rabs_pos|].
+Qed.
+
End relative_error_conversion.
Variable emin p : Z.
@@ -108,8 +135,8 @@ apply Rlt_not_le, bpow_gt_0.
apply Rlt_le_trans with (ulp beta fexp x)%R.
now apply error_lt_ulp...
rewrite ulp_neq_0; trivial.
-unfold canonic_exp.
-destruct (ln_beta beta x) as (ex, He).
+unfold cexp.
+destruct (mag beta x) as (ex, He).
simpl.
specialize (He Hx').
apply Rle_trans with (bpow (-p + 1) * bpow (ex - 1))%R.
@@ -150,7 +177,7 @@ apply relative_error.
unfold x.
rewrite <- F2R_Zabs.
apply bpow_le_F2R.
-apply F2R_lt_reg with beta emin.
+apply lt_F2R with beta emin.
rewrite F2R_0, F2R_Zabs.
now apply Rabs_pos_lt.
Qed.
@@ -179,8 +206,8 @@ apply Rlt_not_le, bpow_gt_0.
apply Rlt_le_trans with (ulp beta fexp x)%R.
now apply error_lt_ulp.
rewrite ulp_neq_0; trivial.
-unfold canonic_exp.
-destruct (ln_beta beta x) as (ex, He).
+unfold cexp.
+destruct (mag beta x) as (ex, He).
simpl.
specialize (He Hx').
assert (He': (emin < ex)%Z).
@@ -218,7 +245,7 @@ exact Hp.
unfold x.
rewrite <- F2R_Zabs.
apply bpow_le_F2R.
-apply F2R_lt_reg with beta emin.
+apply lt_F2R with beta emin.
rewrite F2R_0, F2R_Zabs.
now apply Rabs_pos_lt.
Qed.
@@ -237,15 +264,15 @@ rewrite Rmult_assoc.
apply Rmult_le_compat_l.
apply Rlt_le.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
assert (Hx': (x <> 0)%R).
intros H.
apply Rlt_not_le with (2 := Hx).
rewrite H, Rabs_R0.
apply bpow_gt_0.
rewrite ulp_neq_0; trivial.
-unfold canonic_exp.
-destruct (ln_beta beta x) as (ex, He).
+unfold cexp.
+destruct (mag beta x) as (ex, He).
simpl.
specialize (He Hx').
apply Rle_trans with (bpow (-p + 1) * bpow (ex - 1))%R.
@@ -274,7 +301,7 @@ apply relative_error_le_conversion...
apply Rlt_le.
apply Rmult_lt_0_compat.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
apply bpow_gt_0.
now apply relative_error_N.
Qed.
@@ -296,7 +323,7 @@ apply relative_error_N.
unfold x.
rewrite <- F2R_Zabs.
apply bpow_le_F2R.
-apply F2R_lt_reg with beta emin.
+apply lt_F2R with beta emin.
rewrite F2R_0, F2R_Zabs.
now apply Rabs_pos_lt.
Qed.
@@ -311,7 +338,7 @@ apply relative_error_le_conversion...
apply Rlt_le.
apply Rmult_lt_0_compat.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
apply bpow_gt_0.
now apply relative_error_N_F2R_emin.
Qed.
@@ -329,15 +356,15 @@ rewrite Rmult_assoc.
apply Rmult_le_compat_l.
apply Rlt_le.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
assert (Hx': (x <> 0)%R).
intros H.
apply Rlt_not_le with (2 := Hx).
rewrite H, Rabs_R0.
apply bpow_gt_0.
rewrite ulp_neq_0; trivial.
-unfold canonic_exp.
-destruct (ln_beta beta x) as (ex, He).
+unfold cexp.
+destruct (mag beta x) as (ex, He).
simpl.
specialize (He Hx').
assert (He': (emin < ex)%Z).
@@ -381,17 +408,250 @@ apply relative_error_N_round with (1 := Hp).
unfold x.
rewrite <- F2R_Zabs.
apply bpow_le_F2R.
-apply F2R_lt_reg with beta emin.
+apply lt_F2R with beta emin.
rewrite F2R_0, F2R_Zabs.
now apply Rabs_pos_lt.
Qed.
End Fprop_relative_generic.
+Section Fprop_relative_FLX.
+
+Variable prec : Z.
+Variable Hp : Z.lt 0 prec.
+
+Lemma relative_error_FLX_aux :
+ forall k, (prec <= k - FLX_exp prec k)%Z.
+Proof.
+intros k.
+unfold FLX_exp.
+omega.
+Qed.
+
+Variable rnd : R -> Z.
+Context { valid_rnd : Valid_rnd rnd }.
+
+Theorem relative_error_FLX :
+ forall x,
+ (x <> 0)%R ->
+ (Rabs (round beta (FLX_exp prec) rnd x - x) < bpow (-prec + 1) * Rabs x)%R.
+Proof with auto with typeclass_instances.
+intros x Hx.
+destruct (mag beta x) as (ex, He).
+specialize (He Hx).
+apply relative_error with (ex - 1)%Z...
+intros k _.
+apply relative_error_FLX_aux.
+apply He.
+Qed.
+
+(** 1+#&epsilon;# property in any rounding in FLX *)
+Theorem relative_error_FLX_ex :
+ forall x,
+ exists eps,
+ (Rabs eps < bpow (-prec + 1))%R /\ round beta (FLX_exp prec) rnd x = (x * (1 + eps))%R.
+Proof with auto with typeclass_instances.
+intros x.
+apply relative_error_lt_conversion...
+apply bpow_gt_0.
+now apply relative_error_FLX.
+Qed.
+
+Theorem relative_error_FLX_round :
+ forall x,
+ (x <> 0)%R ->
+ (Rabs (round beta (FLX_exp prec) rnd x - x) < bpow (-prec + 1) * Rabs (round beta (FLX_exp prec) rnd x))%R.
+Proof with auto with typeclass_instances.
+intros x Hx.
+destruct (mag beta x) as (ex, He).
+specialize (He Hx).
+apply relative_error_round with (ex - 1)%Z...
+intros k _.
+apply relative_error_FLX_aux.
+apply He.
+Qed.
+
+Variable choice : Z -> bool.
+
+Theorem relative_error_N_FLX :
+ forall x,
+ (Rabs (round beta (FLX_exp prec) (Znearest choice) x - x) <= /2 * bpow (-prec + 1) * Rabs x)%R.
+Proof with auto with typeclass_instances.
+intros x.
+destruct (Req_dec x 0) as [Hx|Hx].
+(* . *)
+rewrite Hx, round_0...
+unfold Rminus.
+rewrite Rplus_0_l, Rabs_Ropp, Rabs_R0.
+rewrite Rmult_0_r.
+apply Rle_refl.
+(* . *)
+destruct (mag beta x) as (ex, He).
+specialize (He Hx).
+apply relative_error_N with (ex - 1)%Z...
+intros k _.
+apply relative_error_FLX_aux.
+apply He.
+Qed.
+
+(** unit roundoff *)
+Definition u_ro := (/2 * bpow (-prec + 1))%R.
+
+Lemma u_ro_pos : (0 <= u_ro)%R.
+Proof. apply Rmult_le_pos; [lra|apply bpow_ge_0]. Qed.
+
+Lemma u_ro_lt_1 : (u_ro < 1)%R.
+Proof.
+unfold u_ro; apply (Rmult_lt_reg_l 2); [lra|].
+rewrite <-Rmult_assoc, Rinv_r, Rmult_1_l, Rmult_1_r; [|lra].
+apply (Rle_lt_trans _ (bpow 0));
+ [apply bpow_le; omega|simpl; lra].
+Qed.
+
+Lemma u_rod1pu_ro_pos : (0 <= u_ro / (1 + u_ro))%R.
+Proof.
+apply Rmult_le_pos; [|apply Rlt_le, Rinv_0_lt_compat];
+assert (H := u_ro_pos); lra.
+Qed.
+
+Lemma u_rod1pu_ro_le_u_ro : (u_ro / (1 + u_ro) <= u_ro)%R.
+Proof.
+assert (Pu_ro := u_ro_pos).
+apply (Rmult_le_reg_r (1 + u_ro)); [lra|].
+unfold Rdiv; rewrite Rmult_assoc, Rinv_l; [|lra].
+assert (0 <= u_ro * u_ro)%R; [apply Rmult_le_pos|]; lra.
+Qed.
+
+Theorem relative_error_N_FLX' :
+ forall x,
+ (Rabs (round beta (FLX_exp prec) (Znearest choice) x - x)
+ <= u_ro / (1 + u_ro) * Rabs x)%R.
+Proof with auto with typeclass_instances.
+intro x.
+assert (Pu_ro : (0 <= u_ro)%R).
+{ apply Rmult_le_pos; [lra|apply bpow_ge_0]. }
+destruct (Req_dec x 0) as [Zx|Nzx].
+{ rewrite Zx, Rabs_R0, Rmult_0_r, round_0...
+ now unfold Rminus; rewrite Rplus_0_l, Rabs_Ropp, Rabs_R0; right. }
+set (ufpx := bpow (mag beta x - 1)%Z).
+set (rx := round _ _ _ _).
+assert (Pufpx : (0 <= ufpx)%R); [now apply bpow_ge_0|].
+assert (H_2_1 : (Rabs (rx - x) <= u_ro * ufpx)%R).
+{ refine (Rle_trans _ _ _ (error_le_half_ulp _ _ _ _) _);
+ [now apply FLX_exp_valid|right].
+ unfold ulp, cexp, FLX_exp, u_ro, ufpx; rewrite (Req_bool_false _ _ Nzx).
+ rewrite Rmult_assoc, <-bpow_plus; do 2 f_equal; ring. }
+assert (H_2_3 : (ufpx + Rabs (rx - x) <= Rabs x)%R).
+{ apply (Rplus_le_reg_r (- ufpx)); ring_simplify.
+ destruct (Rle_or_lt 0 x) as [Sx|Sx].
+ { apply (Rle_trans _ (Rabs (ufpx - x))).
+ { apply round_N_pt; [now apply FLX_exp_valid|].
+ apply generic_format_bpow; unfold FLX_exp; lia. }
+ rewrite Rabs_minus_sym, Rabs_pos_eq.
+ { now rewrite Rabs_pos_eq; [right; ring|]. }
+ apply (Rplus_le_reg_r ufpx); ring_simplify.
+ now rewrite <-(Rabs_pos_eq _ Sx); apply bpow_mag_le. }
+ apply (Rle_trans _ (Rabs (- ufpx - x))).
+ { apply round_N_pt; [now apply FLX_exp_valid|].
+ apply generic_format_opp, generic_format_bpow; unfold FLX_exp; lia. }
+ rewrite Rabs_pos_eq; [now rewrite Rabs_left; [right|]|].
+ apply (Rplus_le_reg_r x); ring_simplify.
+ rewrite <-(Ropp_involutive x); apply Ropp_le_contravar; unfold ufpx.
+ rewrite <-mag_opp, <-Rabs_pos_eq; [apply bpow_mag_le|]; lra. }
+assert (H : (Rabs ((rx - x) / x) <= u_ro / (1 + u_ro))%R).
+{ assert (H : (0 < ufpx + Rabs (rx - x))%R).
+ { apply Rplus_lt_le_0_compat; [apply bpow_gt_0|apply Rabs_pos]. }
+ apply (Rle_trans _ (Rabs (rx - x) / (ufpx + Rabs (rx - x)))).
+ { unfold Rdiv; rewrite Rabs_mult; apply Rmult_le_compat_l; [apply Rabs_pos|].
+ now rewrite (Rabs_Rinv _ Nzx); apply Rinv_le. }
+ apply (Rmult_le_reg_r ((ufpx + Rabs (rx - x)) * (1 + u_ro))).
+ { apply Rmult_lt_0_compat; lra. }
+ field_simplify; [unfold Rdiv; rewrite Rinv_1, !Rmult_1_r| |]; lra. }
+revert H; unfold Rdiv; rewrite Rabs_mult, (Rabs_Rinv _ Nzx); intro H.
+apply (Rmult_le_reg_r (/ Rabs x)); [now apply Rinv_0_lt_compat, Rabs_pos_lt|].
+now apply (Rle_trans _ _ _ H); right; field; split; [apply Rabs_no_R0|lra].
+Qed.
+
+(** 1+#&epsilon;# property in rounding to nearest in FLX *)
+Theorem relative_error_N_FLX_ex :
+ forall x,
+ exists eps,
+ (Rabs eps <= /2 * bpow (-prec + 1))%R /\ round beta (FLX_exp prec) (Znearest choice) x = (x * (1 + eps))%R.
+Proof with auto with typeclass_instances.
+intros x.
+apply relative_error_le_conversion...
+apply Rlt_le.
+apply Rmult_lt_0_compat.
+apply Rinv_0_lt_compat.
+now apply IZR_lt.
+apply bpow_gt_0.
+now apply relative_error_N_FLX.
+Qed.
+
+Theorem relative_error_N_FLX'_ex :
+ forall x,
+ exists eps,
+ (Rabs eps <= u_ro / (1 + u_ro))%R /\
+ round beta (FLX_exp prec) (Znearest choice) x = (x * (1 + eps))%R.
+Proof with auto with typeclass_instances.
+intros x.
+apply relative_error_le_conversion...
+{ apply u_rod1pu_ro_pos. }
+now apply relative_error_N_FLX'.
+Qed.
+
+Lemma relative_error_N_round_ex_derive :
+ forall x rx,
+ (exists eps, (Rabs eps <= u_ro / (1 + u_ro))%R /\ rx = (x * (1 + eps))%R) ->
+ exists eps, (Rabs eps <= u_ro)%R /\ x = (rx * (1 + eps))%R.
+Proof.
+intros x rx (d, (Bd, Hd)).
+assert (Pu_ro := u_ro_pos).
+assert (H := Rabs_le_inv _ _ Bd).
+assert (H' := u_rod1pu_ro_le_u_ro); assert (H'' := u_ro_lt_1).
+destruct (Req_dec rx 0) as [Zfx|Nzfx].
+{ exists 0%R; split; [now rewrite Rabs_R0|].
+ rewrite Rplus_0_r, Rmult_1_r, Zfx.
+ now rewrite Zfx in Hd; destruct (Rmult_integral _ _ (sym_eq Hd)); [|lra]. }
+destruct (Req_dec x 0) as [Zx|Nzx].
+{ now exfalso; revert Hd; rewrite Zx, Rmult_0_l. }
+set (d' := ((x - rx) / rx)%R).
+assert (Hd' : (Rabs d' <= u_ro)%R).
+{ unfold d'; rewrite Hd.
+ replace (_ / _)%R with (- d / (1 + d))%R; [|now field; split; lra].
+ unfold Rdiv; rewrite Rabs_mult, Rabs_Ropp.
+ rewrite (Rabs_pos_eq (/ _)); [|apply Rlt_le, Rinv_0_lt_compat; lra].
+ apply (Rmult_le_reg_r (1 + d)); [lra|].
+ rewrite Rmult_assoc, Rinv_l, Rmult_1_r; [|lra].
+ apply (Rle_trans _ _ _ Bd).
+ unfold Rdiv; apply Rmult_le_compat_l; [now apply u_ro_pos|].
+ apply (Rle_trans _ (1 - u_ro / (1 + u_ro))); [right; field|]; lra. }
+now exists d'; split; [|unfold d'; field].
+Qed.
+
+Theorem relative_error_N_FLX_round_ex :
+ forall x,
+ exists eps,
+ (Rabs eps <= u_ro)%R /\
+ x = (round beta (FLX_exp prec) (Znearest choice) x * (1 + eps))%R.
+Proof.
+intro x; apply relative_error_N_round_ex_derive, relative_error_N_FLX'_ex.
+Qed.
+
+Theorem relative_error_N_FLX_round :
+ forall x,
+ (Rabs (round beta (FLX_exp prec) (Znearest choice) x - x) <= /2 * bpow (-prec + 1) * Rabs(round beta (FLX_exp prec) (Znearest choice) x))%R.
+Proof.
+intro x.
+apply relative_error_le_conversion_round_inv, relative_error_N_FLX_round_ex.
+Qed.
+
+End Fprop_relative_FLX.
+
Section Fprop_relative_FLT.
Variable emin prec : Z.
-Variable Hp : Zlt 0 prec.
+Variable Hp : Z.lt 0 prec.
Lemma relative_error_FLT_aux :
forall k, (emin + prec - 1 < k)%Z -> (prec <= k - FLT_exp emin prec k)%Z.
@@ -486,7 +746,7 @@ apply relative_error_le_conversion...
apply Rlt_le.
apply Rmult_lt_0_compat.
apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
+now apply IZR_lt.
apply bpow_gt_0.
now apply relative_error_N_FLT.
Qed.
@@ -607,23 +867,84 @@ apply Rlt_le, pos_half_prf.
rewrite ulp_neq_0.
2: now apply Rgt_not_eq.
apply bpow_le.
-unfold FLT_exp, canonic_exp.
+unfold FLT_exp, cexp.
rewrite Zmax_right.
omega.
-destruct (ln_beta beta x) as (e,He); simpl.
+destruct (mag beta x) as (e,He); simpl.
assert (e-1 < emin+prec)%Z.
apply (lt_bpow beta).
apply Rle_lt_trans with (2:=Hx).
-rewrite <- (Rabs_right x).
-apply He; auto with real.
-apply Rle_ge; now left.
+rewrite <- (Rabs_pos_eq x) by now apply Rlt_le.
+now apply He, Rgt_not_eq.
omega.
-split;ring.
+split ; ring.
+Qed.
+
+Theorem relative_error_N_FLT'_ex :
+ forall x,
+ exists eps eta : R,
+ (Rabs eps <= u_ro prec / (1 + u_ro prec))%R /\
+ (Rabs eta <= /2 * bpow emin)%R /\
+ (eps * eta = 0)%R /\
+ round beta (FLT_exp emin prec) (Znearest choice) x
+ = (x * (1 + eps) + eta)%R.
+Proof.
+intro x.
+set (rx := round _ _ _ x).
+assert (Pb := u_rod1pu_ro_pos prec).
+destruct (Rle_or_lt (bpow (emin + prec - 1)) (Rabs x)) as [MX|Mx].
+{ destruct (relative_error_N_FLX'_ex prec Hp choice x) as (d, (Bd, Hd)).
+ exists d, 0%R; split; [exact Bd|]; split.
+ { rewrite Rabs_R0; apply Rmult_le_pos; [lra|apply bpow_ge_0]. }
+ rewrite Rplus_0_r, Rmult_0_r; split; [reflexivity|].
+ now rewrite <- Hd; apply round_FLT_FLX. }
+assert (H : (Rabs (rx - x) <= /2 * bpow emin)%R).
+{ refine (Rle_trans _ _ _ (error_le_half_ulp _ _ _ _) _);
+ [now apply FLT_exp_valid|].
+ rewrite ulp_FLT_small; [now right|now simpl|].
+ apply (Rlt_le_trans _ _ _ Mx), bpow_le; lia. }
+exists 0%R, (rx - x)%R; split; [now rewrite Rabs_R0|]; split; [exact H|].
+now rewrite Rmult_0_l, Rplus_0_r, Rmult_1_r; split; [|ring].
+Qed.
+
+Theorem relative_error_N_FLT'_ex_separate :
+ forall x,
+ exists x' : R,
+ round beta (FLT_exp emin prec) (Znearest choice) x'
+ = round beta (FLT_exp emin prec) (Znearest choice) x /\
+ (exists eta, Rabs eta <= /2 * bpow emin /\ x' = x + eta)%R /\
+ (exists eps, Rabs eps <= u_ro prec / (1 + u_ro prec) /\
+ round beta (FLT_exp emin prec) (Znearest choice) x'
+ = x' * (1 + eps))%R.
+Proof.
+intro x.
+set (rx := round _ _ _ x).
+destruct (relative_error_N_FLT'_ex x) as (d, (e, (Bd, (Be, (Hde0, Hde))))).
+destruct (Rlt_or_le (Rabs (d * x)) (Rabs e)) as [HdxLte|HeLedx].
+{ exists rx; split; [|split].
+ { apply round_generic; [now apply valid_rnd_N|].
+ now apply generic_format_round; [apply FLT_exp_valid|apply valid_rnd_N]. }
+ { exists e; split; [exact Be|].
+ unfold rx; rewrite Hde; destruct (Rmult_integral _ _ Hde0) as [Zd|Ze].
+ { now rewrite Zd, Rplus_0_r, Rmult_1_r. }
+ exfalso; revert HdxLte; rewrite Ze, Rabs_R0; apply Rle_not_lt, Rabs_pos. }
+ exists 0%R; split; [now rewrite Rabs_R0; apply u_rod1pu_ro_pos|].
+ rewrite Rplus_0_r, Rmult_1_r; apply round_generic; [now apply valid_rnd_N|].
+ now apply generic_format_round; [apply FLT_exp_valid|apply valid_rnd_N]. }
+exists x; split; [now simpl|split].
+{ exists 0%R; split;
+ [rewrite Rabs_R0; apply Rmult_le_pos; [lra|apply bpow_ge_0]|ring]. }
+exists d; rewrite Hde; destruct (Rmult_integral _ _ Hde0) as [Zd|Ze].
+{ split; [exact Bd|].
+ assert (Ze : e = 0%R); [|now rewrite Ze, Rplus_0_r].
+ apply Rabs_eq_R0, Rle_antisym; [|now apply Rabs_pos].
+ now revert HeLedx; rewrite Zd, Rmult_0_l, Rabs_R0. }
+now rewrite Ze, Rplus_0_r.
Qed.
End Fprop_relative_FLT.
-Lemma error_N_FLT :
+Theorem error_N_FLT :
forall (emin prec : Z), (0 < prec)%Z ->
forall (choice : Z -> bool),
forall (x : R),
@@ -638,9 +959,9 @@ intros emin prec Pprec choice x.
destruct (Rtotal_order x 0) as [Nx|[Zx|Px]].
{ assert (Pmx : (0 < - x)%R).
{ now rewrite <- Ropp_0; apply Ropp_lt_contravar. }
- destruct (error_N_FLT_aux emin prec Pprec
- (fun t : Z => negb (choice (- (t + 1))%Z))
- (- x)%R Pmx)
+ destruct (@error_N_FLT_aux emin prec Pprec
+ (fun t : Z => negb (choice (- (t + 1))%Z))
+ (- x)%R Pmx)
as (d,(e,(Hd,(He,(Hde,Hr))))).
exists d; exists (- e)%R; split; [exact Hd|split; [|split]].
{ now rewrite Rabs_Ropp. }
@@ -659,124 +980,4 @@ destruct (Rtotal_order x 0) as [Nx|[Zx|Px]].
now apply error_N_FLT_aux.
Qed.
-Section Fprop_relative_FLX.
-
-Variable prec : Z.
-Variable Hp : Zlt 0 prec.
-
-Lemma relative_error_FLX_aux :
- forall k, (prec <= k - FLX_exp prec k)%Z.
-Proof.
-intros k.
-unfold FLX_exp.
-omega.
-Qed.
-
-Variable rnd : R -> Z.
-Context { valid_rnd : Valid_rnd rnd }.
-
-Theorem relative_error_FLX :
- forall x,
- (x <> 0)%R ->
- (Rabs (round beta (FLX_exp prec) rnd x - x) < bpow (-prec + 1) * Rabs x)%R.
-Proof with auto with typeclass_instances.
-intros x Hx.
-destruct (ln_beta beta x) as (ex, He).
-specialize (He Hx).
-apply relative_error with (ex - 1)%Z...
-intros k _.
-apply relative_error_FLX_aux.
-apply He.
-Qed.
-
-(** 1+#&epsilon;# property in any rounding in FLX *)
-Theorem relative_error_FLX_ex :
- forall x,
- exists eps,
- (Rabs eps < bpow (-prec + 1))%R /\ round beta (FLX_exp prec) rnd x = (x * (1 + eps))%R.
-Proof with auto with typeclass_instances.
-intros x.
-apply relative_error_lt_conversion...
-apply bpow_gt_0.
-now apply relative_error_FLX.
-Qed.
-
-Theorem relative_error_FLX_round :
- forall x,
- (x <> 0)%R ->
- (Rabs (round beta (FLX_exp prec) rnd x - x) < bpow (-prec + 1) * Rabs (round beta (FLX_exp prec) rnd x))%R.
-Proof with auto with typeclass_instances.
-intros x Hx.
-destruct (ln_beta beta x) as (ex, He).
-specialize (He Hx).
-apply relative_error_round with (ex - 1)%Z...
-intros k _.
-apply relative_error_FLX_aux.
-apply He.
-Qed.
-
-Variable choice : Z -> bool.
-
-Theorem relative_error_N_FLX :
- forall x,
- (Rabs (round beta (FLX_exp prec) (Znearest choice) x - x) <= /2 * bpow (-prec + 1) * Rabs x)%R.
-Proof with auto with typeclass_instances.
-intros x.
-destruct (Req_dec x 0) as [Hx|Hx].
-(* . *)
-rewrite Hx, round_0...
-unfold Rminus.
-rewrite Rplus_0_l, Rabs_Ropp, Rabs_R0.
-rewrite Rmult_0_r.
-apply Rle_refl.
-(* . *)
-destruct (ln_beta beta x) as (ex, He).
-specialize (He Hx).
-apply relative_error_N with (ex - 1)%Z...
-intros k _.
-apply relative_error_FLX_aux.
-apply He.
-Qed.
-
-(** 1+#&epsilon;# property in rounding to nearest in FLX *)
-Theorem relative_error_N_FLX_ex :
- forall x,
- exists eps,
- (Rabs eps <= /2 * bpow (-prec + 1))%R /\ round beta (FLX_exp prec) (Znearest choice) x = (x * (1 + eps))%R.
-Proof with auto with typeclass_instances.
-intros x.
-apply relative_error_le_conversion...
-apply Rlt_le.
-apply Rmult_lt_0_compat.
-apply Rinv_0_lt_compat.
-now apply (Z2R_lt 0 2).
-apply bpow_gt_0.
-now apply relative_error_N_FLX.
-Qed.
-
-Theorem relative_error_N_FLX_round :
- forall x,
- (Rabs (round beta (FLX_exp prec) (Znearest choice) x - x) <= /2 * bpow (-prec + 1) * Rabs (round beta (FLX_exp prec) (Znearest choice) x))%R.
-Proof with auto with typeclass_instances.
-intros x.
-destruct (Req_dec x 0) as [Hx|Hx].
-(* . *)
-rewrite Hx, round_0...
-unfold Rminus.
-rewrite Rplus_0_l, Rabs_Ropp, Rabs_R0.
-rewrite Rmult_0_r.
-apply Rle_refl.
-(* . *)
-destruct (ln_beta beta x) as (ex, He).
-specialize (He Hx).
-apply relative_error_N_round with (ex - 1)%Z.
-now apply FLX_exp_valid.
-intros k _.
-apply relative_error_FLX_aux.
-exact Hp.
-apply He.
-Qed.
-
-End Fprop_relative_FLX.
-
-End Fprop_relative. \ No newline at end of file
+End Fprop_relative.
diff --git a/flocq/Prop/Round_odd.v b/flocq/Prop/Round_odd.v
new file mode 100644
index 00000000..df2952cc
--- /dev/null
+++ b/flocq/Prop/Round_odd.v
@@ -0,0 +1,1220 @@
+(**
+This file is part of the Flocq formalization of floating-point
+arithmetic in Coq: http://flocq.gforge.inria.fr/
+
+Copyright (C) 2013-2018 Sylvie Boldo
+#<br />#
+Copyright (C) 2013-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.
+*)
+
+(** * Rounding to odd and its properties, including the equivalence
+ between rnd_NE and double rounding with rnd_odd and then rnd_NE *)
+
+Require Import Reals Psatz.
+Require Import Core Operations.
+
+Definition Zrnd_odd x := match Req_EM_T x (IZR (Zfloor x)) with
+ | left _ => Zfloor x
+ | right _ => match (Z.even (Zfloor x)) with
+ | true => Zceil x
+ | false => Zfloor x
+ end
+ end.
+
+
+
+Global Instance valid_rnd_odd : Valid_rnd Zrnd_odd.
+Proof.
+split.
+(* . *)
+intros x y Hxy.
+assert (Zfloor x <= Zrnd_odd y)%Z.
+(* .. *)
+apply Z.le_trans with (Zfloor y).
+now apply Zfloor_le.
+unfold Zrnd_odd; destruct (Req_EM_T y (IZR (Zfloor y))).
+now apply Z.le_refl.
+case (Z.even (Zfloor y)).
+apply le_IZR.
+apply Rle_trans with y.
+apply Zfloor_lb.
+apply Zceil_ub.
+now apply Z.le_refl.
+unfold Zrnd_odd at 1.
+(* . *)
+destruct (Req_EM_T x (IZR (Zfloor x))) as [Hx|Hx].
+(* .. *)
+apply H.
+(* .. *)
+case_eq (Z.even (Zfloor x)); intros Hx2.
+2: apply H.
+unfold Zrnd_odd; destruct (Req_EM_T y (IZR (Zfloor y))) as [Hy|Hy].
+apply Zceil_glb.
+now rewrite <- Hy.
+case_eq (Z.even (Zfloor y)); intros Hy2.
+now apply Zceil_le.
+apply Zceil_glb.
+assert (H0:(Zfloor x <= Zfloor y)%Z) by now apply Zfloor_le.
+case (Zle_lt_or_eq _ _ H0); intros H1.
+apply Rle_trans with (1:=Zceil_ub _).
+rewrite Zceil_floor_neq.
+apply IZR_le; omega.
+now apply sym_not_eq.
+contradict Hy2.
+rewrite <- H1, Hx2; discriminate.
+(* . *)
+intros n; unfold Zrnd_odd.
+rewrite Zfloor_IZR, Zceil_IZR.
+destruct (Req_EM_T (IZR n) (IZR n)); trivial.
+case (Z.even n); trivial.
+Qed.
+
+
+
+Lemma Zrnd_odd_Zodd: forall x, x <> (IZR (Zfloor x)) ->
+ (Z.even (Zrnd_odd x)) = false.
+Proof.
+intros x Hx; unfold Zrnd_odd.
+destruct (Req_EM_T x (IZR (Zfloor x))) as [H|H].
+now contradict H.
+case_eq (Z.even (Zfloor x)).
+(* difficult case *)
+intros H'.
+rewrite Zceil_floor_neq.
+rewrite Z.even_add, H'.
+reflexivity.
+now apply sym_not_eq.
+trivial.
+Qed.
+
+
+Lemma Zfloor_plus: forall (n:Z) y,
+ (Zfloor (IZR n+y) = n + Zfloor y)%Z.
+Proof.
+intros n y; unfold Zfloor.
+unfold Zminus; rewrite Zplus_assoc; f_equal.
+apply sym_eq, tech_up.
+rewrite plus_IZR.
+apply Rplus_lt_compat_l.
+apply archimed.
+rewrite plus_IZR, Rplus_assoc.
+apply Rplus_le_compat_l.
+apply Rplus_le_reg_r with (-y)%R.
+ring_simplify (y+1+-y)%R.
+apply archimed.
+Qed.
+
+Lemma Zceil_plus: forall (n:Z) y,
+ (Zceil (IZR n+y) = n + Zceil y)%Z.
+Proof.
+intros n y; unfold Zceil.
+rewrite Ropp_plus_distr, <- Ropp_Ropp_IZR.
+rewrite Zfloor_plus.
+ring.
+Qed.
+
+
+Lemma Zeven_abs: forall z, Z.even (Z.abs z) = Z.even z.
+Proof.
+intros z; case (Zle_or_lt z 0); intros H1.
+rewrite Z.abs_neq; try assumption.
+apply Z.even_opp.
+rewrite Z.abs_eq; auto with zarith.
+Qed.
+
+
+
+
+Lemma Zrnd_odd_plus: forall x y, (x = IZR (Zfloor x)) ->
+ Z.even (Zfloor x) = true ->
+ (IZR (Zrnd_odd (x+y)) = x+IZR (Zrnd_odd y))%R.
+Proof.
+intros x y Hx H.
+unfold Zrnd_odd; rewrite Hx, Zfloor_plus.
+case (Req_EM_T y (IZR (Zfloor y))); intros Hy.
+rewrite Hy; repeat rewrite <- plus_IZR.
+repeat rewrite Zfloor_IZR.
+case (Req_EM_T _ _); intros K; easy.
+case (Req_EM_T _ _); intros K.
+contradict Hy.
+apply Rplus_eq_reg_l with (IZR (Zfloor x)).
+now rewrite K, plus_IZR.
+rewrite Z.even_add, H; simpl.
+case (Z.even (Zfloor y)).
+now rewrite Zceil_plus, plus_IZR.
+now rewrite plus_IZR.
+Qed.
+
+
+Section Fcore_rnd_odd.
+
+Variable beta : radix.
+
+Notation bpow e := (bpow beta e).
+
+Variable fexp : Z -> Z.
+
+Context { valid_exp : Valid_exp fexp }.
+Context { exists_NE_ : Exists_NE beta fexp }.
+
+Notation format := (generic_format beta fexp).
+Notation canonical := (canonical beta fexp).
+Notation cexp := (cexp beta fexp).
+
+
+Definition Rnd_odd_pt (x f : R) :=
+ format f /\ ((f = x)%R \/
+ ((Rnd_DN_pt format x f \/ Rnd_UP_pt format x f) /\
+ exists g : float beta, f = F2R g /\ canonical g /\ Z.even (Fnum g) = false)).
+
+Definition Rnd_odd (rnd : R -> R) :=
+ forall x : R, Rnd_odd_pt x (rnd x).
+
+Theorem Rnd_odd_pt_opp_inv : forall x f : R,
+ Rnd_odd_pt (-x) (-f) -> Rnd_odd_pt x f.
+Proof with auto with typeclass_instances.
+intros x f (H1,H2).
+split.
+replace f with (-(-f))%R by ring.
+now apply generic_format_opp.
+destruct H2.
+left.
+replace f with (-(-f))%R by ring.
+rewrite H; ring.
+right.
+destruct H as (H2,(g,(Hg1,(Hg2,Hg3)))).
+split.
+destruct H2.
+right.
+replace f with (-(-f))%R by ring.
+replace x with (-(-x))%R by ring.
+apply Rnd_UP_pt_opp...
+apply generic_format_opp.
+left.
+replace f with (-(-f))%R by ring.
+replace x with (-(-x))%R by ring.
+apply Rnd_DN_pt_opp...
+apply generic_format_opp.
+exists (Float beta (-Fnum g) (Fexp g)).
+split.
+rewrite F2R_Zopp.
+replace f with (-(-f))%R by ring.
+rewrite Hg1; reflexivity.
+split.
+now apply canonical_opp.
+simpl.
+now rewrite Z.even_opp.
+Qed.
+
+
+Theorem round_odd_opp :
+ forall x,
+ round beta fexp Zrnd_odd (-x) = (- round beta fexp Zrnd_odd x)%R.
+Proof.
+intros x; unfold round.
+rewrite <- F2R_Zopp.
+unfold F2R; simpl.
+apply f_equal2; apply f_equal.
+rewrite scaled_mantissa_opp.
+generalize (scaled_mantissa beta fexp x); intros r.
+unfold Zrnd_odd.
+case (Req_EM_T (- r) (IZR (Zfloor (- r)))).
+case (Req_EM_T r (IZR (Zfloor r))).
+intros Y1 Y2.
+apply eq_IZR.
+now rewrite opp_IZR, <- Y1, <-Y2.
+intros Y1 Y2.
+absurd (r=IZR (Zfloor r)); trivial.
+pattern r at 2; replace r with (-(-r))%R by ring.
+rewrite Y2, <- opp_IZR.
+rewrite Zfloor_IZR.
+rewrite opp_IZR, <- Y2.
+ring.
+case (Req_EM_T r (IZR (Zfloor r))).
+intros Y1 Y2.
+absurd (-r=IZR (Zfloor (-r)))%R; trivial.
+pattern r at 2; rewrite Y1.
+rewrite <- opp_IZR, Zfloor_IZR.
+now rewrite opp_IZR, <- Y1.
+intros Y1 Y2.
+unfold Zceil; rewrite Ropp_involutive.
+replace (Z.even (Zfloor (- r))) with (negb (Z.even (Zfloor r))).
+case (Z.even (Zfloor r)); simpl; ring.
+apply trans_eq with (Z.even (Zceil r)).
+rewrite Zceil_floor_neq.
+rewrite Z.even_add.
+destruct (Z.even (Zfloor r)); reflexivity.
+now apply sym_not_eq.
+rewrite <- (Z.even_opp (Zfloor (- r))).
+reflexivity.
+apply cexp_opp.
+Qed.
+
+
+
+Theorem round_odd_pt :
+ forall x,
+ Rnd_odd_pt x (round beta fexp Zrnd_odd x).
+Proof with auto with typeclass_instances.
+cut (forall x : R, (0 < x)%R -> Rnd_odd_pt x (round beta fexp Zrnd_odd x)).
+intros H x; case (Rle_or_lt 0 x).
+intros H1; destruct H1.
+now apply H.
+rewrite <- H0.
+rewrite round_0...
+split.
+apply generic_format_0.
+now left.
+intros Hx; apply Rnd_odd_pt_opp_inv.
+rewrite <- round_odd_opp.
+apply H.
+auto with real.
+(* *)
+intros x Hxp.
+generalize (generic_format_round beta fexp Zrnd_odd x).
+set (o:=round beta fexp Zrnd_odd x).
+intros Ho.
+split.
+assumption.
+(* *)
+case (Req_dec o x); intros Hx.
+now left.
+right.
+assert (o=round beta fexp Zfloor x \/ o=round beta fexp Zceil x).
+unfold o, round, F2R;simpl.
+case (Zrnd_DN_or_UP Zrnd_odd (scaled_mantissa beta fexp x))...
+intros H; rewrite H; now left.
+intros H; rewrite H; now right.
+split.
+destruct H; rewrite H.
+left; apply round_DN_pt...
+right; apply round_UP_pt...
+(* *)
+unfold o, Zrnd_odd, round.
+case (Req_EM_T (scaled_mantissa beta fexp x)
+ (IZR (Zfloor (scaled_mantissa beta fexp x)))).
+intros T.
+absurd (o=x); trivial.
+apply round_generic...
+unfold generic_format, F2R; simpl.
+rewrite <- (scaled_mantissa_mult_bpow beta fexp) at 1.
+apply f_equal2; trivial; rewrite T at 1.
+apply f_equal, sym_eq, Ztrunc_floor.
+apply Rmult_le_pos.
+now left.
+apply bpow_ge_0.
+intros L.
+case_eq (Z.even (Zfloor (scaled_mantissa beta fexp x))).
+(* . *)
+generalize (generic_format_round beta fexp Zceil x).
+unfold generic_format.
+set (f:=round beta fexp Zceil x).
+set (ef := cexp f).
+set (mf := Ztrunc (scaled_mantissa beta fexp f)).
+exists (Float beta mf ef).
+unfold canonical.
+rewrite <- H0.
+repeat split; try assumption.
+apply trans_eq with (negb (Z.even (Zfloor (scaled_mantissa beta fexp x)))).
+2: rewrite H1; reflexivity.
+apply trans_eq with (negb (Z.even (Fnum
+ (Float beta (Zfloor (scaled_mantissa beta fexp x)) (cexp x))))).
+2: reflexivity.
+case (Rle_lt_or_eq_dec 0 (round beta fexp Zfloor x)).
+rewrite <- round_0 with beta fexp Zfloor...
+apply round_le...
+now left.
+intros Y.
+generalize (DN_UP_parity_generic beta fexp)...
+unfold DN_UP_parity_prop.
+intros T; apply T with x; clear T.
+unfold generic_format.
+rewrite <- (scaled_mantissa_mult_bpow beta fexp x) at 1.
+unfold F2R; simpl.
+apply Rmult_neq_compat_r.
+apply Rgt_not_eq, bpow_gt_0.
+rewrite Ztrunc_floor.
+assumption.
+apply Rmult_le_pos.
+now left.
+apply bpow_ge_0.
+unfold canonical.
+simpl.
+apply sym_eq, cexp_DN...
+unfold canonical.
+rewrite <- H0; reflexivity.
+reflexivity.
+apply trans_eq with (round beta fexp Ztrunc (round beta fexp Zceil x)).
+reflexivity.
+apply round_generic...
+intros Y.
+replace (Fnum {| Fnum := Zfloor (scaled_mantissa beta fexp x); Fexp := cexp x |})
+ with (Fnum (Float beta 0 (fexp (mag beta 0)))).
+generalize (DN_UP_parity_generic beta fexp)...
+unfold DN_UP_parity_prop.
+intros T; apply T with x; clear T.
+unfold generic_format.
+rewrite <- (scaled_mantissa_mult_bpow beta fexp x) at 1.
+unfold F2R; simpl.
+apply Rmult_neq_compat_r.
+apply Rgt_not_eq, bpow_gt_0.
+rewrite Ztrunc_floor.
+assumption.
+apply Rmult_le_pos.
+now left.
+apply bpow_ge_0.
+apply canonical_0.
+unfold canonical.
+rewrite <- H0; reflexivity.
+rewrite <- Y; unfold F2R; simpl; ring.
+apply trans_eq with (round beta fexp Ztrunc (round beta fexp Zceil x)).
+reflexivity.
+apply round_generic...
+simpl.
+apply eq_IZR, Rmult_eq_reg_r with (bpow (cexp x)).
+unfold round, F2R in Y; simpl in Y; rewrite <- Y.
+simpl; ring.
+apply Rgt_not_eq, bpow_gt_0.
+(* . *)
+intros Y.
+case (Rle_lt_or_eq_dec 0 (round beta fexp Zfloor x)).
+rewrite <- round_0 with beta fexp Zfloor...
+apply round_le...
+now left.
+intros Hrx.
+set (ef := cexp x).
+set (mf := Zfloor (scaled_mantissa beta fexp x)).
+exists (Float beta mf ef).
+unfold canonical.
+repeat split; try assumption.
+simpl.
+apply trans_eq with (cexp (round beta fexp Zfloor x )).
+apply sym_eq, cexp_DN...
+reflexivity.
+intros Hrx; contradict Y.
+replace (Zfloor (scaled_mantissa beta fexp x)) with 0%Z.
+simpl; discriminate.
+apply eq_IZR, Rmult_eq_reg_r with (bpow (cexp x)).
+unfold round, F2R in Hrx; simpl in Hrx; rewrite <- Hrx.
+simpl; ring.
+apply Rgt_not_eq, bpow_gt_0.
+Qed.
+
+Theorem Rnd_odd_pt_unique :
+ forall x f1 f2 : R,
+ Rnd_odd_pt x f1 -> Rnd_odd_pt x f2 ->
+ f1 = f2.
+Proof.
+intros x f1 f2 (Ff1,H1) (Ff2,H2).
+(* *)
+case (generic_format_EM beta fexp x); intros L.
+apply trans_eq with x.
+case H1; try easy.
+intros (J,_); case J; intros J'.
+apply Rnd_DN_pt_idempotent with format; assumption.
+apply Rnd_UP_pt_idempotent with format; assumption.
+case H2; try easy.
+intros (J,_); case J; intros J'; apply sym_eq.
+apply Rnd_DN_pt_idempotent with format; assumption.
+apply Rnd_UP_pt_idempotent with format; assumption.
+(* *)
+destruct H1 as [H1|(H1,H1')].
+contradict L; now rewrite <- H1.
+destruct H2 as [H2|(H2,H2')].
+contradict L; now rewrite <- H2.
+destruct H1 as [H1|H1]; destruct H2 as [H2|H2].
+apply Rnd_DN_pt_unique with format x; assumption.
+destruct H1' as (ff,(K1,(K2,K3))).
+destruct H2' as (gg,(L1,(L2,L3))).
+absurd (true = false); try discriminate.
+rewrite <- L3.
+apply trans_eq with (negb (Z.even (Fnum ff))).
+rewrite K3; easy.
+apply sym_eq.
+generalize (DN_UP_parity_generic beta fexp).
+unfold DN_UP_parity_prop; intros T; apply (T x); clear T; try assumption...
+rewrite <- K1; apply Rnd_DN_pt_unique with (generic_format beta fexp) x; try easy...
+now apply round_DN_pt...
+rewrite <- L1; apply Rnd_UP_pt_unique with (generic_format beta fexp) x; try easy...
+now apply round_UP_pt...
+(* *)
+destruct H1' as (ff,(K1,(K2,K3))).
+destruct H2' as (gg,(L1,(L2,L3))).
+absurd (true = false); try discriminate.
+rewrite <- K3.
+apply trans_eq with (negb (Z.even (Fnum gg))).
+rewrite L3; easy.
+apply sym_eq.
+generalize (DN_UP_parity_generic beta fexp).
+unfold DN_UP_parity_prop; intros T; apply (T x); clear T; try assumption...
+rewrite <- L1; apply Rnd_DN_pt_unique with (generic_format beta fexp) x; try easy...
+now apply round_DN_pt...
+rewrite <- K1; apply Rnd_UP_pt_unique with (generic_format beta fexp) x; try easy...
+now apply round_UP_pt...
+apply Rnd_UP_pt_unique with format x; assumption.
+Qed.
+
+Theorem Rnd_odd_pt_monotone :
+ round_pred_monotone (Rnd_odd_pt).
+Proof with auto with typeclass_instances.
+intros x y f g H1 H2 Hxy.
+apply Rle_trans with (round beta fexp Zrnd_odd x).
+right; apply Rnd_odd_pt_unique with x; try assumption.
+apply round_odd_pt.
+apply Rle_trans with (round beta fexp Zrnd_odd y).
+apply round_le...
+right; apply Rnd_odd_pt_unique with y; try assumption.
+apply round_odd_pt.
+Qed.
+
+End Fcore_rnd_odd.
+
+Section Odd_prop_aux.
+
+Variable beta : radix.
+Hypothesis Even_beta: Z.even (radix_val beta)=true.
+
+Notation bpow e := (bpow beta e).
+
+Variable fexp : Z -> Z.
+Variable fexpe : Z -> Z.
+
+Context { valid_exp : Valid_exp fexp }.
+Context { exists_NE_ : Exists_NE beta fexp }. (* for underflow reason *)
+Context { valid_expe : Valid_exp fexpe }.
+Context { exists_NE_e : Exists_NE beta fexpe }. (* for defining rounding to odd *)
+
+Hypothesis fexpe_fexp: forall e, (fexpe e <= fexp e -2)%Z.
+
+
+Lemma generic_format_fexpe_fexp: forall x,
+ generic_format beta fexp x -> generic_format beta fexpe x.
+Proof.
+intros x Hx.
+apply generic_inclusion_mag with fexp; trivial; intros Hx2.
+generalize (fexpe_fexp (mag beta x)).
+omega.
+Qed.
+
+
+
+Lemma exists_even_fexp_lt: forall (c:Z->Z), forall (x:R),
+ (exists f:float beta, F2R f = x /\ (c (mag beta x) < Fexp f)%Z) ->
+ exists f:float beta, F2R f =x /\ canonical beta c f /\ Z.even (Fnum f) = true.
+Proof with auto with typeclass_instances.
+intros c x (g,(Hg1,Hg2)).
+exists (Float beta
+ (Fnum g*Z.pow (radix_val beta) (Fexp g - c (mag beta x)))
+ (c (mag beta x))).
+assert (F2R (Float beta
+ (Fnum g*Z.pow (radix_val beta) (Fexp g - c (mag beta x)))
+ (c (mag beta x))) = x).
+unfold F2R; simpl.
+rewrite mult_IZR, IZR_Zpower.
+rewrite Rmult_assoc, <- bpow_plus.
+rewrite <- Hg1; unfold F2R.
+apply f_equal, f_equal.
+ring.
+omega.
+split; trivial.
+split.
+unfold canonical, cexp.
+now rewrite H.
+simpl.
+rewrite Z.even_mul.
+rewrite Z.even_pow.
+rewrite Even_beta.
+apply Bool.orb_true_intro.
+now right.
+omega.
+Qed.
+
+
+Variable choice:Z->bool.
+Variable x:R.
+
+
+Variable d u: float beta.
+Hypothesis Hd: Rnd_DN_pt (generic_format beta fexp) x (F2R d).
+Hypothesis Cd: canonical beta fexp d.
+Hypothesis Hu: Rnd_UP_pt (generic_format beta fexp) x (F2R u).
+Hypothesis Cu: canonical beta fexp u.
+
+Hypothesis xPos: (0 < x)%R.
+
+
+Let m:= ((F2R d+F2R u)/2)%R.
+
+
+Lemma d_eq: F2R d= round beta fexp Zfloor x.
+Proof with auto with typeclass_instances.
+apply Rnd_DN_pt_unique with (generic_format beta fexp) x...
+apply round_DN_pt...
+Qed.
+
+
+Lemma u_eq: F2R u= round beta fexp Zceil x.
+Proof with auto with typeclass_instances.
+apply Rnd_UP_pt_unique with (generic_format beta fexp) x...
+apply round_UP_pt...
+Qed.
+
+
+Lemma d_ge_0: (0 <= F2R d)%R.
+Proof with auto with typeclass_instances.
+rewrite d_eq; apply round_ge_generic...
+apply generic_format_0.
+now left.
+Qed.
+
+
+
+Lemma mag_d: (0< F2R d)%R ->
+ (mag beta (F2R d) = mag beta x :>Z).
+Proof with auto with typeclass_instances.
+intros Y.
+rewrite d_eq; apply mag_DN...
+now rewrite <- d_eq.
+Qed.
+
+
+Lemma Fexp_d: (0 < F2R d)%R -> Fexp d =fexp (mag beta x).
+Proof with auto with typeclass_instances.
+intros Y.
+now rewrite Cd, <- mag_d.
+Qed.
+
+
+
+Lemma format_bpow_x: (0 < F2R d)%R
+ -> generic_format beta fexp (bpow (mag beta x)).
+Proof with auto with typeclass_instances.
+intros Y.
+apply generic_format_bpow.
+apply valid_exp.
+rewrite <- Fexp_d; trivial.
+apply Z.lt_le_trans with (mag beta (F2R d))%Z.
+rewrite Cd; apply mag_generic_gt...
+now apply Rgt_not_eq.
+apply Hd.
+apply mag_le; trivial.
+apply Hd.
+Qed.
+
+
+Lemma format_bpow_d: (0 < F2R d)%R ->
+ generic_format beta fexp (bpow (mag beta (F2R d))).
+Proof with auto with typeclass_instances.
+intros Y; apply generic_format_bpow.
+apply valid_exp.
+apply mag_generic_gt...
+now apply Rgt_not_eq.
+now apply generic_format_canonical.
+Qed.
+
+
+Lemma d_le_m: (F2R d <= m)%R.
+Proof.
+assert (F2R d <= F2R u)%R.
+ apply Rle_trans with x.
+ apply Hd.
+ apply Hu.
+unfold m.
+lra.
+Qed.
+
+Lemma m_le_u: (m <= F2R u)%R.
+Proof.
+assert (F2R d <= F2R u)%R.
+ apply Rle_trans with x.
+ apply Hd.
+ apply Hu.
+unfold m.
+lra.
+Qed.
+
+Lemma mag_m: (0 < F2R d)%R -> (mag beta m =mag beta (F2R d) :>Z).
+Proof with auto with typeclass_instances.
+intros dPos; apply mag_unique_pos.
+split.
+apply Rle_trans with (F2R d).
+destruct (mag beta (F2R d)) as (e,He).
+simpl.
+rewrite Rabs_right in He.
+apply He.
+now apply Rgt_not_eq.
+apply Rle_ge; now left.
+apply d_le_m.
+case m_le_u; intros H.
+apply Rlt_le_trans with (1:=H).
+rewrite u_eq.
+apply round_le_generic...
+apply generic_format_bpow.
+apply valid_exp.
+apply mag_generic_gt...
+now apply Rgt_not_eq.
+now apply generic_format_canonical.
+case (Rle_or_lt x (bpow (mag beta (F2R d)))); trivial; intros Z.
+absurd ((bpow (mag beta (F2R d)) <= (F2R d)))%R.
+apply Rlt_not_le.
+destruct (mag beta (F2R d)) as (e,He).
+simpl in *; rewrite Rabs_right in He.
+apply He.
+now apply Rgt_not_eq.
+apply Rle_ge; now left.
+apply Rle_trans with (round beta fexp Zfloor x).
+2: right; apply sym_eq, d_eq.
+apply round_ge_generic...
+apply generic_format_bpow.
+apply valid_exp.
+apply mag_generic_gt...
+now apply Rgt_not_eq.
+now apply generic_format_canonical.
+now left.
+replace m with (F2R d).
+destruct (mag beta (F2R d)) as (e,He).
+simpl in *; rewrite Rabs_right in He.
+apply He.
+now apply Rgt_not_eq.
+apply Rle_ge; now left.
+unfold m in H |- *.
+lra.
+Qed.
+
+
+Lemma mag_m_0: (0 = F2R d)%R
+ -> (mag beta m =mag beta (F2R u)-1:>Z)%Z.
+Proof with auto with typeclass_instances.
+intros Y.
+apply mag_unique_pos.
+unfold m; rewrite <- Y, Rplus_0_l.
+rewrite u_eq.
+destruct (mag beta x) as (e,He).
+rewrite Rabs_pos_eq in He by now apply Rlt_le.
+rewrite round_UP_small_pos with (ex:=e).
+rewrite mag_bpow.
+ring_simplify (fexp e + 1 - 1)%Z.
+split.
+unfold Zminus; rewrite bpow_plus.
+unfold Rdiv; apply Rmult_le_compat_l.
+apply bpow_ge_0.
+simpl; unfold Z.pow_pos; simpl.
+rewrite Zmult_1_r; apply Rinv_le.
+exact Rlt_0_2.
+apply IZR_le.
+specialize (radix_gt_1 beta).
+omega.
+apply Rlt_le_trans with (bpow (fexp e)*1)%R.
+2: right; ring.
+unfold Rdiv; apply Rmult_lt_compat_l.
+apply bpow_gt_0.
+lra.
+now apply He, Rgt_not_eq.
+apply exp_small_round_0_pos with beta (Zfloor) x...
+now apply He, Rgt_not_eq.
+now rewrite <- d_eq, Y.
+Qed.
+
+
+
+
+
+Lemma u'_eq: (0 < F2R d)%R -> exists f:float beta, F2R f = F2R u /\ (Fexp f = Fexp d)%Z.
+Proof with auto with typeclass_instances.
+intros Y.
+rewrite u_eq; unfold round.
+eexists; repeat split.
+simpl; now rewrite Fexp_d.
+Qed.
+
+
+Lemma m_eq :
+ (0 < F2R d)%R ->
+ exists f:float beta,
+ F2R f = m /\ (Fexp f = fexp (mag beta x) - 1)%Z.
+Proof with auto with typeclass_instances.
+intros Y.
+specialize (Zeven_ex (radix_val beta)); rewrite Even_beta.
+intros (b, Hb); rewrite Zplus_0_r in Hb.
+destruct u'_eq as (u', (Hu'1,Hu'2)); trivial.
+exists (Fmult (Float beta b (-1)) (Fplus d u'))%R.
+split.
+rewrite F2R_mult, F2R_plus, Hu'1.
+unfold m; rewrite Rmult_comm.
+unfold Rdiv; apply f_equal.
+unfold F2R; simpl; unfold Z.pow_pos; simpl.
+rewrite Zmult_1_r, Hb, mult_IZR.
+simpl; field.
+apply Rgt_not_eq, Rmult_lt_reg_l with (1 := Rlt_0_2).
+rewrite Rmult_0_r, <- (mult_IZR 2), <-Hb.
+apply radix_pos.
+apply trans_eq with (-1+Fexp (Fplus d u'))%Z.
+unfold Fmult.
+destruct (Fplus d u'); reflexivity.
+rewrite Zplus_comm; unfold Zminus; apply f_equal2.
+2: reflexivity.
+rewrite Fexp_Fplus.
+rewrite Z.min_l.
+now rewrite Fexp_d.
+rewrite Hu'2; omega.
+Qed.
+
+Lemma m_eq_0: (0 = F2R d)%R -> exists f:float beta,
+ F2R f = m /\ (Fexp f = fexp (mag beta (F2R u)) -1)%Z.
+Proof with auto with typeclass_instances.
+intros Y.
+specialize (Zeven_ex (radix_val beta)); rewrite Even_beta.
+intros (b, Hb); rewrite Zplus_0_r in Hb.
+exists (Fmult (Float beta b (-1)) u)%R.
+split.
+rewrite F2R_mult; unfold m; rewrite <- Y, Rplus_0_l.
+rewrite Rmult_comm.
+unfold Rdiv; apply f_equal.
+unfold F2R; simpl; unfold Z.pow_pos; simpl.
+rewrite Zmult_1_r, Hb, mult_IZR.
+simpl; field.
+apply Rgt_not_eq, Rmult_lt_reg_l with (1 := Rlt_0_2).
+rewrite Rmult_0_r, <- (mult_IZR 2), <-Hb.
+apply radix_pos.
+apply trans_eq with (-1+Fexp u)%Z.
+unfold Fmult.
+destruct u; reflexivity.
+rewrite Zplus_comm, Cu; unfold Zminus; now apply f_equal2.
+Qed.
+
+Lemma fexp_m_eq_0: (0 = F2R d)%R ->
+ (fexp (mag beta (F2R u)-1) < fexp (mag beta (F2R u))+1)%Z.
+Proof with auto with typeclass_instances.
+intros Y.
+assert ((fexp (mag beta (F2R u) - 1) <= fexp (mag beta (F2R u))))%Z.
+2: omega.
+destruct (mag beta x) as (e,He).
+rewrite Rabs_right in He.
+2: now left.
+assert (e <= fexp e)%Z.
+apply exp_small_round_0_pos with beta (Zfloor) x...
+now apply He, Rgt_not_eq.
+now rewrite <- d_eq, Y.
+rewrite u_eq, round_UP_small_pos with (ex:=e); trivial.
+2: now apply He, Rgt_not_eq.
+rewrite mag_bpow.
+ring_simplify (fexp e + 1 - 1)%Z.
+replace (fexp (fexp e)) with (fexp e).
+case exists_NE_; intros V.
+contradict V; rewrite Even_beta; discriminate.
+rewrite (proj2 (V e)); omega.
+apply sym_eq, valid_exp; omega.
+Qed.
+
+Lemma Fm: generic_format beta fexpe m.
+Proof.
+case (d_ge_0); intros Y.
+(* *)
+destruct m_eq as (g,(Hg1,Hg2)); trivial.
+apply generic_format_F2R' with g.
+now apply sym_eq.
+intros H; unfold cexp; rewrite Hg2.
+rewrite mag_m; trivial.
+rewrite <- Fexp_d; trivial.
+rewrite Cd.
+unfold cexp.
+generalize (fexpe_fexp (mag beta (F2R d))).
+omega.
+(* *)
+destruct m_eq_0 as (g,(Hg1,Hg2)); trivial.
+apply generic_format_F2R' with g.
+assumption.
+intros H; unfold cexp; rewrite Hg2.
+rewrite mag_m_0; try assumption.
+apply Z.le_trans with (1:=fexpe_fexp _).
+generalize (fexp_m_eq_0 Y).
+omega.
+Qed.
+
+
+
+Lemma Zm:
+ exists g : float beta, F2R g = m /\ canonical beta fexpe g /\ Z.even (Fnum g) = true.
+Proof with auto with typeclass_instances.
+case (d_ge_0); intros Y.
+(* *)
+destruct m_eq as (g,(Hg1,Hg2)); trivial.
+apply exists_even_fexp_lt.
+exists g; split; trivial.
+rewrite Hg2.
+rewrite mag_m; trivial.
+rewrite <- Fexp_d; trivial.
+rewrite Cd.
+unfold cexp.
+generalize (fexpe_fexp (mag beta (F2R d))).
+omega.
+(* *)
+destruct m_eq_0 as (g,(Hg1,Hg2)); trivial.
+apply exists_even_fexp_lt.
+exists g; split; trivial.
+rewrite Hg2.
+rewrite mag_m_0; trivial.
+apply Z.le_lt_trans with (1:=fexpe_fexp _).
+generalize (fexp_m_eq_0 Y).
+omega.
+Qed.
+
+
+Lemma DN_odd_d_aux :
+ forall z, (F2R d <= z < F2R u)%R ->
+ Rnd_DN_pt (generic_format beta fexp) z (F2R d).
+Proof with auto with typeclass_instances.
+intros z Hz1.
+replace (F2R d) with (round beta fexp Zfloor z).
+apply round_DN_pt...
+case (Rnd_DN_UP_pt_split _ _ _ _ Hd Hu (round beta fexp Zfloor z)).
+apply generic_format_round...
+intros Y; apply Rle_antisym; trivial.
+apply round_DN_pt...
+apply Hd.
+apply Hz1.
+intros Y ; elim (Rlt_irrefl z).
+apply Rlt_le_trans with (1:=proj2 Hz1), Rle_trans with (1:=Y).
+apply round_DN_pt...
+Qed.
+
+Lemma UP_odd_d_aux :
+ forall z, (F2R d < z <= F2R u)%R ->
+ Rnd_UP_pt (generic_format beta fexp) z (F2R u).
+Proof with auto with typeclass_instances.
+intros z Hz1.
+replace (F2R u) with (round beta fexp Zceil z).
+apply round_UP_pt...
+case (Rnd_DN_UP_pt_split _ _ _ _ Hd Hu (round beta fexp Zceil z)).
+apply generic_format_round...
+intros Y ; elim (Rlt_irrefl z).
+apply Rle_lt_trans with (2:=proj1 Hz1), Rle_trans with (2:=Y).
+apply round_UP_pt...
+intros Y; apply Rle_antisym; trivial.
+apply round_UP_pt...
+apply Hu.
+apply Hz1.
+Qed.
+
+
+Lemma round_N_odd_pos :
+ round beta fexp (Znearest choice) (round beta fexpe Zrnd_odd x) =
+ round beta fexp (Znearest choice) x.
+Proof with auto with typeclass_instances.
+set (o:=round beta fexpe Zrnd_odd x).
+case (generic_format_EM beta fexp x); intros Hx.
+replace o with x; trivial.
+unfold o; apply sym_eq, round_generic...
+now apply generic_format_fexpe_fexp.
+assert (K1:(F2R d <= o)%R).
+apply round_ge_generic...
+apply generic_format_fexpe_fexp, Hd.
+apply Hd.
+assert (K2:(o <= F2R u)%R).
+apply round_le_generic...
+apply generic_format_fexpe_fexp, Hu.
+apply Hu.
+assert (P:(x <> m -> o=m -> (forall P:Prop, P))).
+intros Y1 Y2.
+assert (H:(Rnd_odd_pt beta fexpe x o)).
+apply round_odd_pt...
+destruct H as (_,H); destruct H.
+absurd (x=m)%R; try trivial.
+now rewrite <- Y2, H.
+destruct H as (_,(k,(Hk1,(Hk2,Hk3)))).
+destruct Zm as (k',(Hk'1,(Hk'2,Hk'3))).
+absurd (true=false).
+discriminate.
+rewrite <- Hk3, <- Hk'3.
+apply f_equal, f_equal.
+apply canonical_unique with fexpe...
+now rewrite Hk'1, <- Y2.
+assert (generic_format beta fexp o -> (forall P:Prop, P)).
+intros Y.
+assert (H:(Rnd_odd_pt beta fexpe x o)).
+apply round_odd_pt...
+destruct H as (_,H); destruct H.
+absurd (generic_format beta fexp x); trivial.
+now rewrite <- H.
+destruct H as (_,(k,(Hk1,(Hk2,Hk3)))).
+destruct (exists_even_fexp_lt fexpe o) as (k',(Hk'1,(Hk'2,Hk'3))).
+eexists; split.
+apply sym_eq, Y.
+simpl; unfold cexp.
+apply Z.le_lt_trans with (1:=fexpe_fexp _).
+omega.
+absurd (true=false).
+discriminate.
+rewrite <- Hk3, <- Hk'3.
+apply f_equal, f_equal.
+apply canonical_unique with fexpe...
+now rewrite Hk'1, <- Hk1.
+case K1; clear K1; intros K1.
+2: apply H; rewrite <- K1; apply Hd.
+case K2; clear K2; intros K2.
+2: apply H; rewrite K2; apply Hu.
+case (Rle_or_lt x m); intros Y;[destruct Y|idtac].
+(* . *)
+apply trans_eq with (F2R d).
+apply round_N_eq_DN_pt with (F2R u)...
+apply DN_odd_d_aux; split; try left; assumption.
+apply UP_odd_d_aux; split; try left; assumption.
+assert (o <= (F2R d + F2R u) / 2)%R.
+apply round_le_generic...
+apply Fm.
+now left.
+destruct H1; trivial.
+apply P.
+now apply Rlt_not_eq.
+trivial.
+apply sym_eq, round_N_eq_DN_pt with (F2R u)...
+(* . *)
+replace o with x.
+reflexivity.
+apply sym_eq, round_generic...
+rewrite H0; apply Fm.
+(* . *)
+apply trans_eq with (F2R u).
+apply round_N_eq_UP_pt with (F2R d)...
+apply DN_odd_d_aux; split; try left; assumption.
+apply UP_odd_d_aux; split; try left; assumption.
+assert ((F2R d + F2R u) / 2 <= o)%R.
+apply round_ge_generic...
+apply Fm.
+now left.
+destruct H0; trivial.
+apply P.
+now apply Rgt_not_eq.
+rewrite <- H0; trivial.
+apply sym_eq, round_N_eq_UP_pt with (F2R d)...
+Qed.
+
+
+End Odd_prop_aux.
+
+Section Odd_prop.
+
+Variable beta : radix.
+Hypothesis Even_beta: Z.even (radix_val beta)=true.
+
+Variable fexp : Z -> Z.
+Variable fexpe : Z -> Z.
+Variable choice:Z->bool.
+
+Context { valid_exp : Valid_exp fexp }.
+Context { exists_NE_ : Exists_NE beta fexp }. (* for underflow reason *)
+Context { valid_expe : Valid_exp fexpe }.
+Context { exists_NE_e : Exists_NE beta fexpe }. (* for defining rounding to odd *)
+
+Hypothesis fexpe_fexp: forall e, (fexpe e <= fexp e -2)%Z.
+
+Theorem round_N_odd :
+ forall x,
+ round beta fexp (Znearest choice) (round beta fexpe Zrnd_odd x) =
+ round beta fexp (Znearest choice) x.
+Proof with auto with typeclass_instances.
+intros x.
+case (total_order_T x 0); intros H; [case H; clear H; intros H | idtac].
+rewrite <- (Ropp_involutive x).
+rewrite round_odd_opp.
+rewrite 2!round_N_opp.
+apply f_equal.
+destruct (canonical_generic_format beta fexp (round beta fexp Zfloor (-x))) as (d,(Hd1,Hd2)).
+apply generic_format_round...
+destruct (canonical_generic_format beta fexp (round beta fexp Zceil (-x))) as (u,(Hu1,Hu2)).
+apply generic_format_round...
+apply round_N_odd_pos with d u...
+rewrite <- Hd1; apply round_DN_pt...
+rewrite <- Hu1; apply round_UP_pt...
+auto with real.
+(* . *)
+rewrite H; repeat rewrite round_0...
+(* . *)
+destruct (canonical_generic_format beta fexp (round beta fexp Zfloor x)) as (d,(Hd1,Hd2)).
+apply generic_format_round...
+destruct (canonical_generic_format beta fexp (round beta fexp Zceil x)) as (u,(Hu1,Hu2)).
+apply generic_format_round...
+apply round_N_odd_pos with d u...
+rewrite <- Hd1; apply round_DN_pt...
+rewrite <- Hu1; apply round_UP_pt...
+Qed.
+
+End Odd_prop.
+
+
+Section Odd_propbis.
+
+Variable beta : radix.
+Hypothesis Even_beta: Z.even (radix_val beta)=true.
+
+Variable emin prec:Z.
+Variable choice:Z->bool.
+
+Hypothesis prec_gt_1: (1 < prec)%Z.
+
+
+Notation format := (generic_format beta (FLT_exp emin prec)).
+Notation round_flt :=(round beta (FLT_exp emin prec) (Znearest choice)).
+Notation cexp_flt := (cexp beta (FLT_exp emin prec)).
+Notation fexpe k := (FLT_exp (emin-k) (prec+k)).
+
+
+
+Lemma Zrnd_odd_plus': forall x y,
+ (exists n:Z, exists e:Z, (x = IZR n*bpow beta e)%R /\ (1 <= e)%Z) ->
+ (IZR (Zrnd_odd (x+y)) = x+IZR (Zrnd_odd y))%R.
+Proof.
+intros x y (n,(e,(H1,H2))).
+apply Zrnd_odd_plus.
+rewrite H1.
+rewrite <- IZR_Zpower.
+2: auto with zarith.
+now rewrite <- mult_IZR, Zfloor_IZR.
+rewrite H1, <- IZR_Zpower.
+2: auto with zarith.
+rewrite <- mult_IZR, Zfloor_IZR.
+rewrite Z.even_mul.
+rewrite Z.even_pow.
+2: auto with zarith.
+rewrite Even_beta.
+apply Bool.orb_true_iff; now right.
+Qed.
+
+
+
+Theorem mag_round_odd: forall (x:R),
+ (emin < mag beta x)%Z ->
+ (mag_val beta _ (mag beta (round beta (FLT_exp emin prec) Zrnd_odd x))
+ = mag_val beta x (mag beta x))%Z.
+Proof with auto with typeclass_instances.
+intros x.
+assert (T:Prec_gt_0 prec).
+unfold Prec_gt_0; auto with zarith.
+case (Req_dec x 0); intros Zx.
+intros _; rewrite Zx, round_0...
+destruct (mag beta x) as (e,He); simpl; intros H.
+apply mag_unique; split.
+apply abs_round_ge_generic...
+apply FLT_format_bpow...
+auto with zarith.
+now apply He.
+assert (V:
+ (Rabs (round beta (FLT_exp emin prec) Zrnd_odd x) <= bpow beta e)%R).
+apply abs_round_le_generic...
+apply FLT_format_bpow...
+auto with zarith.
+left; now apply He.
+case V; try easy; intros K.
+assert (H0:Rnd_odd_pt beta (FLT_exp emin prec) x (round beta (FLT_exp emin prec) Zrnd_odd x)).
+apply round_odd_pt...
+destruct H0 as (_,HH); destruct HH as [H0|(H0,(g,(Hg1,(Hg2,Hg3))))].
+absurd (Rabs x < bpow beta e)%R.
+apply Rle_not_lt; right.
+now rewrite <- H0,K.
+now apply He.
+pose (gg:=Float beta (Zpower beta (e-FLT_exp emin prec (e+1))) (FLT_exp emin prec (e+1))).
+assert (Y1: F2R gg = bpow beta e).
+unfold F2R; simpl.
+rewrite IZR_Zpower.
+rewrite <- bpow_plus.
+f_equal; ring.
+assert (FLT_exp emin prec (e+1) <= e)%Z; [idtac|auto with zarith].
+unfold FLT_exp.
+apply Z.max_case_strong; auto with zarith.
+assert (Y2: canonical beta (FLT_exp emin prec) gg).
+unfold canonical; rewrite Y1; unfold gg; simpl.
+unfold cexp; now rewrite mag_bpow.
+assert (Y3: Fnum gg = Z.abs (Fnum g)).
+apply trans_eq with (Fnum (Fabs g)).
+2: destruct g; unfold Fabs; now simpl.
+f_equal.
+apply canonical_unique with (FLT_exp emin prec); try assumption.
+destruct g; unfold Fabs; apply canonical_abs; easy.
+now rewrite Y1, F2R_abs, <- Hg1,K.
+assert (Y4: Z.even (Fnum gg) = true).
+unfold gg; simpl.
+rewrite Z.even_pow; try assumption.
+assert (FLT_exp emin prec (e+1) < e)%Z; [idtac|auto with zarith].
+unfold FLT_exp.
+apply Z.max_case_strong; auto with zarith.
+absurd (true = false).
+discriminate.
+rewrite <- Hg3.
+rewrite <- Zeven_abs.
+now rewrite <- Y3.
+Qed.
+
+Theorem fexp_round_odd: forall (x:R),
+ (cexp_flt (round beta (FLT_exp emin prec) Zrnd_odd x)
+ = cexp_flt x)%Z.
+Proof with auto with typeclass_instances.
+intros x.
+assert (G0:Valid_exp (FLT_exp emin prec)).
+apply FLT_exp_valid; unfold Prec_gt_0; auto with zarith.
+case (Req_dec x 0); intros Zx.
+rewrite Zx, round_0...
+case (Zle_or_lt (mag beta x) emin).
+unfold cexp; destruct (mag beta x) as (e,He); simpl.
+intros H; unfold FLT_exp at 4.
+rewrite Z.max_r.
+2: auto with zarith.
+apply Z.max_r.
+assert (G: Rabs (round beta (FLT_exp emin prec) Zrnd_odd x) = bpow beta emin).
+assert (G1: (Rabs (round beta (FLT_exp emin prec) Zrnd_odd x) <= bpow beta emin)%R).
+apply abs_round_le_generic...
+apply generic_format_bpow'...
+unfold FLT_exp; rewrite Z.max_r; auto with zarith.
+left; apply Rlt_le_trans with (bpow beta e).
+now apply He.
+now apply bpow_le.
+assert (G2: (0 <= Rabs (round beta (FLT_exp emin prec) Zrnd_odd x))%R).
+apply Rabs_pos.
+assert (G3: (Rabs (round beta (FLT_exp emin prec) Zrnd_odd x) <> 0)%R).
+assert (H0:Rnd_odd_pt beta (FLT_exp emin prec) x
+ (round beta (FLT_exp emin prec) Zrnd_odd x)).
+apply round_odd_pt...
+destruct H0 as (_,H0); destruct H0 as [H0|(_,(g,(Hg1,(Hg2,Hg3))))].
+apply Rgt_not_eq; rewrite H0.
+apply Rlt_le_trans with (bpow beta (e-1)).
+apply bpow_gt_0.
+now apply He.
+rewrite Hg1; intros K.
+contradict Hg3.
+replace (Fnum g) with 0%Z.
+easy.
+case (Z.eq_dec (Fnum g) Z0); intros W; try easy.
+contradict K.
+apply Rabs_no_R0.
+now apply F2R_neq_0.
+apply Rle_antisym; try assumption.
+apply Rle_trans with (succ beta (FLT_exp emin prec) 0).
+right; rewrite succ_0.
+rewrite ulp_FLT_small; try easy.
+unfold Prec_gt_0; auto with zarith.
+rewrite Rabs_R0; apply bpow_gt_0.
+apply succ_le_lt...
+apply generic_format_0.
+apply generic_format_abs; apply generic_format_round...
+case G2; [easy|intros; now contradict G3].
+rewrite <- mag_abs.
+rewrite G, mag_bpow; auto with zarith.
+intros H; unfold cexp.
+now rewrite mag_round_odd.
+Qed.
+
+
+
+
+End Odd_propbis.
+
+
diff --git a/flocq/Prop/Fprop_Sterbenz.v b/flocq/Prop/Sterbenz.v
index 4e74f889..746b7026 100644
--- a/flocq/Prop/Fprop_Sterbenz.v
+++ b/flocq/Prop/Sterbenz.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
@@ -19,10 +19,7 @@ COPYING file for more details.
(** * Sterbenz conditions for exact subtraction *)
-Require Import Fcore_Raux.
-Require Import Fcore_defs.
-Require Import Fcore_generic_fmt.
-Require Import Fcalc_ops.
+Require Import Raux Defs Generic_fmt Operations.
Section Fprop_Sterbenz.
@@ -37,7 +34,7 @@ Notation format := (generic_format beta fexp).
Theorem generic_format_plus :
forall x y,
format x -> format y ->
- (Rabs (x + y) < bpow (Zmin (ln_beta beta x) (ln_beta beta y)))%R ->
+ (Rabs (x + y) <= bpow (Z.min (mag beta x) (mag beta y)))%R ->
format (x + y)%R.
Proof.
intros x y Fx Fy Hxy.
@@ -48,44 +45,51 @@ destruct (Req_dec x R0) as [Zx|Zx].
now rewrite Zx, Rplus_0_l.
destruct (Req_dec y R0) as [Zy|Zy].
now rewrite Zy, Rplus_0_r.
+destruct Hxy as [Hxy|Hxy].
revert Hxy.
-destruct (ln_beta beta x) as (ex, Ex). simpl.
+destruct (mag beta x) as (ex, Ex). simpl.
specialize (Ex Zx).
-destruct (ln_beta beta y) as (ey, Ey). simpl.
+destruct (mag beta y) as (ey, Ey). simpl.
specialize (Ey Zy).
intros Hxy.
set (fx := Float beta (Ztrunc (scaled_mantissa beta fexp x)) (fexp ex)).
assert (Hx: x = F2R fx).
rewrite Fx at 1.
-unfold canonic_exp.
-now rewrite ln_beta_unique with (1 := Ex).
+unfold cexp.
+now rewrite mag_unique with (1 := Ex).
set (fy := Float beta (Ztrunc (scaled_mantissa beta fexp y)) (fexp ey)).
assert (Hy: y = F2R fy).
rewrite Fy at 1.
-unfold canonic_exp.
-now rewrite ln_beta_unique with (1 := Ey).
+unfold cexp.
+now rewrite mag_unique with (1 := Ey).
rewrite Hx, Hy.
rewrite <- F2R_plus.
apply generic_format_F2R.
intros _.
-case_eq (Fplus beta fx fy).
+case_eq (Fplus fx fy).
intros mxy exy Pxy.
rewrite <- Pxy, F2R_plus, <- Hx, <- Hy.
-unfold canonic_exp.
-replace exy with (fexp (Zmin ex ey)).
+unfold cexp.
+replace exy with (fexp (Z.min ex ey)).
apply monotone_exp.
-now apply ln_beta_le_bpow.
-replace exy with (Fexp (Fplus beta fx fy)) by exact (f_equal Fexp Pxy).
+now apply mag_le_bpow.
+replace exy with (Fexp (Fplus fx fy)) by exact (f_equal Fexp Pxy).
rewrite Fexp_Fplus.
simpl. clear -monotone_exp.
apply sym_eq.
destruct (Zmin_spec ex ey) as [(H1,H2)|(H1,H2)] ; rewrite H2.
-apply Zmin_l.
+apply Z.min_l.
now apply monotone_exp.
-apply Zmin_r.
+apply Z.min_r.
apply monotone_exp.
apply Zlt_le_weak.
-now apply Zgt_lt.
+now apply Z.gt_lt.
+apply generic_format_abs_inv.
+rewrite Hxy.
+apply generic_format_bpow.
+apply valid_exp.
+case (Zmin_spec (mag beta x) (mag beta y)); intros (H1,H2);
+ rewrite H2; now apply mag_generic_gt.
Qed.
Theorem generic_format_plus_weak :
@@ -100,17 +104,17 @@ now rewrite Zx, Rplus_0_l.
destruct (Req_dec y R0) as [Zy|Zy].
now rewrite Zy, Rplus_0_r.
apply generic_format_plus ; try assumption.
-apply Rle_lt_trans with (1 := Hxy).
+apply Rle_trans with (1 := Hxy).
unfold Rmin.
destruct (Rle_dec (Rabs x) (Rabs y)) as [Hxy'|Hxy'].
-rewrite Zmin_l.
-destruct (ln_beta beta x) as (ex, Hx).
-now apply Hx.
-now apply ln_beta_le_abs.
-rewrite Zmin_r.
-destruct (ln_beta beta y) as (ex, Hy).
-now apply Hy.
-apply ln_beta_le_abs.
+rewrite Z.min_l.
+destruct (mag beta x) as (ex, Hx).
+apply Rlt_le; now apply Hx.
+now apply mag_le_abs.
+rewrite Z.min_r.
+destruct (mag beta y) as (ex, Hy).
+apply Rlt_le; now apply Hy.
+apply mag_le_abs.
exact Zy.
apply Rlt_le.
now apply Rnot_le_lt.